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_or *) | id_or : symb (* id_gen_tag *) | id_gen_tag : symb (* id_record_new *) | id_record_new : symb (* id_a *) | id_a : symb (* id_locker2_release_lock *) | id_locker2_release_lock : symb (* id_eqt *) | id_eqt : symb (* id_istops *) | id_istops : symb (* id_locker2_map_add_pending *) | id_locker2_map_add_pending : symb (* id_release *) | id_release : symb (* id_locker2_obtainable *) | id_locker2_obtainable : symb (* id_imp *) | id_imp : symb (* id_stack *) | id_stack : symb (* id_locker2_map_promote_pending *) | id_locker2_map_promote_pending : symb (* id_locker *) | id_locker : symb (* id_case4 *) | id_case4 : symb (* id_int *) | id_int : symb (* id_push *) | id_push : symb (* id_locker2_add_pending *) | id_locker2_add_pending : symb (* id_true *) | id_true : symb (* id_locker2_check_availables *) | id_locker2_check_availables : symb (* id_F *) | id_F : symb (* id_eqs *) | id_eqs : symb (* id_record_update *) | id_record_update : symb (* id_false *) | id_false : symb (* id_gen_modtageq *) | id_gen_modtageq : symb (* id_undefined *) | id_undefined : symb (* id_nocalls *) | id_nocalls : symb (* id_locker2_remove_pending *) | id_locker2_remove_pending : symb (* id_resource *) | id_resource : symb (* id_case6 *) | id_case6 : symb (* id_if *) | id_if : symb (* id_pops *) | id_pops : symb (* id_locker2_map_claim_lock *) | id_locker2_map_claim_lock : symb (* id_ok *) | id_ok : symb (* id_case5 *) | id_case5 : symb (* id_tuple *) | id_tuple : symb (* id_member *) | id_member : symb (* id_s *) | id_s : symb (* id_delete *) | id_delete : symb (* id_T *) | id_T : symb (* id_case9 *) | id_case9 : symb (* id_record_extract *) | id_record_extract : symb (* id_excl *) | id_excl : symb (* id_case2 *) | id_case2 : symb (* id_nil *) | id_nil : symb (* id_eqc *) | id_eqc : symb (* id_case0 *) | id_case0 : symb (* id_request *) | id_request : symb (* id_locker2_check_available *) | id_locker2_check_available : symb (* id_not *) | id_not : symb (* id_pushs *) | id_pushs : symb (* id_locker2_promote_pending *) | id_locker2_promote_pending : symb (* id_mcrlrecord *) | id_mcrlrecord : symb (* id_locker2_obtainables *) | id_locker2_obtainables : symb (* id_cons *) | id_cons : symb (* id_push1 *) | id_push1 : symb (* id_case1 *) | id_case1 : symb (* id_element *) | id_element : symb (* id_locker2_adduniq *) | id_locker2_adduniq : symb (* id_and *) | id_and : symb (* id_empty *) | id_empty : symb (* id_record_updates *) | id_record_updates : symb (* id_lock *) | id_lock : symb (* id_excllock *) | id_excllock : symb (* id_pid *) | id_pid : symb (* id_calls *) | id_calls : symb (* id_subtract *) | id_subtract : symb (* id_tag *) | id_tag : symb (* id_equal *) | id_equal : symb (* id_eq *) | id_eq : symb (* id_tops *) | id_tops : symb (* id_locker2_claim_lock *) | id_locker2_claim_lock : symb (* id_pending *) | id_pending : symb (* id_andt *) | id_andt : symb (* id_tuplenil *) | id_tuplenil : symb (* id_append *) | id_append : symb (* id_0 *) | id_0 : symb (* id_case8 *) | id_case8 : symb . Definition symb_eq_bool (f1 f2:symb) : bool := match f1,f2 with | id_or,id_or => true | id_gen_tag,id_gen_tag => true | id_record_new,id_record_new => true | id_a,id_a => true | id_locker2_release_lock,id_locker2_release_lock => true | id_eqt,id_eqt => true | id_istops,id_istops => true | id_locker2_map_add_pending,id_locker2_map_add_pending => true | id_release,id_release => true | id_locker2_obtainable,id_locker2_obtainable => true | id_imp,id_imp => true | id_stack,id_stack => true | id_locker2_map_promote_pending,id_locker2_map_promote_pending => true | id_locker,id_locker => true | id_case4,id_case4 => true | id_int,id_int => true | id_push,id_push => true | id_locker2_add_pending,id_locker2_add_pending => true | id_true,id_true => true | id_locker2_check_availables,id_locker2_check_availables => true | id_F,id_F => true | id_eqs,id_eqs => true | id_record_update,id_record_update => true | id_false,id_false => true | id_gen_modtageq,id_gen_modtageq => true | id_undefined,id_undefined => true | id_nocalls,id_nocalls => true | id_locker2_remove_pending,id_locker2_remove_pending => true | id_resource,id_resource => true | id_case6,id_case6 => true | id_if,id_if => true | id_pops,id_pops => true | id_locker2_map_claim_lock,id_locker2_map_claim_lock => true | id_ok,id_ok => true | id_case5,id_case5 => true | id_tuple,id_tuple => true | id_member,id_member => true | id_s,id_s => true | id_delete,id_delete => true | id_T,id_T => true | id_case9,id_case9 => true | id_record_extract,id_record_extract => true | id_excl,id_excl => true | id_case2,id_case2 => true | id_nil,id_nil => true | id_eqc,id_eqc => true | id_case0,id_case0 => true | id_request,id_request => true | id_locker2_check_available,id_locker2_check_available => true | id_not,id_not => true | id_pushs,id_pushs => true | id_locker2_promote_pending,id_locker2_promote_pending => true | id_mcrlrecord,id_mcrlrecord => true | id_locker2_obtainables,id_locker2_obtainables => true | id_cons,id_cons => true | id_push1,id_push1 => true | id_case1,id_case1 => true | id_element,id_element => true | id_locker2_adduniq,id_locker2_adduniq => true | id_and,id_and => true | id_empty,id_empty => true | id_record_updates,id_record_updates => true | id_lock,id_lock => true | id_excllock,id_excllock => true | id_pid,id_pid => true | id_calls,id_calls => true | id_subtract,id_subtract => true | id_tag,id_tag => true | id_equal,id_equal => true | id_eq,id_eq => true | id_tops,id_tops => true | id_locker2_claim_lock,id_locker2_claim_lock => true | id_pending,id_pending => true | id_andt,id_andt => true | id_tuplenil,id_tuplenil => true | id_append,id_append => true | id_0,id_0 => true | id_case8,id_case8 => 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_or,id_or => refl_equal _ | id_gen_tag,id_gen_tag => refl_equal _ | id_record_new,id_record_new => refl_equal _ | id_a,id_a => refl_equal _ | id_locker2_release_lock,id_locker2_release_lock => refl_equal _ | id_eqt,id_eqt => refl_equal _ | id_istops,id_istops => refl_equal _ | id_locker2_map_add_pending,id_locker2_map_add_pending => refl_equal _ | id_release,id_release => refl_equal _ | id_locker2_obtainable,id_locker2_obtainable => refl_equal _ | id_imp,id_imp => refl_equal _ | id_stack,id_stack => refl_equal _ | id_locker2_map_promote_pending, id_locker2_map_promote_pending => refl_equal _ | id_locker,id_locker => refl_equal _ | id_case4,id_case4 => refl_equal _ | id_int,id_int => refl_equal _ | id_push,id_push => refl_equal _ | id_locker2_add_pending,id_locker2_add_pending => refl_equal _ | id_true,id_true => refl_equal _ | id_locker2_check_availables,id_locker2_check_availables => refl_equal _ | id_F,id_F => refl_equal _ | id_eqs,id_eqs => refl_equal _ | id_record_update,id_record_update => refl_equal _ | id_false,id_false => refl_equal _ | id_gen_modtageq,id_gen_modtageq => refl_equal _ | id_undefined,id_undefined => refl_equal _ | id_nocalls,id_nocalls => refl_equal _ | id_locker2_remove_pending,id_locker2_remove_pending => refl_equal _ | id_resource,id_resource => refl_equal _ | id_case6,id_case6 => refl_equal _ | id_if,id_if => refl_equal _ | id_pops,id_pops => refl_equal _ | id_locker2_map_claim_lock,id_locker2_map_claim_lock => refl_equal _ | id_ok,id_ok => refl_equal _ | id_case5,id_case5 => refl_equal _ | id_tuple,id_tuple => refl_equal _ | id_member,id_member => refl_equal _ | id_s,id_s => refl_equal _ | id_delete,id_delete => refl_equal _ | id_T,id_T => refl_equal _ | id_case9,id_case9 => refl_equal _ | id_record_extract,id_record_extract => refl_equal _ | id_excl,id_excl => refl_equal _ | id_case2,id_case2 => refl_equal _ | id_nil,id_nil => refl_equal _ | id_eqc,id_eqc => refl_equal _ | id_case0,id_case0 => refl_equal _ | id_request,id_request => refl_equal _ | id_locker2_check_available,id_locker2_check_available => refl_equal _ | id_not,id_not => refl_equal _ | id_pushs,id_pushs => refl_equal _ | id_locker2_promote_pending,id_locker2_promote_pending => refl_equal _ | id_mcrlrecord,id_mcrlrecord => refl_equal _ | id_locker2_obtainables,id_locker2_obtainables => refl_equal _ | id_cons,id_cons => refl_equal _ | id_push1,id_push1 => refl_equal _ | id_case1,id_case1 => refl_equal _ | id_element,id_element => refl_equal _ | id_locker2_adduniq,id_locker2_adduniq => refl_equal _ | id_and,id_and => refl_equal _ | id_empty,id_empty => refl_equal _ | id_record_updates,id_record_updates => refl_equal _ | id_lock,id_lock => refl_equal _ | id_excllock,id_excllock => refl_equal _ | id_pid,id_pid => refl_equal _ | id_calls,id_calls => refl_equal _ | id_subtract,id_subtract => refl_equal _ | id_tag,id_tag => refl_equal _ | id_equal,id_equal => refl_equal _ | id_eq,id_eq => refl_equal _ | id_tops,id_tops => refl_equal _ | id_locker2_claim_lock,id_locker2_claim_lock => refl_equal _ | id_pending,id_pending => refl_equal _ | id_andt,id_andt => refl_equal _ | id_tuplenil,id_tuplenil => refl_equal _ | id_append,id_append => refl_equal _ | id_0,id_0 => refl_equal _ | id_case8,id_case8 => refl_equal _ | _,_ => _ end;intros abs;discriminate. Defined. Definition arity (f:symb) := match f with | id_or => term.Free 2 | id_gen_tag => term.Free 1 | id_record_new => term.Free 1 | id_a => term.Free 0 | id_locker2_release_lock => term.Free 2 | id_eqt => term.Free 2 | id_istops => term.Free 2 | id_locker2_map_add_pending => term.Free 3 | id_release => term.Free 0 | id_locker2_obtainable => term.Free 2 | id_imp => term.Free 2 | id_stack => term.Free 2 | id_locker2_map_promote_pending => term.Free 2 | id_locker => term.Free 0 | id_case4 => term.Free 3 | id_int => term.Free 1 | id_push => term.Free 3 | id_locker2_add_pending => term.Free 3 | id_true => term.Free 0 | id_locker2_check_availables => term.Free 2 | id_F => term.Free 0 | id_eqs => term.Free 2 | id_record_update => term.Free 4 | id_false => term.Free 0 | id_gen_modtageq => term.Free 2 | id_undefined => term.Free 0 | id_nocalls => term.Free 0 | id_locker2_remove_pending => term.Free 2 | id_resource => term.Free 0 | id_case6 => term.Free 4 | id_if => term.Free 3 | id_pops => term.Free 1 | id_locker2_map_claim_lock => term.Free 3 | id_ok => term.Free 0 | id_case5 => term.Free 4 | id_tuple => term.Free 2 | id_member => term.Free 2 | id_s => term.Free 1 | id_delete => term.Free 2 | id_T => term.Free 0 | id_case9 => term.Free 4 | id_record_extract => term.Free 3 | id_excl => term.Free 0 | id_case2 => term.Free 3 | id_nil => term.Free 0 | id_eqc => term.Free 2 | id_case0 => term.Free 3 | id_request => term.Free 0 | id_locker2_check_available => term.Free 2 | id_not => term.Free 1 | id_pushs => term.Free 2 | id_locker2_promote_pending => term.Free 2 | id_mcrlrecord => term.Free 0 | id_locker2_obtainables => term.Free 2 | id_cons => term.Free 2 | id_push1 => term.Free 6 | id_case1 => term.Free 4 | id_element => term.Free 2 | id_locker2_adduniq => term.Free 2 | id_and => term.Free 2 | id_empty => term.Free 0 | id_record_updates => term.Free 3 | id_lock => term.Free 0 | id_excllock => term.Free 0 | id_pid => term.Free 1 | id_calls => term.Free 3 | id_subtract => term.Free 2 | id_tag => term.Free 0 | id_equal => term.Free 2 | id_eq => term.Free 2 | id_tops => term.Free 1 | id_locker2_claim_lock => term.Free 3 | id_pending => term.Free 0 | id_andt => term.Free 2 | id_tuplenil => term.Free 1 | id_append => term.Free 2 | id_0 => term.Free 0 | id_case8 => term.Free 4 end. Definition symb_order (f1 f2:symb) : bool := match f1,f2 with | id_or,id_or => true | id_or,id_gen_tag => false | id_or,id_record_new => false | id_or,id_a => false | id_or,id_locker2_release_lock => false | id_or,id_eqt => false | id_or,id_istops => false | id_or,id_locker2_map_add_pending => false | id_or,id_release => false | id_or,id_locker2_obtainable => false | id_or,id_imp => false | id_or,id_stack => false | id_or,id_locker2_map_promote_pending => false | id_or,id_locker => false | id_or,id_case4 => false | id_or,id_int => false | id_or,id_push => false | id_or,id_locker2_add_pending => false | id_or,id_true => false | id_or,id_locker2_check_availables => false | id_or,id_F => false | id_or,id_eqs => false | id_or,id_record_update => false | id_or,id_false => false | id_or,id_gen_modtageq => false | id_or,id_undefined => false | id_or,id_nocalls => false | id_or,id_locker2_remove_pending => false | id_or,id_resource => false | id_or,id_case6 => false | id_or,id_if => false | id_or,id_pops => false | id_or,id_locker2_map_claim_lock => false | id_or,id_ok => false | id_or,id_case5 => false | id_or,id_tuple => false | id_or,id_member => false | id_or,id_s => false | id_or,id_delete => false | id_or,id_T => false | id_or,id_case9 => false | id_or,id_record_extract => false | id_or,id_excl => false | id_or,id_case2 => false | id_or,id_nil => false | id_or,id_eqc => false | id_or,id_case0 => false | id_or,id_request => false | id_or,id_locker2_check_available => false | id_or,id_not => false | id_or,id_pushs => false | id_or,id_locker2_promote_pending => false | id_or,id_mcrlrecord => false | id_or,id_locker2_obtainables => false | id_or,id_cons => false | id_or,id_push1 => false | id_or,id_case1 => false | id_or,id_element => false | id_or,id_locker2_adduniq => false | id_or,id_and => false | id_or,id_empty => false | id_or,id_record_updates => false | id_or,id_lock => false | id_or,id_excllock => false | id_or,id_pid => false | id_or,id_calls => false | id_or,id_subtract => false | id_or,id_tag => false | id_or,id_equal => false | id_or,id_eq => false | id_or,id_tops => false | id_or,id_locker2_claim_lock => false | id_or,id_pending => false | id_or,id_andt => false | id_or,id_tuplenil => false | id_or,id_append => false | id_or,id_0 => false | id_or,id_case8 => false | id_gen_tag,id_or => true | id_gen_tag,id_gen_tag => true | id_gen_tag,id_record_new => false | id_gen_tag,id_a => false | id_gen_tag,id_locker2_release_lock => false | id_gen_tag,id_eqt => false | id_gen_tag,id_istops => false | id_gen_tag,id_locker2_map_add_pending => false | id_gen_tag,id_release => false | id_gen_tag,id_locker2_obtainable => false | id_gen_tag,id_imp => false | id_gen_tag,id_stack => false | id_gen_tag,id_locker2_map_promote_pending => false | id_gen_tag,id_locker => false | id_gen_tag,id_case4 => false | id_gen_tag,id_int => false | id_gen_tag,id_push => false | id_gen_tag,id_locker2_add_pending => false | id_gen_tag,id_true => false | id_gen_tag,id_locker2_check_availables => false | id_gen_tag,id_F => false | id_gen_tag,id_eqs => false | id_gen_tag,id_record_update => false | id_gen_tag,id_false => false | id_gen_tag,id_gen_modtageq => false | id_gen_tag,id_undefined => false | id_gen_tag,id_nocalls => false | id_gen_tag,id_locker2_remove_pending => false | id_gen_tag,id_resource => false | id_gen_tag,id_case6 => false | id_gen_tag,id_if => false | id_gen_tag,id_pops => false | id_gen_tag,id_locker2_map_claim_lock => false | id_gen_tag,id_ok => false | id_gen_tag,id_case5 => false | id_gen_tag,id_tuple => false | id_gen_tag,id_member => false | id_gen_tag,id_s => false | id_gen_tag,id_delete => false | id_gen_tag,id_T => false | id_gen_tag,id_case9 => false | id_gen_tag,id_record_extract => false | id_gen_tag,id_excl => false | id_gen_tag,id_case2 => false | id_gen_tag,id_nil => false | id_gen_tag,id_eqc => false | id_gen_tag,id_case0 => false | id_gen_tag,id_request => false | id_gen_tag,id_locker2_check_available => false | id_gen_tag,id_not => false | id_gen_tag,id_pushs => false | id_gen_tag,id_locker2_promote_pending => false | id_gen_tag,id_mcrlrecord => false | id_gen_tag,id_locker2_obtainables => false | id_gen_tag,id_cons => false | id_gen_tag,id_push1 => false | id_gen_tag,id_case1 => false | id_gen_tag,id_element => false | id_gen_tag,id_locker2_adduniq => false | id_gen_tag,id_and => false | id_gen_tag,id_empty => false | id_gen_tag,id_record_updates => false | id_gen_tag,id_lock => false | id_gen_tag,id_excllock => false | id_gen_tag,id_pid => false | id_gen_tag,id_calls => false | id_gen_tag,id_subtract => false | id_gen_tag,id_tag => false | id_gen_tag,id_equal => false | id_gen_tag,id_eq => false | id_gen_tag,id_tops => false | id_gen_tag,id_locker2_claim_lock => false | id_gen_tag,id_pending => false | id_gen_tag,id_andt => false | id_gen_tag,id_tuplenil => false | id_gen_tag,id_append => false | id_gen_tag,id_0 => false | id_gen_tag,id_case8 => false | id_record_new,id_or => true | id_record_new,id_gen_tag => true | id_record_new,id_record_new => true | id_record_new,id_a => false | id_record_new,id_locker2_release_lock => false | id_record_new,id_eqt => false | id_record_new,id_istops => false | id_record_new,id_locker2_map_add_pending => false | id_record_new,id_release => false | id_record_new,id_locker2_obtainable => false | id_record_new,id_imp => false | id_record_new,id_stack => false | id_record_new,id_locker2_map_promote_pending => false | id_record_new,id_locker => false | id_record_new,id_case4 => false | id_record_new,id_int => false | id_record_new,id_push => false | id_record_new,id_locker2_add_pending => false | id_record_new,id_true => false | id_record_new,id_locker2_check_availables => false | id_record_new,id_F => false | id_record_new,id_eqs => false | id_record_new,id_record_update => false | id_record_new,id_false => false | id_record_new,id_gen_modtageq => false | id_record_new,id_undefined => false | id_record_new,id_nocalls => false | id_record_new,id_locker2_remove_pending => false | id_record_new,id_resource => false | id_record_new,id_case6 => false | id_record_new,id_if => false | id_record_new,id_pops => false | id_record_new,id_locker2_map_claim_lock => false | id_record_new,id_ok => false | id_record_new,id_case5 => false | id_record_new,id_tuple => false | id_record_new,id_member => false | id_record_new,id_s => false | id_record_new,id_delete => false | id_record_new,id_T => false | id_record_new,id_case9 => false | id_record_new,id_record_extract => false | id_record_new,id_excl => false | id_record_new,id_case2 => false | id_record_new,id_nil => false | id_record_new,id_eqc => false | id_record_new,id_case0 => false | id_record_new,id_request => false | id_record_new,id_locker2_check_available => false | id_record_new,id_not => false | id_record_new,id_pushs => false | id_record_new,id_locker2_promote_pending => false | id_record_new,id_mcrlrecord => false | id_record_new,id_locker2_obtainables => false | id_record_new,id_cons => false | id_record_new,id_push1 => false | id_record_new,id_case1 => false | id_record_new,id_element => false | id_record_new,id_locker2_adduniq => false | id_record_new,id_and => false | id_record_new,id_empty => false | id_record_new,id_record_updates => false | id_record_new,id_lock => false | id_record_new,id_excllock => false | id_record_new,id_pid => false | id_record_new,id_calls => false | id_record_new,id_subtract => false | id_record_new,id_tag => false | id_record_new,id_equal => false | id_record_new,id_eq => false | id_record_new,id_tops => false | id_record_new,id_locker2_claim_lock => false | id_record_new,id_pending => false | id_record_new,id_andt => false | id_record_new,id_tuplenil => false | id_record_new,id_append => false | id_record_new,id_0 => false | id_record_new,id_case8 => false | id_a,id_or => true | id_a,id_gen_tag => true | id_a,id_record_new => true | id_a,id_a => true | id_a,id_locker2_release_lock => false | id_a,id_eqt => false | id_a,id_istops => false | id_a,id_locker2_map_add_pending => false | id_a,id_release => false | id_a,id_locker2_obtainable => false | id_a,id_imp => false | id_a,id_stack => false | id_a,id_locker2_map_promote_pending => false | id_a,id_locker => false | id_a,id_case4 => false | id_a,id_int => false | id_a,id_push => false | id_a,id_locker2_add_pending => false | id_a,id_true => false | id_a,id_locker2_check_availables => false | id_a,id_F => false | id_a,id_eqs => false | id_a,id_record_update => false | id_a,id_false => false | id_a,id_gen_modtageq => false | id_a,id_undefined => false | id_a,id_nocalls => false | id_a,id_locker2_remove_pending => false | id_a,id_resource => false | id_a,id_case6 => false | id_a,id_if => false | id_a,id_pops => false | id_a,id_locker2_map_claim_lock => false | id_a,id_ok => false | id_a,id_case5 => false | id_a,id_tuple => false | id_a,id_member => false | id_a,id_s => false | id_a,id_delete => false | id_a,id_T => false | id_a,id_case9 => false | id_a,id_record_extract => false | id_a,id_excl => false | id_a,id_case2 => false | id_a,id_nil => false | id_a,id_eqc => false | id_a,id_case0 => false | id_a,id_request => false | id_a,id_locker2_check_available => false | id_a,id_not => false | id_a,id_pushs => false | id_a,id_locker2_promote_pending => false | id_a,id_mcrlrecord => false | id_a,id_locker2_obtainables => false | id_a,id_cons => false | id_a,id_push1 => false | id_a,id_case1 => false | id_a,id_element => false | id_a,id_locker2_adduniq => false | id_a,id_and => false | id_a,id_empty => false | id_a,id_record_updates => false | id_a,id_lock => false | id_a,id_excllock => false | id_a,id_pid => false | id_a,id_calls => false | id_a,id_subtract => false | id_a,id_tag => false | id_a,id_equal => false | id_a,id_eq => false | id_a,id_tops => false | id_a,id_locker2_claim_lock => false | id_a,id_pending => false | id_a,id_andt => false | id_a,id_tuplenil => false | id_a,id_append => false | id_a,id_0 => false | id_a,id_case8 => false | id_locker2_release_lock,id_or => true | id_locker2_release_lock,id_gen_tag => true | id_locker2_release_lock,id_record_new => true | id_locker2_release_lock,id_a => true | id_locker2_release_lock,id_locker2_release_lock => true | id_locker2_release_lock,id_eqt => false | id_locker2_release_lock,id_istops => false | id_locker2_release_lock,id_locker2_map_add_pending => false | id_locker2_release_lock,id_release => false | id_locker2_release_lock,id_locker2_obtainable => false | id_locker2_release_lock,id_imp => false | id_locker2_release_lock,id_stack => false | id_locker2_release_lock,id_locker2_map_promote_pending => false | id_locker2_release_lock,id_locker => false | id_locker2_release_lock,id_case4 => false | id_locker2_release_lock,id_int => false | id_locker2_release_lock,id_push => false | id_locker2_release_lock,id_locker2_add_pending => false | id_locker2_release_lock,id_true => false | id_locker2_release_lock,id_locker2_check_availables => false | id_locker2_release_lock,id_F => false | id_locker2_release_lock,id_eqs => false | id_locker2_release_lock,id_record_update => false | id_locker2_release_lock,id_false => false | id_locker2_release_lock,id_gen_modtageq => false | id_locker2_release_lock,id_undefined => false | id_locker2_release_lock,id_nocalls => false | id_locker2_release_lock,id_locker2_remove_pending => false | id_locker2_release_lock,id_resource => false | id_locker2_release_lock,id_case6 => false | id_locker2_release_lock,id_if => false | id_locker2_release_lock,id_pops => false | id_locker2_release_lock,id_locker2_map_claim_lock => false | id_locker2_release_lock,id_ok => false | id_locker2_release_lock,id_case5 => false | id_locker2_release_lock,id_tuple => false | id_locker2_release_lock,id_member => false | id_locker2_release_lock,id_s => false | id_locker2_release_lock,id_delete => false | id_locker2_release_lock,id_T => false | id_locker2_release_lock,id_case9 => false | id_locker2_release_lock,id_record_extract => false | id_locker2_release_lock,id_excl => false | id_locker2_release_lock,id_case2 => false | id_locker2_release_lock,id_nil => false | id_locker2_release_lock,id_eqc => false | id_locker2_release_lock,id_case0 => false | id_locker2_release_lock,id_request => false | id_locker2_release_lock,id_locker2_check_available => false | id_locker2_release_lock,id_not => false | id_locker2_release_lock,id_pushs => false | id_locker2_release_lock,id_locker2_promote_pending => false | id_locker2_release_lock,id_mcrlrecord => false | id_locker2_release_lock,id_locker2_obtainables => false | id_locker2_release_lock,id_cons => false | id_locker2_release_lock,id_push1 => false | id_locker2_release_lock,id_case1 => false | id_locker2_release_lock,id_element => false | id_locker2_release_lock,id_locker2_adduniq => false | id_locker2_release_lock,id_and => false | id_locker2_release_lock,id_empty => false | id_locker2_release_lock,id_record_updates => false | id_locker2_release_lock,id_lock => false | id_locker2_release_lock,id_excllock => false | id_locker2_release_lock,id_pid => false | id_locker2_release_lock,id_calls => false | id_locker2_release_lock,id_subtract => false | id_locker2_release_lock,id_tag => false | id_locker2_release_lock,id_equal => false | id_locker2_release_lock,id_eq => false | id_locker2_release_lock,id_tops => false | id_locker2_release_lock,id_locker2_claim_lock => false | id_locker2_release_lock,id_pending => false | id_locker2_release_lock,id_andt => false | id_locker2_release_lock,id_tuplenil => false | id_locker2_release_lock,id_append => false | id_locker2_release_lock,id_0 => false | id_locker2_release_lock,id_case8 => false | id_eqt,id_or => true | id_eqt,id_gen_tag => true | id_eqt,id_record_new => true | id_eqt,id_a => true | id_eqt,id_locker2_release_lock => true | id_eqt,id_eqt => true | id_eqt,id_istops => false | id_eqt,id_locker2_map_add_pending => false | id_eqt,id_release => false | id_eqt,id_locker2_obtainable => false | id_eqt,id_imp => false | id_eqt,id_stack => false | id_eqt,id_locker2_map_promote_pending => false | id_eqt,id_locker => false | id_eqt,id_case4 => false | id_eqt,id_int => false | id_eqt,id_push => false | id_eqt,id_locker2_add_pending => false | id_eqt,id_true => false | id_eqt,id_locker2_check_availables => false | id_eqt,id_F => false | id_eqt,id_eqs => false | id_eqt,id_record_update => false | id_eqt,id_false => false | id_eqt,id_gen_modtageq => false | id_eqt,id_undefined => false | id_eqt,id_nocalls => false | id_eqt,id_locker2_remove_pending => false | id_eqt,id_resource => false | id_eqt,id_case6 => false | id_eqt,id_if => false | id_eqt,id_pops => false | id_eqt,id_locker2_map_claim_lock => false | id_eqt,id_ok => false | id_eqt,id_case5 => false | id_eqt,id_tuple => false | id_eqt,id_member => false | id_eqt,id_s => false | id_eqt,id_delete => false | id_eqt,id_T => false | id_eqt,id_case9 => false | id_eqt,id_record_extract => false | id_eqt,id_excl => false | id_eqt,id_case2 => false | id_eqt,id_nil => false | id_eqt,id_eqc => false | id_eqt,id_case0 => false | id_eqt,id_request => false | id_eqt,id_locker2_check_available => false | id_eqt,id_not => false | id_eqt,id_pushs => false | id_eqt,id_locker2_promote_pending => false | id_eqt,id_mcrlrecord => false | id_eqt,id_locker2_obtainables => false | id_eqt,id_cons => false | id_eqt,id_push1 => false | id_eqt,id_case1 => false | id_eqt,id_element => false | id_eqt,id_locker2_adduniq => false | id_eqt,id_and => false | id_eqt,id_empty => false | id_eqt,id_record_updates => false | id_eqt,id_lock => false | id_eqt,id_excllock => false | id_eqt,id_pid => false | id_eqt,id_calls => false | id_eqt,id_subtract => false | id_eqt,id_tag => false | id_eqt,id_equal => false | id_eqt,id_eq => false | id_eqt,id_tops => false | id_eqt,id_locker2_claim_lock => false | id_eqt,id_pending => false | id_eqt,id_andt => false | id_eqt,id_tuplenil => false | id_eqt,id_append => false | id_eqt,id_0 => false | id_eqt,id_case8 => false | id_istops,id_or => true | id_istops,id_gen_tag => true | id_istops,id_record_new => true | id_istops,id_a => true | id_istops,id_locker2_release_lock => true | id_istops,id_eqt => true | id_istops,id_istops => true | id_istops,id_locker2_map_add_pending => false | id_istops,id_release => false | id_istops,id_locker2_obtainable => false | id_istops,id_imp => false | id_istops,id_stack => false | id_istops,id_locker2_map_promote_pending => false | id_istops,id_locker => false | id_istops,id_case4 => false | id_istops,id_int => false | id_istops,id_push => false | id_istops,id_locker2_add_pending => false | id_istops,id_true => false | id_istops,id_locker2_check_availables => false | id_istops,id_F => false | id_istops,id_eqs => false | id_istops,id_record_update => false | id_istops,id_false => false | id_istops,id_gen_modtageq => false | id_istops,id_undefined => false | id_istops,id_nocalls => false | id_istops,id_locker2_remove_pending => false | id_istops,id_resource => false | id_istops,id_case6 => false | id_istops,id_if => false | id_istops,id_pops => false | id_istops,id_locker2_map_claim_lock => false | id_istops,id_ok => false | id_istops,id_case5 => false | id_istops,id_tuple => false | id_istops,id_member => false | id_istops,id_s => false | id_istops,id_delete => false | id_istops,id_T => false | id_istops,id_case9 => false | id_istops,id_record_extract => false | id_istops,id_excl => false | id_istops,id_case2 => false | id_istops,id_nil => false | id_istops,id_eqc => false | id_istops,id_case0 => false | id_istops,id_request => false | id_istops,id_locker2_check_available => false | id_istops,id_not => false | id_istops,id_pushs => false | id_istops,id_locker2_promote_pending => false | id_istops,id_mcrlrecord => false | id_istops,id_locker2_obtainables => false | id_istops,id_cons => false | id_istops,id_push1 => false | id_istops,id_case1 => false | id_istops,id_element => false | id_istops,id_locker2_adduniq => false | id_istops,id_and => false | id_istops,id_empty => false | id_istops,id_record_updates => false | id_istops,id_lock => false | id_istops,id_excllock => false | id_istops,id_pid => false | id_istops,id_calls => false | id_istops,id_subtract => false | id_istops,id_tag => false | id_istops,id_equal => false | id_istops,id_eq => false | id_istops,id_tops => false | id_istops,id_locker2_claim_lock => false | id_istops,id_pending => false | id_istops,id_andt => false | id_istops,id_tuplenil => false | id_istops,id_append => false | id_istops,id_0 => false | id_istops,id_case8 => false | id_locker2_map_add_pending,id_or => true | id_locker2_map_add_pending,id_gen_tag => true | id_locker2_map_add_pending,id_record_new => true | id_locker2_map_add_pending,id_a => true | id_locker2_map_add_pending,id_locker2_release_lock => true | id_locker2_map_add_pending,id_eqt => true | id_locker2_map_add_pending,id_istops => true | id_locker2_map_add_pending,id_locker2_map_add_pending => true | id_locker2_map_add_pending,id_release => false | id_locker2_map_add_pending,id_locker2_obtainable => false | id_locker2_map_add_pending,id_imp => false | id_locker2_map_add_pending,id_stack => false | id_locker2_map_add_pending,id_locker2_map_promote_pending => false | id_locker2_map_add_pending,id_locker => false | id_locker2_map_add_pending,id_case4 => false | id_locker2_map_add_pending,id_int => false | id_locker2_map_add_pending,id_push => false | id_locker2_map_add_pending,id_locker2_add_pending => false | id_locker2_map_add_pending,id_true => false | id_locker2_map_add_pending,id_locker2_check_availables => false | id_locker2_map_add_pending,id_F => false | id_locker2_map_add_pending,id_eqs => false | id_locker2_map_add_pending,id_record_update => false | id_locker2_map_add_pending,id_false => false | id_locker2_map_add_pending,id_gen_modtageq => false | id_locker2_map_add_pending,id_undefined => false | id_locker2_map_add_pending,id_nocalls => false | id_locker2_map_add_pending,id_locker2_remove_pending => false | id_locker2_map_add_pending,id_resource => false | id_locker2_map_add_pending,id_case6 => false | id_locker2_map_add_pending,id_if => false | id_locker2_map_add_pending,id_pops => false | id_locker2_map_add_pending,id_locker2_map_claim_lock => false | id_locker2_map_add_pending,id_ok => false | id_locker2_map_add_pending,id_case5 => false | id_locker2_map_add_pending,id_tuple => false | id_locker2_map_add_pending,id_member => false | id_locker2_map_add_pending,id_s => false | id_locker2_map_add_pending,id_delete => false | id_locker2_map_add_pending,id_T => false | id_locker2_map_add_pending,id_case9 => false | id_locker2_map_add_pending,id_record_extract => false | id_locker2_map_add_pending,id_excl => false | id_locker2_map_add_pending,id_case2 => false | id_locker2_map_add_pending,id_nil => false | id_locker2_map_add_pending,id_eqc => false | id_locker2_map_add_pending,id_case0 => false | id_locker2_map_add_pending,id_request => false | id_locker2_map_add_pending,id_locker2_check_available => false | id_locker2_map_add_pending,id_not => false | id_locker2_map_add_pending,id_pushs => false | id_locker2_map_add_pending,id_locker2_promote_pending => false | id_locker2_map_add_pending,id_mcrlrecord => false | id_locker2_map_add_pending,id_locker2_obtainables => false | id_locker2_map_add_pending,id_cons => false | id_locker2_map_add_pending,id_push1 => false | id_locker2_map_add_pending,id_case1 => false | id_locker2_map_add_pending,id_element => false | id_locker2_map_add_pending,id_locker2_adduniq => false | id_locker2_map_add_pending,id_and => false | id_locker2_map_add_pending,id_empty => false | id_locker2_map_add_pending,id_record_updates => false | id_locker2_map_add_pending,id_lock => false | id_locker2_map_add_pending,id_excllock => false | id_locker2_map_add_pending,id_pid => false | id_locker2_map_add_pending,id_calls => false | id_locker2_map_add_pending,id_subtract => false | id_locker2_map_add_pending,id_tag => false | id_locker2_map_add_pending,id_equal => false | id_locker2_map_add_pending,id_eq => false | id_locker2_map_add_pending,id_tops => false | id_locker2_map_add_pending,id_locker2_claim_lock => false | id_locker2_map_add_pending,id_pending => false | id_locker2_map_add_pending,id_andt => false | id_locker2_map_add_pending,id_tuplenil => false | id_locker2_map_add_pending,id_append => false | id_locker2_map_add_pending,id_0 => false | id_locker2_map_add_pending,id_case8 => false | id_release,id_or => true | id_release,id_gen_tag => true | id_release,id_record_new => true | id_release,id_a => true | id_release,id_locker2_release_lock => true | id_release,id_eqt => true | id_release,id_istops => true | id_release,id_locker2_map_add_pending => true | id_release,id_release => true | id_release,id_locker2_obtainable => false | id_release,id_imp => false | id_release,id_stack => false | id_release,id_locker2_map_promote_pending => false | id_release,id_locker => false | id_release,id_case4 => false | id_release,id_int => false | id_release,id_push => false | id_release,id_locker2_add_pending => false | id_release,id_true => false | id_release,id_locker2_check_availables => false | id_release,id_F => false | id_release,id_eqs => false | id_release,id_record_update => false | id_release,id_false => false | id_release,id_gen_modtageq => false | id_release,id_undefined => false | id_release,id_nocalls => false | id_release,id_locker2_remove_pending => false | id_release,id_resource => false | id_release,id_case6 => false | id_release,id_if => false | id_release,id_pops => false | id_release,id_locker2_map_claim_lock => false | id_release,id_ok => false | id_release,id_case5 => false | id_release,id_tuple => false | id_release,id_member => false | id_release,id_s => false | id_release,id_delete => false | id_release,id_T => false | id_release,id_case9 => false | id_release,id_record_extract => false | id_release,id_excl => false | id_release,id_case2 => false | id_release,id_nil => false | id_release,id_eqc => false | id_release,id_case0 => false | id_release,id_request => false | id_release,id_locker2_check_available => false | id_release,id_not => false | id_release,id_pushs => false | id_release,id_locker2_promote_pending => false | id_release,id_mcrlrecord => false | id_release,id_locker2_obtainables => false | id_release,id_cons => false | id_release,id_push1 => false | id_release,id_case1 => false | id_release,id_element => false | id_release,id_locker2_adduniq => false | id_release,id_and => false | id_release,id_empty => false | id_release,id_record_updates => false | id_release,id_lock => false | id_release,id_excllock => false | id_release,id_pid => false | id_release,id_calls => false | id_release,id_subtract => false | id_release,id_tag => false | id_release,id_equal => false | id_release,id_eq => false | id_release,id_tops => false | id_release,id_locker2_claim_lock => false | id_release,id_pending => false | id_release,id_andt => false | id_release,id_tuplenil => false | id_release,id_append => false | id_release,id_0 => false | id_release,id_case8 => false | id_locker2_obtainable,id_or => true | id_locker2_obtainable,id_gen_tag => true | id_locker2_obtainable,id_record_new => true | id_locker2_obtainable,id_a => true | id_locker2_obtainable,id_locker2_release_lock => true | id_locker2_obtainable,id_eqt => true | id_locker2_obtainable,id_istops => true | id_locker2_obtainable,id_locker2_map_add_pending => true | id_locker2_obtainable,id_release => true | id_locker2_obtainable,id_locker2_obtainable => true | id_locker2_obtainable,id_imp => false | id_locker2_obtainable,id_stack => false | id_locker2_obtainable,id_locker2_map_promote_pending => false | id_locker2_obtainable,id_locker => false | id_locker2_obtainable,id_case4 => false | id_locker2_obtainable,id_int => false | id_locker2_obtainable,id_push => false | id_locker2_obtainable,id_locker2_add_pending => false | id_locker2_obtainable,id_true => false | id_locker2_obtainable,id_locker2_check_availables => false | id_locker2_obtainable,id_F => false | id_locker2_obtainable,id_eqs => false | id_locker2_obtainable,id_record_update => false | id_locker2_obtainable,id_false => false | id_locker2_obtainable,id_gen_modtageq => false | id_locker2_obtainable,id_undefined => false | id_locker2_obtainable,id_nocalls => false | id_locker2_obtainable,id_locker2_remove_pending => false | id_locker2_obtainable,id_resource => false | id_locker2_obtainable,id_case6 => false | id_locker2_obtainable,id_if => false | id_locker2_obtainable,id_pops => false | id_locker2_obtainable,id_locker2_map_claim_lock => false | id_locker2_obtainable,id_ok => false | id_locker2_obtainable,id_case5 => false | id_locker2_obtainable,id_tuple => false | id_locker2_obtainable,id_member => false | id_locker2_obtainable,id_s => false | id_locker2_obtainable,id_delete => false | id_locker2_obtainable,id_T => false | id_locker2_obtainable,id_case9 => false | id_locker2_obtainable,id_record_extract => false | id_locker2_obtainable,id_excl => false | id_locker2_obtainable,id_case2 => false | id_locker2_obtainable,id_nil => false | id_locker2_obtainable,id_eqc => false | id_locker2_obtainable,id_case0 => false | id_locker2_obtainable,id_request => false | id_locker2_obtainable,id_locker2_check_available => false | id_locker2_obtainable,id_not => false | id_locker2_obtainable,id_pushs => false | id_locker2_obtainable,id_locker2_promote_pending => false | id_locker2_obtainable,id_mcrlrecord => false | id_locker2_obtainable,id_locker2_obtainables => false | id_locker2_obtainable,id_cons => false | id_locker2_obtainable,id_push1 => false | id_locker2_obtainable,id_case1 => false | id_locker2_obtainable,id_element => false | id_locker2_obtainable,id_locker2_adduniq => false | id_locker2_obtainable,id_and => false | id_locker2_obtainable,id_empty => false | id_locker2_obtainable,id_record_updates => false | id_locker2_obtainable,id_lock => false | id_locker2_obtainable,id_excllock => false | id_locker2_obtainable,id_pid => false | id_locker2_obtainable,id_calls => false | id_locker2_obtainable,id_subtract => false | id_locker2_obtainable,id_tag => false | id_locker2_obtainable,id_equal => false | id_locker2_obtainable,id_eq => false | id_locker2_obtainable,id_tops => false | id_locker2_obtainable,id_locker2_claim_lock => false | id_locker2_obtainable,id_pending => false | id_locker2_obtainable,id_andt => false | id_locker2_obtainable,id_tuplenil => false | id_locker2_obtainable,id_append => false | id_locker2_obtainable,id_0 => false | id_locker2_obtainable,id_case8 => false | id_imp,id_or => true | id_imp,id_gen_tag => true | id_imp,id_record_new => true | id_imp,id_a => true | id_imp,id_locker2_release_lock => true | id_imp,id_eqt => true | id_imp,id_istops => true | id_imp,id_locker2_map_add_pending => true | id_imp,id_release => true | id_imp,id_locker2_obtainable => true | id_imp,id_imp => true | id_imp,id_stack => false | id_imp,id_locker2_map_promote_pending => false | id_imp,id_locker => false | id_imp,id_case4 => false | id_imp,id_int => false | id_imp,id_push => false | id_imp,id_locker2_add_pending => false | id_imp,id_true => false | id_imp,id_locker2_check_availables => false | id_imp,id_F => false | id_imp,id_eqs => false | id_imp,id_record_update => false | id_imp,id_false => false | id_imp,id_gen_modtageq => false | id_imp,id_undefined => false | id_imp,id_nocalls => false | id_imp,id_locker2_remove_pending => false | id_imp,id_resource => false | id_imp,id_case6 => false | id_imp,id_if => false | id_imp,id_pops => false | id_imp,id_locker2_map_claim_lock => false | id_imp,id_ok => false | id_imp,id_case5 => false | id_imp,id_tuple => false | id_imp,id_member => false | id_imp,id_s => false | id_imp,id_delete => false | id_imp,id_T => false | id_imp,id_case9 => false | id_imp,id_record_extract => false | id_imp,id_excl => false | id_imp,id_case2 => false | id_imp,id_nil => false | id_imp,id_eqc => false | id_imp,id_case0 => false | id_imp,id_request => false | id_imp,id_locker2_check_available => false | id_imp,id_not => false | id_imp,id_pushs => false | id_imp,id_locker2_promote_pending => false | id_imp,id_mcrlrecord => false | id_imp,id_locker2_obtainables => false | id_imp,id_cons => false | id_imp,id_push1 => false | id_imp,id_case1 => false | id_imp,id_element => false | id_imp,id_locker2_adduniq => false | id_imp,id_and => false | id_imp,id_empty => false | id_imp,id_record_updates => false | id_imp,id_lock => false | id_imp,id_excllock => false | id_imp,id_pid => false | id_imp,id_calls => false | id_imp,id_subtract => false | id_imp,id_tag => false | id_imp,id_equal => false | id_imp,id_eq => false | id_imp,id_tops => false | id_imp,id_locker2_claim_lock => false | id_imp,id_pending => false | id_imp,id_andt => false | id_imp,id_tuplenil => false | id_imp,id_append => false | id_imp,id_0 => false | id_imp,id_case8 => false | id_stack,id_or => true | id_stack,id_gen_tag => true | id_stack,id_record_new => true | id_stack,id_a => true | id_stack,id_locker2_release_lock => true | id_stack,id_eqt => true | id_stack,id_istops => true | id_stack,id_locker2_map_add_pending => true | id_stack,id_release => true | id_stack,id_locker2_obtainable => true | id_stack,id_imp => true | id_stack,id_stack => true | id_stack,id_locker2_map_promote_pending => false | id_stack,id_locker => false | id_stack,id_case4 => false | id_stack,id_int => false | id_stack,id_push => false | id_stack,id_locker2_add_pending => false | id_stack,id_true => false | id_stack,id_locker2_check_availables => false | id_stack,id_F => false | id_stack,id_eqs => false | id_stack,id_record_update => false | id_stack,id_false => false | id_stack,id_gen_modtageq => false | id_stack,id_undefined => false | id_stack,id_nocalls => false | id_stack,id_locker2_remove_pending => false | id_stack,id_resource => false | id_stack,id_case6 => false | id_stack,id_if => false | id_stack,id_pops => false | id_stack,id_locker2_map_claim_lock => false | id_stack,id_ok => false | id_stack,id_case5 => false | id_stack,id_tuple => false | id_stack,id_member => false | id_stack,id_s => false | id_stack,id_delete => false | id_stack,id_T => false | id_stack,id_case9 => false | id_stack,id_record_extract => false | id_stack,id_excl => false | id_stack,id_case2 => false | id_stack,id_nil => false | id_stack,id_eqc => false | id_stack,id_case0 => false | id_stack,id_request => false | id_stack,id_locker2_check_available => false | id_stack,id_not => false | id_stack,id_pushs => false | id_stack,id_locker2_promote_pending => false | id_stack,id_mcrlrecord => false | id_stack,id_locker2_obtainables => false | id_stack,id_cons => false | id_stack,id_push1 => false | id_stack,id_case1 => false | id_stack,id_element => false | id_stack,id_locker2_adduniq => false | id_stack,id_and => false | id_stack,id_empty => false | id_stack,id_record_updates => false | id_stack,id_lock => false | id_stack,id_excllock => false | id_stack,id_pid => false | id_stack,id_calls => false | id_stack,id_subtract => false | id_stack,id_tag => false | id_stack,id_equal => false | id_stack,id_eq => false | id_stack,id_tops => false | id_stack,id_locker2_claim_lock => false | id_stack,id_pending => false | id_stack,id_andt => false | id_stack,id_tuplenil => false | id_stack,id_append => false | id_stack,id_0 => false | id_stack,id_case8 => false | id_locker2_map_promote_pending,id_or => true | id_locker2_map_promote_pending,id_gen_tag => true | id_locker2_map_promote_pending,id_record_new => true | id_locker2_map_promote_pending,id_a => true | id_locker2_map_promote_pending,id_locker2_release_lock => true | id_locker2_map_promote_pending,id_eqt => true | id_locker2_map_promote_pending,id_istops => true | id_locker2_map_promote_pending,id_locker2_map_add_pending => true | id_locker2_map_promote_pending,id_release => true | id_locker2_map_promote_pending,id_locker2_obtainable => true | id_locker2_map_promote_pending,id_imp => true | id_locker2_map_promote_pending,id_stack => true | id_locker2_map_promote_pending,id_locker2_map_promote_pending => true | id_locker2_map_promote_pending,id_locker => false | id_locker2_map_promote_pending,id_case4 => false | id_locker2_map_promote_pending,id_int => false | id_locker2_map_promote_pending,id_push => false | id_locker2_map_promote_pending,id_locker2_add_pending => false | id_locker2_map_promote_pending,id_true => false | id_locker2_map_promote_pending,id_locker2_check_availables => false | id_locker2_map_promote_pending,id_F => false | id_locker2_map_promote_pending,id_eqs => false | id_locker2_map_promote_pending,id_record_update => false | id_locker2_map_promote_pending,id_false => false | id_locker2_map_promote_pending,id_gen_modtageq => false | id_locker2_map_promote_pending,id_undefined => false | id_locker2_map_promote_pending,id_nocalls => false | id_locker2_map_promote_pending,id_locker2_remove_pending => false | id_locker2_map_promote_pending,id_resource => false | id_locker2_map_promote_pending,id_case6 => false | id_locker2_map_promote_pending,id_if => false | id_locker2_map_promote_pending,id_pops => false | id_locker2_map_promote_pending,id_locker2_map_claim_lock => false | id_locker2_map_promote_pending,id_ok => false | id_locker2_map_promote_pending,id_case5 => false | id_locker2_map_promote_pending,id_tuple => false | id_locker2_map_promote_pending,id_member => false | id_locker2_map_promote_pending,id_s => false | id_locker2_map_promote_pending,id_delete => false | id_locker2_map_promote_pending,id_T => false | id_locker2_map_promote_pending,id_case9 => false | id_locker2_map_promote_pending,id_record_extract => false | id_locker2_map_promote_pending,id_excl => false | id_locker2_map_promote_pending,id_case2 => false | id_locker2_map_promote_pending,id_nil => false | id_locker2_map_promote_pending,id_eqc => false | id_locker2_map_promote_pending,id_case0 => false | id_locker2_map_promote_pending,id_request => false | id_locker2_map_promote_pending,id_locker2_check_available => false | id_locker2_map_promote_pending,id_not => false | id_locker2_map_promote_pending,id_pushs => false | id_locker2_map_promote_pending,id_locker2_promote_pending => false | id_locker2_map_promote_pending,id_mcrlrecord => false | id_locker2_map_promote_pending,id_locker2_obtainables => false | id_locker2_map_promote_pending,id_cons => false | id_locker2_map_promote_pending,id_push1 => false | id_locker2_map_promote_pending,id_case1 => false | id_locker2_map_promote_pending,id_element => false | id_locker2_map_promote_pending,id_locker2_adduniq => false | id_locker2_map_promote_pending,id_and => false | id_locker2_map_promote_pending,id_empty => false | id_locker2_map_promote_pending,id_record_updates => false | id_locker2_map_promote_pending,id_lock => false | id_locker2_map_promote_pending,id_excllock => false | id_locker2_map_promote_pending,id_pid => false | id_locker2_map_promote_pending,id_calls => false | id_locker2_map_promote_pending,id_subtract => false | id_locker2_map_promote_pending,id_tag => false | id_locker2_map_promote_pending,id_equal => false | id_locker2_map_promote_pending,id_eq => false | id_locker2_map_promote_pending,id_tops => false | id_locker2_map_promote_pending,id_locker2_claim_lock => false | id_locker2_map_promote_pending,id_pending => false | id_locker2_map_promote_pending,id_andt => false | id_locker2_map_promote_pending,id_tuplenil => false | id_locker2_map_promote_pending,id_append => false | id_locker2_map_promote_pending,id_0 => false | id_locker2_map_promote_pending,id_case8 => false | id_locker,id_or => true | id_locker,id_gen_tag => true | id_locker,id_record_new => true | id_locker,id_a => true | id_locker,id_locker2_release_lock => true | id_locker,id_eqt => true | id_locker,id_istops => true | id_locker,id_locker2_map_add_pending => true | id_locker,id_release => true | id_locker,id_locker2_obtainable => true | id_locker,id_imp => true | id_locker,id_stack => true | id_locker,id_locker2_map_promote_pending => true | id_locker,id_locker => true | id_locker,id_case4 => false | id_locker,id_int => false | id_locker,id_push => false | id_locker,id_locker2_add_pending => false | id_locker,id_true => false | id_locker,id_locker2_check_availables => false | id_locker,id_F => false | id_locker,id_eqs => false | id_locker,id_record_update => false | id_locker,id_false => false | id_locker,id_gen_modtageq => false | id_locker,id_undefined => false | id_locker,id_nocalls => false | id_locker,id_locker2_remove_pending => false | id_locker,id_resource => false | id_locker,id_case6 => false | id_locker,id_if => false | id_locker,id_pops => false | id_locker,id_locker2_map_claim_lock => false | id_locker,id_ok => false | id_locker,id_case5 => false | id_locker,id_tuple => false | id_locker,id_member => false | id_locker,id_s => false | id_locker,id_delete => false | id_locker,id_T => false | id_locker,id_case9 => false | id_locker,id_record_extract => false | id_locker,id_excl => false | id_locker,id_case2 => false | id_locker,id_nil => false | id_locker,id_eqc => false | id_locker,id_case0 => false | id_locker,id_request => false | id_locker,id_locker2_check_available => false | id_locker,id_not => false | id_locker,id_pushs => false | id_locker,id_locker2_promote_pending => false | id_locker,id_mcrlrecord => false | id_locker,id_locker2_obtainables => false | id_locker,id_cons => false | id_locker,id_push1 => false | id_locker,id_case1 => false | id_locker,id_element => false | id_locker,id_locker2_adduniq => false | id_locker,id_and => false | id_locker,id_empty => false | id_locker,id_record_updates => false | id_locker,id_lock => false | id_locker,id_excllock => false | id_locker,id_pid => false | id_locker,id_calls => false | id_locker,id_subtract => false | id_locker,id_tag => false | id_locker,id_equal => false | id_locker,id_eq => false | id_locker,id_tops => false | id_locker,id_locker2_claim_lock => false | id_locker,id_pending => false | id_locker,id_andt => false | id_locker,id_tuplenil => false | id_locker,id_append => false | id_locker,id_0 => false | id_locker,id_case8 => false | id_case4,id_or => true | id_case4,id_gen_tag => true | id_case4,id_record_new => true | id_case4,id_a => true | id_case4,id_locker2_release_lock => true | id_case4,id_eqt => true | id_case4,id_istops => true | id_case4,id_locker2_map_add_pending => true | id_case4,id_release => true | id_case4,id_locker2_obtainable => true | id_case4,id_imp => true | id_case4,id_stack => true | id_case4,id_locker2_map_promote_pending => true | id_case4,id_locker => true | id_case4,id_case4 => true | id_case4,id_int => false | id_case4,id_push => false | id_case4,id_locker2_add_pending => false | id_case4,id_true => false | id_case4,id_locker2_check_availables => false | id_case4,id_F => false | id_case4,id_eqs => false | id_case4,id_record_update => false | id_case4,id_false => false | id_case4,id_gen_modtageq => false | id_case4,id_undefined => false | id_case4,id_nocalls => false | id_case4,id_locker2_remove_pending => false | id_case4,id_resource => false | id_case4,id_case6 => false | id_case4,id_if => false | id_case4,id_pops => false | id_case4,id_locker2_map_claim_lock => false | id_case4,id_ok => false | id_case4,id_case5 => false | id_case4,id_tuple => false | id_case4,id_member => false | id_case4,id_s => false | id_case4,id_delete => false | id_case4,id_T => false | id_case4,id_case9 => false | id_case4,id_record_extract => false | id_case4,id_excl => false | id_case4,id_case2 => false | id_case4,id_nil => false | id_case4,id_eqc => false | id_case4,id_case0 => false | id_case4,id_request => false | id_case4,id_locker2_check_available => false | id_case4,id_not => false | id_case4,id_pushs => false | id_case4,id_locker2_promote_pending => false | id_case4,id_mcrlrecord => false | id_case4,id_locker2_obtainables => false | id_case4,id_cons => false | id_case4,id_push1 => false | id_case4,id_case1 => false | id_case4,id_element => false | id_case4,id_locker2_adduniq => false | id_case4,id_and => false | id_case4,id_empty => false | id_case4,id_record_updates => false | id_case4,id_lock => false | id_case4,id_excllock => false | id_case4,id_pid => false | id_case4,id_calls => false | id_case4,id_subtract => false | id_case4,id_tag => false | id_case4,id_equal => false | id_case4,id_eq => false | id_case4,id_tops => false | id_case4,id_locker2_claim_lock => false | id_case4,id_pending => false | id_case4,id_andt => false | id_case4,id_tuplenil => false | id_case4,id_append => false | id_case4,id_0 => false | id_case4,id_case8 => false | id_int,id_or => true | id_int,id_gen_tag => true | id_int,id_record_new => true | id_int,id_a => true | id_int,id_locker2_release_lock => true | id_int,id_eqt => true | id_int,id_istops => true | id_int,id_locker2_map_add_pending => true | id_int,id_release => true | id_int,id_locker2_obtainable => true | id_int,id_imp => true | id_int,id_stack => true | id_int,id_locker2_map_promote_pending => true | id_int,id_locker => true | id_int,id_case4 => true | id_int,id_int => true | id_int,id_push => false | id_int,id_locker2_add_pending => false | id_int,id_true => false | id_int,id_locker2_check_availables => false | id_int,id_F => false | id_int,id_eqs => false | id_int,id_record_update => false | id_int,id_false => false | id_int,id_gen_modtageq => false | id_int,id_undefined => false | id_int,id_nocalls => false | id_int,id_locker2_remove_pending => false | id_int,id_resource => false | id_int,id_case6 => false | id_int,id_if => false | id_int,id_pops => false | id_int,id_locker2_map_claim_lock => false | id_int,id_ok => false | id_int,id_case5 => false | id_int,id_tuple => false | id_int,id_member => false | id_int,id_s => false | id_int,id_delete => false | id_int,id_T => false | id_int,id_case9 => false | id_int,id_record_extract => false | id_int,id_excl => false | id_int,id_case2 => false | id_int,id_nil => false | id_int,id_eqc => false | id_int,id_case0 => false | id_int,id_request => false | id_int,id_locker2_check_available => false | id_int,id_not => false | id_int,id_pushs => false | id_int,id_locker2_promote_pending => false | id_int,id_mcrlrecord => false | id_int,id_locker2_obtainables => false | id_int,id_cons => false | id_int,id_push1 => false | id_int,id_case1 => false | id_int,id_element => false | id_int,id_locker2_adduniq => false | id_int,id_and => false | id_int,id_empty => false | id_int,id_record_updates => false | id_int,id_lock => false | id_int,id_excllock => false | id_int,id_pid => false | id_int,id_calls => false | id_int,id_subtract => false | id_int,id_tag => false | id_int,id_equal => false | id_int,id_eq => false | id_int,id_tops => false | id_int,id_locker2_claim_lock => false | id_int,id_pending => false | id_int,id_andt => false | id_int,id_tuplenil => false | id_int,id_append => false | id_int,id_0 => false | id_int,id_case8 => false | id_push,id_or => true | id_push,id_gen_tag => true | id_push,id_record_new => true | id_push,id_a => true | id_push,id_locker2_release_lock => true | id_push,id_eqt => true | id_push,id_istops => true | id_push,id_locker2_map_add_pending => true | id_push,id_release => true | id_push,id_locker2_obtainable => true | id_push,id_imp => true | id_push,id_stack => true | id_push,id_locker2_map_promote_pending => true | id_push,id_locker => true | id_push,id_case4 => true | id_push,id_int => true | id_push,id_push => true | id_push,id_locker2_add_pending => false | id_push,id_true => false | id_push,id_locker2_check_availables => false | id_push,id_F => false | id_push,id_eqs => false | id_push,id_record_update => false | id_push,id_false => false | id_push,id_gen_modtageq => false | id_push,id_undefined => false | id_push,id_nocalls => false | id_push,id_locker2_remove_pending => false | id_push,id_resource => false | id_push,id_case6 => false | id_push,id_if => false | id_push,id_pops => false | id_push,id_locker2_map_claim_lock => false | id_push,id_ok => false | id_push,id_case5 => false | id_push,id_tuple => false | id_push,id_member => false | id_push,id_s => false | id_push,id_delete => false | id_push,id_T => false | id_push,id_case9 => false | id_push,id_record_extract => false | id_push,id_excl => false | id_push,id_case2 => false | id_push,id_nil => false | id_push,id_eqc => false | id_push,id_case0 => false | id_push,id_request => false | id_push,id_locker2_check_available => false | id_push,id_not => false | id_push,id_pushs => false | id_push,id_locker2_promote_pending => false | id_push,id_mcrlrecord => false | id_push,id_locker2_obtainables => false | id_push,id_cons => false | id_push,id_push1 => false | id_push,id_case1 => false | id_push,id_element => false | id_push,id_locker2_adduniq => false | id_push,id_and => false | id_push,id_empty => false | id_push,id_record_updates => false | id_push,id_lock => false | id_push,id_excllock => false | id_push,id_pid => false | id_push,id_calls => false | id_push,id_subtract => false | id_push,id_tag => false | id_push,id_equal => false | id_push,id_eq => false | id_push,id_tops => false | id_push,id_locker2_claim_lock => false | id_push,id_pending => false | id_push,id_andt => false | id_push,id_tuplenil => false | id_push,id_append => false | id_push,id_0 => false | id_push,id_case8 => false | id_locker2_add_pending,id_or => true | id_locker2_add_pending,id_gen_tag => true | id_locker2_add_pending,id_record_new => true | id_locker2_add_pending,id_a => true | id_locker2_add_pending,id_locker2_release_lock => true | id_locker2_add_pending,id_eqt => true | id_locker2_add_pending,id_istops => true | id_locker2_add_pending,id_locker2_map_add_pending => true | id_locker2_add_pending,id_release => true | id_locker2_add_pending,id_locker2_obtainable => true | id_locker2_add_pending,id_imp => true | id_locker2_add_pending,id_stack => true | id_locker2_add_pending,id_locker2_map_promote_pending => true | id_locker2_add_pending,id_locker => true | id_locker2_add_pending,id_case4 => true | id_locker2_add_pending,id_int => true | id_locker2_add_pending,id_push => true | id_locker2_add_pending,id_locker2_add_pending => true | id_locker2_add_pending,id_true => false | id_locker2_add_pending,id_locker2_check_availables => false | id_locker2_add_pending,id_F => false | id_locker2_add_pending,id_eqs => false | id_locker2_add_pending,id_record_update => false | id_locker2_add_pending,id_false => false | id_locker2_add_pending,id_gen_modtageq => false | id_locker2_add_pending,id_undefined => false | id_locker2_add_pending,id_nocalls => false | id_locker2_add_pending,id_locker2_remove_pending => false | id_locker2_add_pending,id_resource => false | id_locker2_add_pending,id_case6 => false | id_locker2_add_pending,id_if => false | id_locker2_add_pending,id_pops => false | id_locker2_add_pending,id_locker2_map_claim_lock => false | id_locker2_add_pending,id_ok => false | id_locker2_add_pending,id_case5 => false | id_locker2_add_pending,id_tuple => false | id_locker2_add_pending,id_member => false | id_locker2_add_pending,id_s => false | id_locker2_add_pending,id_delete => false | id_locker2_add_pending,id_T => false | id_locker2_add_pending,id_case9 => false | id_locker2_add_pending,id_record_extract => false | id_locker2_add_pending,id_excl => false | id_locker2_add_pending,id_case2 => false | id_locker2_add_pending,id_nil => false | id_locker2_add_pending,id_eqc => false | id_locker2_add_pending,id_case0 => false | id_locker2_add_pending,id_request => false | id_locker2_add_pending,id_locker2_check_available => false | id_locker2_add_pending,id_not => false | id_locker2_add_pending,id_pushs => false | id_locker2_add_pending,id_locker2_promote_pending => false | id_locker2_add_pending,id_mcrlrecord => false | id_locker2_add_pending,id_locker2_obtainables => false | id_locker2_add_pending,id_cons => false | id_locker2_add_pending,id_push1 => false | id_locker2_add_pending,id_case1 => false | id_locker2_add_pending,id_element => false | id_locker2_add_pending,id_locker2_adduniq => false | id_locker2_add_pending,id_and => false | id_locker2_add_pending,id_empty => false | id_locker2_add_pending,id_record_updates => false | id_locker2_add_pending,id_lock => false | id_locker2_add_pending,id_excllock => false | id_locker2_add_pending,id_pid => false | id_locker2_add_pending,id_calls => false | id_locker2_add_pending,id_subtract => false | id_locker2_add_pending,id_tag => false | id_locker2_add_pending,id_equal => false | id_locker2_add_pending,id_eq => false | id_locker2_add_pending,id_tops => false | id_locker2_add_pending,id_locker2_claim_lock => false | id_locker2_add_pending,id_pending => false | id_locker2_add_pending,id_andt => false | id_locker2_add_pending,id_tuplenil => false | id_locker2_add_pending,id_append => false | id_locker2_add_pending,id_0 => false | id_locker2_add_pending,id_case8 => false | id_true,id_or => true | id_true,id_gen_tag => true | id_true,id_record_new => true | id_true,id_a => true | id_true,id_locker2_release_lock => true | id_true,id_eqt => true | id_true,id_istops => true | id_true,id_locker2_map_add_pending => true | id_true,id_release => true | id_true,id_locker2_obtainable => true | id_true,id_imp => true | id_true,id_stack => true | id_true,id_locker2_map_promote_pending => true | id_true,id_locker => true | id_true,id_case4 => true | id_true,id_int => true | id_true,id_push => true | id_true,id_locker2_add_pending => true | id_true,id_true => true | id_true,id_locker2_check_availables => false | id_true,id_F => false | id_true,id_eqs => false | id_true,id_record_update => false | id_true,id_false => false | id_true,id_gen_modtageq => false | id_true,id_undefined => false | id_true,id_nocalls => false | id_true,id_locker2_remove_pending => false | id_true,id_resource => false | id_true,id_case6 => false | id_true,id_if => false | id_true,id_pops => false | id_true,id_locker2_map_claim_lock => false | id_true,id_ok => false | id_true,id_case5 => false | id_true,id_tuple => false | id_true,id_member => false | id_true,id_s => false | id_true,id_delete => false | id_true,id_T => false | id_true,id_case9 => false | id_true,id_record_extract => false | id_true,id_excl => false | id_true,id_case2 => false | id_true,id_nil => false | id_true,id_eqc => false | id_true,id_case0 => false | id_true,id_request => false | id_true,id_locker2_check_available => false | id_true,id_not => false | id_true,id_pushs => false | id_true,id_locker2_promote_pending => false | id_true,id_mcrlrecord => false | id_true,id_locker2_obtainables => false | id_true,id_cons => false | id_true,id_push1 => false | id_true,id_case1 => false | id_true,id_element => false | id_true,id_locker2_adduniq => false | id_true,id_and => false | id_true,id_empty => false | id_true,id_record_updates => false | id_true,id_lock => false | id_true,id_excllock => false | id_true,id_pid => false | id_true,id_calls => false | id_true,id_subtract => false | id_true,id_tag => false | id_true,id_equal => false | id_true,id_eq => false | id_true,id_tops => false | id_true,id_locker2_claim_lock => false | id_true,id_pending => false | id_true,id_andt => false | id_true,id_tuplenil => false | id_true,id_append => false | id_true,id_0 => false | id_true,id_case8 => false | id_locker2_check_availables,id_or => true | id_locker2_check_availables,id_gen_tag => true | id_locker2_check_availables,id_record_new => true | id_locker2_check_availables,id_a => true | id_locker2_check_availables,id_locker2_release_lock => true | id_locker2_check_availables,id_eqt => true | id_locker2_check_availables,id_istops => true | id_locker2_check_availables,id_locker2_map_add_pending => true | id_locker2_check_availables,id_release => true | id_locker2_check_availables,id_locker2_obtainable => true | id_locker2_check_availables,id_imp => true | id_locker2_check_availables,id_stack => true | id_locker2_check_availables,id_locker2_map_promote_pending => true | id_locker2_check_availables,id_locker => true | id_locker2_check_availables,id_case4 => true | id_locker2_check_availables,id_int => true | id_locker2_check_availables,id_push => true | id_locker2_check_availables,id_locker2_add_pending => true | id_locker2_check_availables,id_true => true | id_locker2_check_availables,id_locker2_check_availables => true | id_locker2_check_availables,id_F => false | id_locker2_check_availables,id_eqs => false | id_locker2_check_availables,id_record_update => false | id_locker2_check_availables,id_false => false | id_locker2_check_availables,id_gen_modtageq => false | id_locker2_check_availables,id_undefined => false | id_locker2_check_availables,id_nocalls => false | id_locker2_check_availables,id_locker2_remove_pending => false | id_locker2_check_availables,id_resource => false | id_locker2_check_availables,id_case6 => false | id_locker2_check_availables,id_if => false | id_locker2_check_availables,id_pops => false | id_locker2_check_availables,id_locker2_map_claim_lock => false | id_locker2_check_availables,id_ok => false | id_locker2_check_availables,id_case5 => false | id_locker2_check_availables,id_tuple => false | id_locker2_check_availables,id_member => false | id_locker2_check_availables,id_s => false | id_locker2_check_availables,id_delete => false | id_locker2_check_availables,id_T => false | id_locker2_check_availables,id_case9 => false | id_locker2_check_availables,id_record_extract => false | id_locker2_check_availables,id_excl => false | id_locker2_check_availables,id_case2 => false | id_locker2_check_availables,id_nil => false | id_locker2_check_availables,id_eqc => false | id_locker2_check_availables,id_case0 => false | id_locker2_check_availables,id_request => false | id_locker2_check_availables,id_locker2_check_available => false | id_locker2_check_availables,id_not => false | id_locker2_check_availables,id_pushs => false | id_locker2_check_availables,id_locker2_promote_pending => false | id_locker2_check_availables,id_mcrlrecord => false | id_locker2_check_availables,id_locker2_obtainables => false | id_locker2_check_availables,id_cons => false | id_locker2_check_availables,id_push1 => false | id_locker2_check_availables,id_case1 => false | id_locker2_check_availables,id_element => false | id_locker2_check_availables,id_locker2_adduniq => false | id_locker2_check_availables,id_and => false | id_locker2_check_availables,id_empty => false | id_locker2_check_availables,id_record_updates => false | id_locker2_check_availables,id_lock => false | id_locker2_check_availables,id_excllock => false | id_locker2_check_availables,id_pid => false | id_locker2_check_availables,id_calls => false | id_locker2_check_availables,id_subtract => false | id_locker2_check_availables,id_tag => false | id_locker2_check_availables,id_equal => false | id_locker2_check_availables,id_eq => false | id_locker2_check_availables,id_tops => false | id_locker2_check_availables,id_locker2_claim_lock => false | id_locker2_check_availables,id_pending => false | id_locker2_check_availables,id_andt => false | id_locker2_check_availables,id_tuplenil => false | id_locker2_check_availables,id_append => false | id_locker2_check_availables,id_0 => false | id_locker2_check_availables,id_case8 => false | id_F,id_or => true | id_F,id_gen_tag => true | id_F,id_record_new => true | id_F,id_a => true | id_F,id_locker2_release_lock => true | id_F,id_eqt => true | id_F,id_istops => true | id_F,id_locker2_map_add_pending => true | id_F,id_release => true | id_F,id_locker2_obtainable => true | id_F,id_imp => true | id_F,id_stack => true | id_F,id_locker2_map_promote_pending => true | id_F,id_locker => true | id_F,id_case4 => true | id_F,id_int => true | id_F,id_push => true | id_F,id_locker2_add_pending => true | id_F,id_true => true | id_F,id_locker2_check_availables => true | id_F,id_F => true | id_F,id_eqs => false | id_F,id_record_update => false | id_F,id_false => false | id_F,id_gen_modtageq => false | id_F,id_undefined => false | id_F,id_nocalls => false | id_F,id_locker2_remove_pending => false | id_F,id_resource => false | id_F,id_case6 => false | id_F,id_if => false | id_F,id_pops => false | id_F,id_locker2_map_claim_lock => false | id_F,id_ok => false | id_F,id_case5 => false | id_F,id_tuple => false | id_F,id_member => false | id_F,id_s => false | id_F,id_delete => false | id_F,id_T => false | id_F,id_case9 => false | id_F,id_record_extract => false | id_F,id_excl => false | id_F,id_case2 => false | id_F,id_nil => false | id_F,id_eqc => false | id_F,id_case0 => false | id_F,id_request => false | id_F,id_locker2_check_available => false | id_F,id_not => false | id_F,id_pushs => false | id_F,id_locker2_promote_pending => false | id_F,id_mcrlrecord => false | id_F,id_locker2_obtainables => false | id_F,id_cons => false | id_F,id_push1 => false | id_F,id_case1 => false | id_F,id_element => false | id_F,id_locker2_adduniq => false | id_F,id_and => false | id_F,id_empty => false | id_F,id_record_updates => false | id_F,id_lock => false | id_F,id_excllock => false | id_F,id_pid => false | id_F,id_calls => false | id_F,id_subtract => false | id_F,id_tag => false | id_F,id_equal => false | id_F,id_eq => false | id_F,id_tops => false | id_F,id_locker2_claim_lock => false | id_F,id_pending => false | id_F,id_andt => false | id_F,id_tuplenil => false | id_F,id_append => false | id_F,id_0 => false | id_F,id_case8 => false | id_eqs,id_or => true | id_eqs,id_gen_tag => true | id_eqs,id_record_new => true | id_eqs,id_a => true | id_eqs,id_locker2_release_lock => true | id_eqs,id_eqt => true | id_eqs,id_istops => true | id_eqs,id_locker2_map_add_pending => true | id_eqs,id_release => true | id_eqs,id_locker2_obtainable => true | id_eqs,id_imp => true | id_eqs,id_stack => true | id_eqs,id_locker2_map_promote_pending => true | id_eqs,id_locker => true | id_eqs,id_case4 => true | id_eqs,id_int => true | id_eqs,id_push => true | id_eqs,id_locker2_add_pending => true | id_eqs,id_true => true | id_eqs,id_locker2_check_availables => true | id_eqs,id_F => true | id_eqs,id_eqs => true | id_eqs,id_record_update => false | id_eqs,id_false => false | id_eqs,id_gen_modtageq => false | id_eqs,id_undefined => false | id_eqs,id_nocalls => false | id_eqs,id_locker2_remove_pending => false | id_eqs,id_resource => false | id_eqs,id_case6 => false | id_eqs,id_if => false | id_eqs,id_pops => false | id_eqs,id_locker2_map_claim_lock => false | id_eqs,id_ok => false | id_eqs,id_case5 => false | id_eqs,id_tuple => false | id_eqs,id_member => false | id_eqs,id_s => false | id_eqs,id_delete => false | id_eqs,id_T => false | id_eqs,id_case9 => false | id_eqs,id_record_extract => false | id_eqs,id_excl => false | id_eqs,id_case2 => false | id_eqs,id_nil => false | id_eqs,id_eqc => false | id_eqs,id_case0 => false | id_eqs,id_request => false | id_eqs,id_locker2_check_available => false | id_eqs,id_not => false | id_eqs,id_pushs => false | id_eqs,id_locker2_promote_pending => false | id_eqs,id_mcrlrecord => false | id_eqs,id_locker2_obtainables => false | id_eqs,id_cons => false | id_eqs,id_push1 => false | id_eqs,id_case1 => false | id_eqs,id_element => false | id_eqs,id_locker2_adduniq => false | id_eqs,id_and => false | id_eqs,id_empty => false | id_eqs,id_record_updates => false | id_eqs,id_lock => false | id_eqs,id_excllock => false | id_eqs,id_pid => false | id_eqs,id_calls => false | id_eqs,id_subtract => false | id_eqs,id_tag => false | id_eqs,id_equal => false | id_eqs,id_eq => false | id_eqs,id_tops => false | id_eqs,id_locker2_claim_lock => false | id_eqs,id_pending => false | id_eqs,id_andt => false | id_eqs,id_tuplenil => false | id_eqs,id_append => false | id_eqs,id_0 => false | id_eqs,id_case8 => false | id_record_update,id_or => true | id_record_update,id_gen_tag => true | id_record_update,id_record_new => true | id_record_update,id_a => true | id_record_update,id_locker2_release_lock => true | id_record_update,id_eqt => true | id_record_update,id_istops => true | id_record_update,id_locker2_map_add_pending => true | id_record_update,id_release => true | id_record_update,id_locker2_obtainable => true | id_record_update,id_imp => true | id_record_update,id_stack => true | id_record_update,id_locker2_map_promote_pending => true | id_record_update,id_locker => true | id_record_update,id_case4 => true | id_record_update,id_int => true | id_record_update,id_push => true | id_record_update,id_locker2_add_pending => true | id_record_update,id_true => true | id_record_update,id_locker2_check_availables => true | id_record_update,id_F => true | id_record_update,id_eqs => true | id_record_update,id_record_update => true | id_record_update,id_false => false | id_record_update,id_gen_modtageq => false | id_record_update,id_undefined => false | id_record_update,id_nocalls => false | id_record_update,id_locker2_remove_pending => false | id_record_update,id_resource => false | id_record_update,id_case6 => false | id_record_update,id_if => false | id_record_update,id_pops => false | id_record_update,id_locker2_map_claim_lock => false | id_record_update,id_ok => false | id_record_update,id_case5 => false | id_record_update,id_tuple => false | id_record_update,id_member => false | id_record_update,id_s => false | id_record_update,id_delete => false | id_record_update,id_T => false | id_record_update,id_case9 => false | id_record_update,id_record_extract => false | id_record_update,id_excl => false | id_record_update,id_case2 => false | id_record_update,id_nil => false | id_record_update,id_eqc => false | id_record_update,id_case0 => false | id_record_update,id_request => false | id_record_update,id_locker2_check_available => false | id_record_update,id_not => false | id_record_update,id_pushs => false | id_record_update,id_locker2_promote_pending => false | id_record_update,id_mcrlrecord => false | id_record_update,id_locker2_obtainables => false | id_record_update,id_cons => false | id_record_update,id_push1 => false | id_record_update,id_case1 => false | id_record_update,id_element => false | id_record_update,id_locker2_adduniq => false | id_record_update,id_and => false | id_record_update,id_empty => false | id_record_update,id_record_updates => false | id_record_update,id_lock => false | id_record_update,id_excllock => false | id_record_update,id_pid => false | id_record_update,id_calls => false | id_record_update,id_subtract => false | id_record_update,id_tag => false | id_record_update,id_equal => false | id_record_update,id_eq => false | id_record_update,id_tops => false | id_record_update,id_locker2_claim_lock => false | id_record_update,id_pending => false | id_record_update,id_andt => false | id_record_update,id_tuplenil => false | id_record_update,id_append => false | id_record_update,id_0 => false | id_record_update,id_case8 => false | id_false,id_or => true | id_false,id_gen_tag => true | id_false,id_record_new => true | id_false,id_a => true | id_false,id_locker2_release_lock => true | id_false,id_eqt => true | id_false,id_istops => true | id_false,id_locker2_map_add_pending => true | id_false,id_release => true | id_false,id_locker2_obtainable => true | id_false,id_imp => true | id_false,id_stack => true | id_false,id_locker2_map_promote_pending => true | id_false,id_locker => true | id_false,id_case4 => true | id_false,id_int => true | id_false,id_push => true | id_false,id_locker2_add_pending => true | id_false,id_true => true | id_false,id_locker2_check_availables => true | id_false,id_F => true | id_false,id_eqs => true | id_false,id_record_update => true | id_false,id_false => true | id_false,id_gen_modtageq => false | id_false,id_undefined => false | id_false,id_nocalls => false | id_false,id_locker2_remove_pending => false | id_false,id_resource => false | id_false,id_case6 => false | id_false,id_if => false | id_false,id_pops => false | id_false,id_locker2_map_claim_lock => false | id_false,id_ok => false | id_false,id_case5 => false | id_false,id_tuple => false | id_false,id_member => false | id_false,id_s => false | id_false,id_delete => false | id_false,id_T => false | id_false,id_case9 => false | id_false,id_record_extract => false | id_false,id_excl => false | id_false,id_case2 => false | id_false,id_nil => false | id_false,id_eqc => false | id_false,id_case0 => false | id_false,id_request => false | id_false,id_locker2_check_available => false | id_false,id_not => false | id_false,id_pushs => false | id_false,id_locker2_promote_pending => false | id_false,id_mcrlrecord => false | id_false,id_locker2_obtainables => false | id_false,id_cons => false | id_false,id_push1 => false | id_false,id_case1 => false | id_false,id_element => false | id_false,id_locker2_adduniq => false | id_false,id_and => false | id_false,id_empty => false | id_false,id_record_updates => false | id_false,id_lock => false | id_false,id_excllock => false | id_false,id_pid => false | id_false,id_calls => false | id_false,id_subtract => false | id_false,id_tag => false | id_false,id_equal => false | id_false,id_eq => false | id_false,id_tops => false | id_false,id_locker2_claim_lock => false | id_false,id_pending => false | id_false,id_andt => false | id_false,id_tuplenil => false | id_false,id_append => false | id_false,id_0 => false | id_false,id_case8 => false | id_gen_modtageq,id_or => true | id_gen_modtageq,id_gen_tag => true | id_gen_modtageq,id_record_new => true | id_gen_modtageq,id_a => true | id_gen_modtageq,id_locker2_release_lock => true | id_gen_modtageq,id_eqt => true | id_gen_modtageq,id_istops => true | id_gen_modtageq,id_locker2_map_add_pending => true | id_gen_modtageq,id_release => true | id_gen_modtageq,id_locker2_obtainable => true | id_gen_modtageq,id_imp => true | id_gen_modtageq,id_stack => true | id_gen_modtageq,id_locker2_map_promote_pending => true | id_gen_modtageq,id_locker => true | id_gen_modtageq,id_case4 => true | id_gen_modtageq,id_int => true | id_gen_modtageq,id_push => true | id_gen_modtageq,id_locker2_add_pending => true | id_gen_modtageq,id_true => true | id_gen_modtageq,id_locker2_check_availables => true | id_gen_modtageq,id_F => true | id_gen_modtageq,id_eqs => true | id_gen_modtageq,id_record_update => true | id_gen_modtageq,id_false => true | id_gen_modtageq,id_gen_modtageq => true | id_gen_modtageq,id_undefined => false | id_gen_modtageq,id_nocalls => false | id_gen_modtageq,id_locker2_remove_pending => false | id_gen_modtageq,id_resource => false | id_gen_modtageq,id_case6 => false | id_gen_modtageq,id_if => false | id_gen_modtageq,id_pops => false | id_gen_modtageq,id_locker2_map_claim_lock => false | id_gen_modtageq,id_ok => false | id_gen_modtageq,id_case5 => false | id_gen_modtageq,id_tuple => false | id_gen_modtageq,id_member => false | id_gen_modtageq,id_s => false | id_gen_modtageq,id_delete => false | id_gen_modtageq,id_T => false | id_gen_modtageq,id_case9 => false | id_gen_modtageq,id_record_extract => false | id_gen_modtageq,id_excl => false | id_gen_modtageq,id_case2 => false | id_gen_modtageq,id_nil => false | id_gen_modtageq,id_eqc => false | id_gen_modtageq,id_case0 => false | id_gen_modtageq,id_request => false | id_gen_modtageq,id_locker2_check_available => false | id_gen_modtageq,id_not => false | id_gen_modtageq,id_pushs => false | id_gen_modtageq,id_locker2_promote_pending => false | id_gen_modtageq,id_mcrlrecord => false | id_gen_modtageq,id_locker2_obtainables => false | id_gen_modtageq,id_cons => false | id_gen_modtageq,id_push1 => false | id_gen_modtageq,id_case1 => false | id_gen_modtageq,id_element => false | id_gen_modtageq,id_locker2_adduniq => false | id_gen_modtageq,id_and => false | id_gen_modtageq,id_empty => false | id_gen_modtageq,id_record_updates => false | id_gen_modtageq,id_lock => false | id_gen_modtageq,id_excllock => false | id_gen_modtageq,id_pid => false | id_gen_modtageq,id_calls => false | id_gen_modtageq,id_subtract => false | id_gen_modtageq,id_tag => false | id_gen_modtageq,id_equal => false | id_gen_modtageq,id_eq => false | id_gen_modtageq,id_tops => false | id_gen_modtageq,id_locker2_claim_lock => false | id_gen_modtageq,id_pending => false | id_gen_modtageq,id_andt => false | id_gen_modtageq,id_tuplenil => false | id_gen_modtageq,id_append => false | id_gen_modtageq,id_0 => false | id_gen_modtageq,id_case8 => false | id_undefined,id_or => true | id_undefined,id_gen_tag => true | id_undefined,id_record_new => true | id_undefined,id_a => true | id_undefined,id_locker2_release_lock => true | id_undefined,id_eqt => true | id_undefined,id_istops => true | id_undefined,id_locker2_map_add_pending => true | id_undefined,id_release => true | id_undefined,id_locker2_obtainable => true | id_undefined,id_imp => true | id_undefined,id_stack => true | id_undefined,id_locker2_map_promote_pending => true | id_undefined,id_locker => true | id_undefined,id_case4 => true | id_undefined,id_int => true | id_undefined,id_push => true | id_undefined,id_locker2_add_pending => true | id_undefined,id_true => true | id_undefined,id_locker2_check_availables => true | id_undefined,id_F => true | id_undefined,id_eqs => true | id_undefined,id_record_update => true | id_undefined,id_false => true | id_undefined,id_gen_modtageq => true | id_undefined,id_undefined => true | id_undefined,id_nocalls => false | id_undefined,id_locker2_remove_pending => false | id_undefined,id_resource => false | id_undefined,id_case6 => false | id_undefined,id_if => false | id_undefined,id_pops => false | id_undefined,id_locker2_map_claim_lock => false | id_undefined,id_ok => false | id_undefined,id_case5 => false | id_undefined,id_tuple => false | id_undefined,id_member => false | id_undefined,id_s => false | id_undefined,id_delete => false | id_undefined,id_T => false | id_undefined,id_case9 => false | id_undefined,id_record_extract => false | id_undefined,id_excl => false | id_undefined,id_case2 => false | id_undefined,id_nil => false | id_undefined,id_eqc => false | id_undefined,id_case0 => false | id_undefined,id_request => false | id_undefined,id_locker2_check_available => false | id_undefined,id_not => false | id_undefined,id_pushs => false | id_undefined,id_locker2_promote_pending => false | id_undefined,id_mcrlrecord => false | id_undefined,id_locker2_obtainables => false | id_undefined,id_cons => false | id_undefined,id_push1 => false | id_undefined,id_case1 => false | id_undefined,id_element => false | id_undefined,id_locker2_adduniq => false | id_undefined,id_and => false | id_undefined,id_empty => false | id_undefined,id_record_updates => false | id_undefined,id_lock => false | id_undefined,id_excllock => false | id_undefined,id_pid => false | id_undefined,id_calls => false | id_undefined,id_subtract => false | id_undefined,id_tag => false | id_undefined,id_equal => false | id_undefined,id_eq => false | id_undefined,id_tops => false | id_undefined,id_locker2_claim_lock => false | id_undefined,id_pending => false | id_undefined,id_andt => false | id_undefined,id_tuplenil => false | id_undefined,id_append => false | id_undefined,id_0 => false | id_undefined,id_case8 => false | id_nocalls,id_or => true | id_nocalls,id_gen_tag => true | id_nocalls,id_record_new => true | id_nocalls,id_a => true | id_nocalls,id_locker2_release_lock => true | id_nocalls,id_eqt => true | id_nocalls,id_istops => true | id_nocalls,id_locker2_map_add_pending => true | id_nocalls,id_release => true | id_nocalls,id_locker2_obtainable => true | id_nocalls,id_imp => true | id_nocalls,id_stack => true | id_nocalls,id_locker2_map_promote_pending => true | id_nocalls,id_locker => true | id_nocalls,id_case4 => true | id_nocalls,id_int => true | id_nocalls,id_push => true | id_nocalls,id_locker2_add_pending => true | id_nocalls,id_true => true | id_nocalls,id_locker2_check_availables => true | id_nocalls,id_F => true | id_nocalls,id_eqs => true | id_nocalls,id_record_update => true | id_nocalls,id_false => true | id_nocalls,id_gen_modtageq => true | id_nocalls,id_undefined => true | id_nocalls,id_nocalls => true | id_nocalls,id_locker2_remove_pending => false | id_nocalls,id_resource => false | id_nocalls,id_case6 => false | id_nocalls,id_if => false | id_nocalls,id_pops => false | id_nocalls,id_locker2_map_claim_lock => false | id_nocalls,id_ok => false | id_nocalls,id_case5 => false | id_nocalls,id_tuple => false | id_nocalls,id_member => false | id_nocalls,id_s => false | id_nocalls,id_delete => false | id_nocalls,id_T => false | id_nocalls,id_case9 => false | id_nocalls,id_record_extract => false | id_nocalls,id_excl => false | id_nocalls,id_case2 => false | id_nocalls,id_nil => false | id_nocalls,id_eqc => false | id_nocalls,id_case0 => false | id_nocalls,id_request => false | id_nocalls,id_locker2_check_available => false | id_nocalls,id_not => false | id_nocalls,id_pushs => false | id_nocalls,id_locker2_promote_pending => false | id_nocalls,id_mcrlrecord => false | id_nocalls,id_locker2_obtainables => false | id_nocalls,id_cons => false | id_nocalls,id_push1 => false | id_nocalls,id_case1 => false | id_nocalls,id_element => false | id_nocalls,id_locker2_adduniq => false | id_nocalls,id_and => false | id_nocalls,id_empty => false | id_nocalls,id_record_updates => false | id_nocalls,id_lock => false | id_nocalls,id_excllock => false | id_nocalls,id_pid => false | id_nocalls,id_calls => false | id_nocalls,id_subtract => false | id_nocalls,id_tag => false | id_nocalls,id_equal => false | id_nocalls,id_eq => false | id_nocalls,id_tops => false | id_nocalls,id_locker2_claim_lock => false | id_nocalls,id_pending => false | id_nocalls,id_andt => false | id_nocalls,id_tuplenil => false | id_nocalls,id_append => false | id_nocalls,id_0 => false | id_nocalls,id_case8 => false | id_locker2_remove_pending,id_or => true | id_locker2_remove_pending,id_gen_tag => true | id_locker2_remove_pending,id_record_new => true | id_locker2_remove_pending,id_a => true | id_locker2_remove_pending,id_locker2_release_lock => true | id_locker2_remove_pending,id_eqt => true | id_locker2_remove_pending,id_istops => true | id_locker2_remove_pending,id_locker2_map_add_pending => true | id_locker2_remove_pending,id_release => true | id_locker2_remove_pending,id_locker2_obtainable => true | id_locker2_remove_pending,id_imp => true | id_locker2_remove_pending,id_stack => true | id_locker2_remove_pending,id_locker2_map_promote_pending => true | id_locker2_remove_pending,id_locker => true | id_locker2_remove_pending,id_case4 => true | id_locker2_remove_pending,id_int => true | id_locker2_remove_pending,id_push => true | id_locker2_remove_pending,id_locker2_add_pending => true | id_locker2_remove_pending,id_true => true | id_locker2_remove_pending,id_locker2_check_availables => true | id_locker2_remove_pending,id_F => true | id_locker2_remove_pending,id_eqs => true | id_locker2_remove_pending,id_record_update => true | id_locker2_remove_pending,id_false => true | id_locker2_remove_pending,id_gen_modtageq => true | id_locker2_remove_pending,id_undefined => true | id_locker2_remove_pending,id_nocalls => true | id_locker2_remove_pending,id_locker2_remove_pending => true | id_locker2_remove_pending,id_resource => false | id_locker2_remove_pending,id_case6 => false | id_locker2_remove_pending,id_if => false | id_locker2_remove_pending,id_pops => false | id_locker2_remove_pending,id_locker2_map_claim_lock => false | id_locker2_remove_pending,id_ok => false | id_locker2_remove_pending,id_case5 => false | id_locker2_remove_pending,id_tuple => false | id_locker2_remove_pending,id_member => false | id_locker2_remove_pending,id_s => false | id_locker2_remove_pending,id_delete => false | id_locker2_remove_pending,id_T => false | id_locker2_remove_pending,id_case9 => false | id_locker2_remove_pending,id_record_extract => false | id_locker2_remove_pending,id_excl => false | id_locker2_remove_pending,id_case2 => false | id_locker2_remove_pending,id_nil => false | id_locker2_remove_pending,id_eqc => false | id_locker2_remove_pending,id_case0 => false | id_locker2_remove_pending,id_request => false | id_locker2_remove_pending,id_locker2_check_available => false | id_locker2_remove_pending,id_not => false | id_locker2_remove_pending,id_pushs => false | id_locker2_remove_pending,id_locker2_promote_pending => false | id_locker2_remove_pending,id_mcrlrecord => false | id_locker2_remove_pending,id_locker2_obtainables => false | id_locker2_remove_pending,id_cons => false | id_locker2_remove_pending,id_push1 => false | id_locker2_remove_pending,id_case1 => false | id_locker2_remove_pending,id_element => false | id_locker2_remove_pending,id_locker2_adduniq => false | id_locker2_remove_pending,id_and => false | id_locker2_remove_pending,id_empty => false | id_locker2_remove_pending,id_record_updates => false | id_locker2_remove_pending,id_lock => false | id_locker2_remove_pending,id_excllock => false | id_locker2_remove_pending,id_pid => false | id_locker2_remove_pending,id_calls => false | id_locker2_remove_pending,id_subtract => false | id_locker2_remove_pending,id_tag => false | id_locker2_remove_pending,id_equal => false | id_locker2_remove_pending,id_eq => false | id_locker2_remove_pending,id_tops => false | id_locker2_remove_pending,id_locker2_claim_lock => false | id_locker2_remove_pending,id_pending => false | id_locker2_remove_pending,id_andt => false | id_locker2_remove_pending,id_tuplenil => false | id_locker2_remove_pending,id_append => false | id_locker2_remove_pending,id_0 => false | id_locker2_remove_pending,id_case8 => false | id_resource,id_or => true | id_resource,id_gen_tag => true | id_resource,id_record_new => true | id_resource,id_a => true | id_resource,id_locker2_release_lock => true | id_resource,id_eqt => true | id_resource,id_istops => true | id_resource,id_locker2_map_add_pending => true | id_resource,id_release => true | id_resource,id_locker2_obtainable => true | id_resource,id_imp => true | id_resource,id_stack => true | id_resource,id_locker2_map_promote_pending => true | id_resource,id_locker => true | id_resource,id_case4 => true | id_resource,id_int => true | id_resource,id_push => true | id_resource,id_locker2_add_pending => true | id_resource,id_true => true | id_resource,id_locker2_check_availables => true | id_resource,id_F => true | id_resource,id_eqs => true | id_resource,id_record_update => true | id_resource,id_false => true | id_resource,id_gen_modtageq => true | id_resource,id_undefined => true | id_resource,id_nocalls => true | id_resource,id_locker2_remove_pending => true | id_resource,id_resource => true | id_resource,id_case6 => false | id_resource,id_if => false | id_resource,id_pops => false | id_resource,id_locker2_map_claim_lock => false | id_resource,id_ok => false | id_resource,id_case5 => false | id_resource,id_tuple => false | id_resource,id_member => false | id_resource,id_s => false | id_resource,id_delete => false | id_resource,id_T => false | id_resource,id_case9 => false | id_resource,id_record_extract => false | id_resource,id_excl => false | id_resource,id_case2 => false | id_resource,id_nil => false | id_resource,id_eqc => false | id_resource,id_case0 => false | id_resource,id_request => false | id_resource,id_locker2_check_available => false | id_resource,id_not => false | id_resource,id_pushs => false | id_resource,id_locker2_promote_pending => false | id_resource,id_mcrlrecord => false | id_resource,id_locker2_obtainables => false | id_resource,id_cons => false | id_resource,id_push1 => false | id_resource,id_case1 => false | id_resource,id_element => false | id_resource,id_locker2_adduniq => false | id_resource,id_and => false | id_resource,id_empty => false | id_resource,id_record_updates => false | id_resource,id_lock => false | id_resource,id_excllock => false | id_resource,id_pid => false | id_resource,id_calls => false | id_resource,id_subtract => false | id_resource,id_tag => false | id_resource,id_equal => false | id_resource,id_eq => false | id_resource,id_tops => false | id_resource,id_locker2_claim_lock => false | id_resource,id_pending => false | id_resource,id_andt => false | id_resource,id_tuplenil => false | id_resource,id_append => false | id_resource,id_0 => false | id_resource,id_case8 => false | id_case6,id_or => true | id_case6,id_gen_tag => true | id_case6,id_record_new => true | id_case6,id_a => true | id_case6,id_locker2_release_lock => true | id_case6,id_eqt => true | id_case6,id_istops => true | id_case6,id_locker2_map_add_pending => true | id_case6,id_release => true | id_case6,id_locker2_obtainable => true | id_case6,id_imp => true | id_case6,id_stack => true | id_case6,id_locker2_map_promote_pending => true | id_case6,id_locker => true | id_case6,id_case4 => true | id_case6,id_int => true | id_case6,id_push => true | id_case6,id_locker2_add_pending => true | id_case6,id_true => true | id_case6,id_locker2_check_availables => true | id_case6,id_F => true | id_case6,id_eqs => true | id_case6,id_record_update => true | id_case6,id_false => true | id_case6,id_gen_modtageq => true | id_case6,id_undefined => true | id_case6,id_nocalls => true | id_case6,id_locker2_remove_pending => true | id_case6,id_resource => true | id_case6,id_case6 => true | id_case6,id_if => false | id_case6,id_pops => false | id_case6,id_locker2_map_claim_lock => false | id_case6,id_ok => false | id_case6,id_case5 => false | id_case6,id_tuple => false | id_case6,id_member => false | id_case6,id_s => false | id_case6,id_delete => false | id_case6,id_T => false | id_case6,id_case9 => false | id_case6,id_record_extract => false | id_case6,id_excl => false | id_case6,id_case2 => false | id_case6,id_nil => false | id_case6,id_eqc => false | id_case6,id_case0 => false | id_case6,id_request => false | id_case6,id_locker2_check_available => false | id_case6,id_not => false | id_case6,id_pushs => false | id_case6,id_locker2_promote_pending => false | id_case6,id_mcrlrecord => false | id_case6,id_locker2_obtainables => false | id_case6,id_cons => false | id_case6,id_push1 => false | id_case6,id_case1 => false | id_case6,id_element => false | id_case6,id_locker2_adduniq => false | id_case6,id_and => false | id_case6,id_empty => false | id_case6,id_record_updates => false | id_case6,id_lock => false | id_case6,id_excllock => false | id_case6,id_pid => false | id_case6,id_calls => false | id_case6,id_subtract => false | id_case6,id_tag => false | id_case6,id_equal => false | id_case6,id_eq => false | id_case6,id_tops => false | id_case6,id_locker2_claim_lock => false | id_case6,id_pending => false | id_case6,id_andt => false | id_case6,id_tuplenil => false | id_case6,id_append => false | id_case6,id_0 => false | id_case6,id_case8 => false | id_if,id_or => true | id_if,id_gen_tag => true | id_if,id_record_new => true | id_if,id_a => true | id_if,id_locker2_release_lock => true | id_if,id_eqt => true | id_if,id_istops => true | id_if,id_locker2_map_add_pending => true | id_if,id_release => true | id_if,id_locker2_obtainable => true | id_if,id_imp => true | id_if,id_stack => true | id_if,id_locker2_map_promote_pending => true | id_if,id_locker => true | id_if,id_case4 => true | id_if,id_int => true | id_if,id_push => true | id_if,id_locker2_add_pending => true | id_if,id_true => true | id_if,id_locker2_check_availables => true | id_if,id_F => true | id_if,id_eqs => true | id_if,id_record_update => true | id_if,id_false => true | id_if,id_gen_modtageq => true | id_if,id_undefined => true | id_if,id_nocalls => true | id_if,id_locker2_remove_pending => true | id_if,id_resource => true | id_if,id_case6 => true | id_if,id_if => true | id_if,id_pops => false | id_if,id_locker2_map_claim_lock => false | id_if,id_ok => false | id_if,id_case5 => false | id_if,id_tuple => false | id_if,id_member => false | id_if,id_s => false | id_if,id_delete => false | id_if,id_T => false | id_if,id_case9 => false | id_if,id_record_extract => false | id_if,id_excl => false | id_if,id_case2 => false | id_if,id_nil => false | id_if,id_eqc => false | id_if,id_case0 => false | id_if,id_request => false | id_if,id_locker2_check_available => false | id_if,id_not => false | id_if,id_pushs => false | id_if,id_locker2_promote_pending => false | id_if,id_mcrlrecord => false | id_if,id_locker2_obtainables => false | id_if,id_cons => false | id_if,id_push1 => false | id_if,id_case1 => false | id_if,id_element => false | id_if,id_locker2_adduniq => false | id_if,id_and => false | id_if,id_empty => false | id_if,id_record_updates => false | id_if,id_lock => false | id_if,id_excllock => false | id_if,id_pid => false | id_if,id_calls => false | id_if,id_subtract => false | id_if,id_tag => false | id_if,id_equal => false | id_if,id_eq => false | id_if,id_tops => false | id_if,id_locker2_claim_lock => false | id_if,id_pending => false | id_if,id_andt => false | id_if,id_tuplenil => false | id_if,id_append => false | id_if,id_0 => false | id_if,id_case8 => false | id_pops,id_or => true | id_pops,id_gen_tag => true | id_pops,id_record_new => true | id_pops,id_a => true | id_pops,id_locker2_release_lock => true | id_pops,id_eqt => true | id_pops,id_istops => true | id_pops,id_locker2_map_add_pending => true | id_pops,id_release => true | id_pops,id_locker2_obtainable => true | id_pops,id_imp => true | id_pops,id_stack => true | id_pops,id_locker2_map_promote_pending => true | id_pops,id_locker => true | id_pops,id_case4 => true | id_pops,id_int => true | id_pops,id_push => true | id_pops,id_locker2_add_pending => true | id_pops,id_true => true | id_pops,id_locker2_check_availables => true | id_pops,id_F => true | id_pops,id_eqs => true | id_pops,id_record_update => true | id_pops,id_false => true | id_pops,id_gen_modtageq => true | id_pops,id_undefined => true | id_pops,id_nocalls => true | id_pops,id_locker2_remove_pending => true | id_pops,id_resource => true | id_pops,id_case6 => true | id_pops,id_if => true | id_pops,id_pops => true | id_pops,id_locker2_map_claim_lock => false | id_pops,id_ok => false | id_pops,id_case5 => false | id_pops,id_tuple => false | id_pops,id_member => false | id_pops,id_s => false | id_pops,id_delete => false | id_pops,id_T => false | id_pops,id_case9 => false | id_pops,id_record_extract => false | id_pops,id_excl => false | id_pops,id_case2 => false | id_pops,id_nil => false | id_pops,id_eqc => false | id_pops,id_case0 => false | id_pops,id_request => false | id_pops,id_locker2_check_available => false | id_pops,id_not => false | id_pops,id_pushs => false | id_pops,id_locker2_promote_pending => false | id_pops,id_mcrlrecord => false | id_pops,id_locker2_obtainables => false | id_pops,id_cons => false | id_pops,id_push1 => false | id_pops,id_case1 => false | id_pops,id_element => false | id_pops,id_locker2_adduniq => false | id_pops,id_and => false | id_pops,id_empty => false | id_pops,id_record_updates => false | id_pops,id_lock => false | id_pops,id_excllock => false | id_pops,id_pid => false | id_pops,id_calls => false | id_pops,id_subtract => false | id_pops,id_tag => false | id_pops,id_equal => false | id_pops,id_eq => false | id_pops,id_tops => false | id_pops,id_locker2_claim_lock => false | id_pops,id_pending => false | id_pops,id_andt => false | id_pops,id_tuplenil => false | id_pops,id_append => false | id_pops,id_0 => false | id_pops,id_case8 => false | id_locker2_map_claim_lock,id_or => true | id_locker2_map_claim_lock,id_gen_tag => true | id_locker2_map_claim_lock,id_record_new => true | id_locker2_map_claim_lock,id_a => true | id_locker2_map_claim_lock,id_locker2_release_lock => true | id_locker2_map_claim_lock,id_eqt => true | id_locker2_map_claim_lock,id_istops => true | id_locker2_map_claim_lock,id_locker2_map_add_pending => true | id_locker2_map_claim_lock,id_release => true | id_locker2_map_claim_lock,id_locker2_obtainable => true | id_locker2_map_claim_lock,id_imp => true | id_locker2_map_claim_lock,id_stack => true | id_locker2_map_claim_lock,id_locker2_map_promote_pending => true | id_locker2_map_claim_lock,id_locker => true | id_locker2_map_claim_lock,id_case4 => true | id_locker2_map_claim_lock,id_int => true | id_locker2_map_claim_lock,id_push => true | id_locker2_map_claim_lock,id_locker2_add_pending => true | id_locker2_map_claim_lock,id_true => true | id_locker2_map_claim_lock,id_locker2_check_availables => true | id_locker2_map_claim_lock,id_F => true | id_locker2_map_claim_lock,id_eqs => true | id_locker2_map_claim_lock,id_record_update => true | id_locker2_map_claim_lock,id_false => true | id_locker2_map_claim_lock,id_gen_modtageq => true | id_locker2_map_claim_lock,id_undefined => true | id_locker2_map_claim_lock,id_nocalls => true | id_locker2_map_claim_lock,id_locker2_remove_pending => true | id_locker2_map_claim_lock,id_resource => true | id_locker2_map_claim_lock,id_case6 => true | id_locker2_map_claim_lock,id_if => true | id_locker2_map_claim_lock,id_pops => true | id_locker2_map_claim_lock,id_locker2_map_claim_lock => true | id_locker2_map_claim_lock,id_ok => false | id_locker2_map_claim_lock,id_case5 => false | id_locker2_map_claim_lock,id_tuple => false | id_locker2_map_claim_lock,id_member => false | id_locker2_map_claim_lock,id_s => false | id_locker2_map_claim_lock,id_delete => false | id_locker2_map_claim_lock,id_T => false | id_locker2_map_claim_lock,id_case9 => false | id_locker2_map_claim_lock,id_record_extract => false | id_locker2_map_claim_lock,id_excl => false | id_locker2_map_claim_lock,id_case2 => false | id_locker2_map_claim_lock,id_nil => false | id_locker2_map_claim_lock,id_eqc => false | id_locker2_map_claim_lock,id_case0 => false | id_locker2_map_claim_lock,id_request => false | id_locker2_map_claim_lock,id_locker2_check_available => false | id_locker2_map_claim_lock,id_not => false | id_locker2_map_claim_lock,id_pushs => false | id_locker2_map_claim_lock,id_locker2_promote_pending => false | id_locker2_map_claim_lock,id_mcrlrecord => false | id_locker2_map_claim_lock,id_locker2_obtainables => false | id_locker2_map_claim_lock,id_cons => false | id_locker2_map_claim_lock,id_push1 => false | id_locker2_map_claim_lock,id_case1 => false | id_locker2_map_claim_lock,id_element => false | id_locker2_map_claim_lock,id_locker2_adduniq => false | id_locker2_map_claim_lock,id_and => false | id_locker2_map_claim_lock,id_empty => false | id_locker2_map_claim_lock,id_record_updates => false | id_locker2_map_claim_lock,id_lock => false | id_locker2_map_claim_lock,id_excllock => false | id_locker2_map_claim_lock,id_pid => false | id_locker2_map_claim_lock,id_calls => false | id_locker2_map_claim_lock,id_subtract => false | id_locker2_map_claim_lock,id_tag => false | id_locker2_map_claim_lock,id_equal => false | id_locker2_map_claim_lock,id_eq => false | id_locker2_map_claim_lock,id_tops => false | id_locker2_map_claim_lock,id_locker2_claim_lock => false | id_locker2_map_claim_lock,id_pending => false | id_locker2_map_claim_lock,id_andt => false | id_locker2_map_claim_lock,id_tuplenil => false | id_locker2_map_claim_lock,id_append => false | id_locker2_map_claim_lock,id_0 => false | id_locker2_map_claim_lock,id_case8 => false | id_ok,id_or => true | id_ok,id_gen_tag => true | id_ok,id_record_new => true | id_ok,id_a => true | id_ok,id_locker2_release_lock => true | id_ok,id_eqt => true | id_ok,id_istops => true | id_ok,id_locker2_map_add_pending => true | id_ok,id_release => true | id_ok,id_locker2_obtainable => true | id_ok,id_imp => true | id_ok,id_stack => true | id_ok,id_locker2_map_promote_pending => true | id_ok,id_locker => true | id_ok,id_case4 => true | id_ok,id_int => true | id_ok,id_push => true | id_ok,id_locker2_add_pending => true | id_ok,id_true => true | id_ok,id_locker2_check_availables => true | id_ok,id_F => true | id_ok,id_eqs => true | id_ok,id_record_update => true | id_ok,id_false => true | id_ok,id_gen_modtageq => true | id_ok,id_undefined => true | id_ok,id_nocalls => true | id_ok,id_locker2_remove_pending => true | id_ok,id_resource => true | id_ok,id_case6 => true | id_ok,id_if => true | id_ok,id_pops => true | id_ok,id_locker2_map_claim_lock => true | id_ok,id_ok => true | id_ok,id_case5 => false | id_ok,id_tuple => false | id_ok,id_member => false | id_ok,id_s => false | id_ok,id_delete => false | id_ok,id_T => false | id_ok,id_case9 => false | id_ok,id_record_extract => false | id_ok,id_excl => false | id_ok,id_case2 => false | id_ok,id_nil => false | id_ok,id_eqc => false | id_ok,id_case0 => false | id_ok,id_request => false | id_ok,id_locker2_check_available => false | id_ok,id_not => false | id_ok,id_pushs => false | id_ok,id_locker2_promote_pending => false | id_ok,id_mcrlrecord => false | id_ok,id_locker2_obtainables => false | id_ok,id_cons => false | id_ok,id_push1 => false | id_ok,id_case1 => false | id_ok,id_element => false | id_ok,id_locker2_adduniq => false | id_ok,id_and => false | id_ok,id_empty => false | id_ok,id_record_updates => false | id_ok,id_lock => false | id_ok,id_excllock => false | id_ok,id_pid => false | id_ok,id_calls => false | id_ok,id_subtract => false | id_ok,id_tag => false | id_ok,id_equal => false | id_ok,id_eq => false | id_ok,id_tops => false | id_ok,id_locker2_claim_lock => false | id_ok,id_pending => false | id_ok,id_andt => false | id_ok,id_tuplenil => false | id_ok,id_append => false | id_ok,id_0 => false | id_ok,id_case8 => false | id_case5,id_or => true | id_case5,id_gen_tag => true | id_case5,id_record_new => true | id_case5,id_a => true | id_case5,id_locker2_release_lock => true | id_case5,id_eqt => true | id_case5,id_istops => true | id_case5,id_locker2_map_add_pending => true | id_case5,id_release => true | id_case5,id_locker2_obtainable => true | id_case5,id_imp => true | id_case5,id_stack => true | id_case5,id_locker2_map_promote_pending => true | id_case5,id_locker => true | id_case5,id_case4 => true | id_case5,id_int => true | id_case5,id_push => true | id_case5,id_locker2_add_pending => true | id_case5,id_true => true | id_case5,id_locker2_check_availables => true | id_case5,id_F => true | id_case5,id_eqs => true | id_case5,id_record_update => true | id_case5,id_false => true | id_case5,id_gen_modtageq => true | id_case5,id_undefined => true | id_case5,id_nocalls => true | id_case5,id_locker2_remove_pending => true | id_case5,id_resource => true | id_case5,id_case6 => true | id_case5,id_if => true | id_case5,id_pops => true | id_case5,id_locker2_map_claim_lock => true | id_case5,id_ok => true | id_case5,id_case5 => true | id_case5,id_tuple => false | id_case5,id_member => false | id_case5,id_s => false | id_case5,id_delete => false | id_case5,id_T => false | id_case5,id_case9 => false | id_case5,id_record_extract => false | id_case5,id_excl => false | id_case5,id_case2 => false | id_case5,id_nil => false | id_case5,id_eqc => false | id_case5,id_case0 => false | id_case5,id_request => false | id_case5,id_locker2_check_available => false | id_case5,id_not => false | id_case5,id_pushs => false | id_case5,id_locker2_promote_pending => false | id_case5,id_mcrlrecord => false | id_case5,id_locker2_obtainables => false | id_case5,id_cons => false | id_case5,id_push1 => false | id_case5,id_case1 => false | id_case5,id_element => false | id_case5,id_locker2_adduniq => false | id_case5,id_and => false | id_case5,id_empty => false | id_case5,id_record_updates => false | id_case5,id_lock => false | id_case5,id_excllock => false | id_case5,id_pid => false | id_case5,id_calls => false | id_case5,id_subtract => false | id_case5,id_tag => false | id_case5,id_equal => false | id_case5,id_eq => false | id_case5,id_tops => false | id_case5,id_locker2_claim_lock => false | id_case5,id_pending => false | id_case5,id_andt => false | id_case5,id_tuplenil => false | id_case5,id_append => false | id_case5,id_0 => false | id_case5,id_case8 => false | id_tuple,id_or => true | id_tuple,id_gen_tag => true | id_tuple,id_record_new => true | id_tuple,id_a => true | id_tuple,id_locker2_release_lock => true | id_tuple,id_eqt => true | id_tuple,id_istops => true | id_tuple,id_locker2_map_add_pending => true | id_tuple,id_release => true | id_tuple,id_locker2_obtainable => true | id_tuple,id_imp => true | id_tuple,id_stack => true | id_tuple,id_locker2_map_promote_pending => true | id_tuple,id_locker => true | id_tuple,id_case4 => true | id_tuple,id_int => true | id_tuple,id_push => true | id_tuple,id_locker2_add_pending => true | id_tuple,id_true => true | id_tuple,id_locker2_check_availables => true | id_tuple,id_F => true | id_tuple,id_eqs => true | id_tuple,id_record_update => true | id_tuple,id_false => true | id_tuple,id_gen_modtageq => true | id_tuple,id_undefined => true | id_tuple,id_nocalls => true | id_tuple,id_locker2_remove_pending => true | id_tuple,id_resource => true | id_tuple,id_case6 => true | id_tuple,id_if => true | id_tuple,id_pops => true | id_tuple,id_locker2_map_claim_lock => true | id_tuple,id_ok => true | id_tuple,id_case5 => true | id_tuple,id_tuple => true | id_tuple,id_member => false | id_tuple,id_s => false | id_tuple,id_delete => false | id_tuple,id_T => false | id_tuple,id_case9 => false | id_tuple,id_record_extract => false | id_tuple,id_excl => false | id_tuple,id_case2 => false | id_tuple,id_nil => false | id_tuple,id_eqc => false | id_tuple,id_case0 => false | id_tuple,id_request => false | id_tuple,id_locker2_check_available => false | id_tuple,id_not => false | id_tuple,id_pushs => false | id_tuple,id_locker2_promote_pending => false | id_tuple,id_mcrlrecord => false | id_tuple,id_locker2_obtainables => false | id_tuple,id_cons => false | id_tuple,id_push1 => false | id_tuple,id_case1 => false | id_tuple,id_element => false | id_tuple,id_locker2_adduniq => false | id_tuple,id_and => false | id_tuple,id_empty => false | id_tuple,id_record_updates => false | id_tuple,id_lock => false | id_tuple,id_excllock => false | id_tuple,id_pid => false | id_tuple,id_calls => false | id_tuple,id_subtract => false | id_tuple,id_tag => false | id_tuple,id_equal => false | id_tuple,id_eq => false | id_tuple,id_tops => false | id_tuple,id_locker2_claim_lock => false | id_tuple,id_pending => false | id_tuple,id_andt => false | id_tuple,id_tuplenil => false | id_tuple,id_append => false | id_tuple,id_0 => false | id_tuple,id_case8 => false | id_member,id_or => true | id_member,id_gen_tag => true | id_member,id_record_new => true | id_member,id_a => true | id_member,id_locker2_release_lock => true | id_member,id_eqt => true | id_member,id_istops => true | id_member,id_locker2_map_add_pending => true | id_member,id_release => true | id_member,id_locker2_obtainable => true | id_member,id_imp => true | id_member,id_stack => true | id_member,id_locker2_map_promote_pending => true | id_member,id_locker => true | id_member,id_case4 => true | id_member,id_int => true | id_member,id_push => true | id_member,id_locker2_add_pending => true | id_member,id_true => true | id_member,id_locker2_check_availables => true | id_member,id_F => true | id_member,id_eqs => true | id_member,id_record_update => true | id_member,id_false => true | id_member,id_gen_modtageq => true | id_member,id_undefined => true | id_member,id_nocalls => true | id_member,id_locker2_remove_pending => true | id_member,id_resource => true | id_member,id_case6 => true | id_member,id_if => true | id_member,id_pops => true | id_member,id_locker2_map_claim_lock => true | id_member,id_ok => true | id_member,id_case5 => true | id_member,id_tuple => true | id_member,id_member => true | id_member,id_s => false | id_member,id_delete => false | id_member,id_T => false | id_member,id_case9 => false | id_member,id_record_extract => false | id_member,id_excl => false | id_member,id_case2 => false | id_member,id_nil => false | id_member,id_eqc => false | id_member,id_case0 => false | id_member,id_request => false | id_member,id_locker2_check_available => false | id_member,id_not => false | id_member,id_pushs => false | id_member,id_locker2_promote_pending => false | id_member,id_mcrlrecord => false | id_member,id_locker2_obtainables => false | id_member,id_cons => false | id_member,id_push1 => false | id_member,id_case1 => false | id_member,id_element => false | id_member,id_locker2_adduniq => false | id_member,id_and => false | id_member,id_empty => false | id_member,id_record_updates => false | id_member,id_lock => false | id_member,id_excllock => false | id_member,id_pid => false | id_member,id_calls => false | id_member,id_subtract => false | id_member,id_tag => false | id_member,id_equal => false | id_member,id_eq => false | id_member,id_tops => false | id_member,id_locker2_claim_lock => false | id_member,id_pending => false | id_member,id_andt => false | id_member,id_tuplenil => false | id_member,id_append => false | id_member,id_0 => false | id_member,id_case8 => false | id_s,id_or => true | id_s,id_gen_tag => true | id_s,id_record_new => true | id_s,id_a => true | id_s,id_locker2_release_lock => true | id_s,id_eqt => true | id_s,id_istops => true | id_s,id_locker2_map_add_pending => true | id_s,id_release => true | id_s,id_locker2_obtainable => true | id_s,id_imp => true | id_s,id_stack => true | id_s,id_locker2_map_promote_pending => true | id_s,id_locker => true | id_s,id_case4 => true | id_s,id_int => true | id_s,id_push => true | id_s,id_locker2_add_pending => true | id_s,id_true => true | id_s,id_locker2_check_availables => true | id_s,id_F => true | id_s,id_eqs => true | id_s,id_record_update => true | id_s,id_false => true | id_s,id_gen_modtageq => true | id_s,id_undefined => true | id_s,id_nocalls => true | id_s,id_locker2_remove_pending => true | id_s,id_resource => true | id_s,id_case6 => true | id_s,id_if => true | id_s,id_pops => true | id_s,id_locker2_map_claim_lock => true | id_s,id_ok => true | id_s,id_case5 => true | id_s,id_tuple => true | id_s,id_member => true | id_s,id_s => true | id_s,id_delete => false | id_s,id_T => false | id_s,id_case9 => false | id_s,id_record_extract => false | id_s,id_excl => false | id_s,id_case2 => false | id_s,id_nil => false | id_s,id_eqc => false | id_s,id_case0 => false | id_s,id_request => false | id_s,id_locker2_check_available => false | id_s,id_not => false | id_s,id_pushs => false | id_s,id_locker2_promote_pending => false | id_s,id_mcrlrecord => false | id_s,id_locker2_obtainables => false | id_s,id_cons => false | id_s,id_push1 => false | id_s,id_case1 => false | id_s,id_element => false | id_s,id_locker2_adduniq => false | id_s,id_and => false | id_s,id_empty => false | id_s,id_record_updates => false | id_s,id_lock => false | id_s,id_excllock => false | id_s,id_pid => false | id_s,id_calls => false | id_s,id_subtract => false | id_s,id_tag => false | id_s,id_equal => false | id_s,id_eq => false | id_s,id_tops => false | id_s,id_locker2_claim_lock => false | id_s,id_pending => false | id_s,id_andt => false | id_s,id_tuplenil => false | id_s,id_append => false | id_s,id_0 => false | id_s,id_case8 => false | id_delete,id_or => true | id_delete,id_gen_tag => true | id_delete,id_record_new => true | id_delete,id_a => true | id_delete,id_locker2_release_lock => true | id_delete,id_eqt => true | id_delete,id_istops => true | id_delete,id_locker2_map_add_pending => true | id_delete,id_release => true | id_delete,id_locker2_obtainable => true | id_delete,id_imp => true | id_delete,id_stack => true | id_delete,id_locker2_map_promote_pending => true | id_delete,id_locker => true | id_delete,id_case4 => true | id_delete,id_int => true | id_delete,id_push => true | id_delete,id_locker2_add_pending => true | id_delete,id_true => true | id_delete,id_locker2_check_availables => true | id_delete,id_F => true | id_delete,id_eqs => true | id_delete,id_record_update => true | id_delete,id_false => true | id_delete,id_gen_modtageq => true | id_delete,id_undefined => true | id_delete,id_nocalls => true | id_delete,id_locker2_remove_pending => true | id_delete,id_resource => true | id_delete,id_case6 => true | id_delete,id_if => true | id_delete,id_pops => true | id_delete,id_locker2_map_claim_lock => true | id_delete,id_ok => true | id_delete,id_case5 => true | id_delete,id_tuple => true | id_delete,id_member => true | id_delete,id_s => true | id_delete,id_delete => true | id_delete,id_T => false | id_delete,id_case9 => false | id_delete,id_record_extract => false | id_delete,id_excl => false | id_delete,id_case2 => false | id_delete,id_nil => false | id_delete,id_eqc => false | id_delete,id_case0 => false | id_delete,id_request => false | id_delete,id_locker2_check_available => false | id_delete,id_not => false | id_delete,id_pushs => false | id_delete,id_locker2_promote_pending => false | id_delete,id_mcrlrecord => false | id_delete,id_locker2_obtainables => false | id_delete,id_cons => false | id_delete,id_push1 => false | id_delete,id_case1 => false | id_delete,id_element => false | id_delete,id_locker2_adduniq => false | id_delete,id_and => false | id_delete,id_empty => false | id_delete,id_record_updates => false | id_delete,id_lock => false | id_delete,id_excllock => false | id_delete,id_pid => false | id_delete,id_calls => false | id_delete,id_subtract => false | id_delete,id_tag => false | id_delete,id_equal => false | id_delete,id_eq => false | id_delete,id_tops => false | id_delete,id_locker2_claim_lock => false | id_delete,id_pending => false | id_delete,id_andt => false | id_delete,id_tuplenil => false | id_delete,id_append => false | id_delete,id_0 => false | id_delete,id_case8 => false | id_T,id_or => true | id_T,id_gen_tag => true | id_T,id_record_new => true | id_T,id_a => true | id_T,id_locker2_release_lock => true | id_T,id_eqt => true | id_T,id_istops => true | id_T,id_locker2_map_add_pending => true | id_T,id_release => true | id_T,id_locker2_obtainable => true | id_T,id_imp => true | id_T,id_stack => true | id_T,id_locker2_map_promote_pending => true | id_T,id_locker => true | id_T,id_case4 => true | id_T,id_int => true | id_T,id_push => true | id_T,id_locker2_add_pending => true | id_T,id_true => true | id_T,id_locker2_check_availables => true | id_T,id_F => true | id_T,id_eqs => true | id_T,id_record_update => true | id_T,id_false => true | id_T,id_gen_modtageq => true | id_T,id_undefined => true | id_T,id_nocalls => true | id_T,id_locker2_remove_pending => true | id_T,id_resource => true | id_T,id_case6 => true | id_T,id_if => true | id_T,id_pops => true | id_T,id_locker2_map_claim_lock => true | id_T,id_ok => true | id_T,id_case5 => true | id_T,id_tuple => true | id_T,id_member => true | id_T,id_s => true | id_T,id_delete => true | id_T,id_T => true | id_T,id_case9 => false | id_T,id_record_extract => false | id_T,id_excl => false | id_T,id_case2 => false | id_T,id_nil => false | id_T,id_eqc => false | id_T,id_case0 => false | id_T,id_request => false | id_T,id_locker2_check_available => false | id_T,id_not => false | id_T,id_pushs => false | id_T,id_locker2_promote_pending => false | id_T,id_mcrlrecord => false | id_T,id_locker2_obtainables => false | id_T,id_cons => false | id_T,id_push1 => false | id_T,id_case1 => false | id_T,id_element => false | id_T,id_locker2_adduniq => false | id_T,id_and => false | id_T,id_empty => false | id_T,id_record_updates => false | id_T,id_lock => false | id_T,id_excllock => false | id_T,id_pid => false | id_T,id_calls => false | id_T,id_subtract => false | id_T,id_tag => false | id_T,id_equal => false | id_T,id_eq => false | id_T,id_tops => false | id_T,id_locker2_claim_lock => false | id_T,id_pending => false | id_T,id_andt => false | id_T,id_tuplenil => false | id_T,id_append => false | id_T,id_0 => false | id_T,id_case8 => false | id_case9,id_or => true | id_case9,id_gen_tag => true | id_case9,id_record_new => true | id_case9,id_a => true | id_case9,id_locker2_release_lock => true | id_case9,id_eqt => true | id_case9,id_istops => true | id_case9,id_locker2_map_add_pending => true | id_case9,id_release => true | id_case9,id_locker2_obtainable => true | id_case9,id_imp => true | id_case9,id_stack => true | id_case9,id_locker2_map_promote_pending => true | id_case9,id_locker => true | id_case9,id_case4 => true | id_case9,id_int => true | id_case9,id_push => true | id_case9,id_locker2_add_pending => true | id_case9,id_true => true | id_case9,id_locker2_check_availables => true | id_case9,id_F => true | id_case9,id_eqs => true | id_case9,id_record_update => true | id_case9,id_false => true | id_case9,id_gen_modtageq => true | id_case9,id_undefined => true | id_case9,id_nocalls => true | id_case9,id_locker2_remove_pending => true | id_case9,id_resource => true | id_case9,id_case6 => true | id_case9,id_if => true | id_case9,id_pops => true | id_case9,id_locker2_map_claim_lock => true | id_case9,id_ok => true | id_case9,id_case5 => true | id_case9,id_tuple => true | id_case9,id_member => true | id_case9,id_s => true | id_case9,id_delete => true | id_case9,id_T => true | id_case9,id_case9 => true | id_case9,id_record_extract => false | id_case9,id_excl => false | id_case9,id_case2 => false | id_case9,id_nil => false | id_case9,id_eqc => false | id_case9,id_case0 => false | id_case9,id_request => false | id_case9,id_locker2_check_available => false | id_case9,id_not => false | id_case9,id_pushs => false | id_case9,id_locker2_promote_pending => false | id_case9,id_mcrlrecord => false | id_case9,id_locker2_obtainables => false | id_case9,id_cons => false | id_case9,id_push1 => false | id_case9,id_case1 => false | id_case9,id_element => false | id_case9,id_locker2_adduniq => false | id_case9,id_and => false | id_case9,id_empty => false | id_case9,id_record_updates => false | id_case9,id_lock => false | id_case9,id_excllock => false | id_case9,id_pid => false | id_case9,id_calls => false | id_case9,id_subtract => false | id_case9,id_tag => false | id_case9,id_equal => false | id_case9,id_eq => false | id_case9,id_tops => false | id_case9,id_locker2_claim_lock => false | id_case9,id_pending => false | id_case9,id_andt => false | id_case9,id_tuplenil => false | id_case9,id_append => false | id_case9,id_0 => false | id_case9,id_case8 => false | id_record_extract,id_or => true | id_record_extract,id_gen_tag => true | id_record_extract,id_record_new => true | id_record_extract,id_a => true | id_record_extract,id_locker2_release_lock => true | id_record_extract,id_eqt => true | id_record_extract,id_istops => true | id_record_extract,id_locker2_map_add_pending => true | id_record_extract,id_release => true | id_record_extract,id_locker2_obtainable => true | id_record_extract,id_imp => true | id_record_extract,id_stack => true | id_record_extract,id_locker2_map_promote_pending => true | id_record_extract,id_locker => true | id_record_extract,id_case4 => true | id_record_extract,id_int => true | id_record_extract,id_push => true | id_record_extract,id_locker2_add_pending => true | id_record_extract,id_true => true | id_record_extract,id_locker2_check_availables => true | id_record_extract,id_F => true | id_record_extract,id_eqs => true | id_record_extract,id_record_update => true | id_record_extract,id_false => true | id_record_extract,id_gen_modtageq => true | id_record_extract,id_undefined => true | id_record_extract,id_nocalls => true | id_record_extract,id_locker2_remove_pending => true | id_record_extract,id_resource => true | id_record_extract,id_case6 => true | id_record_extract,id_if => true | id_record_extract,id_pops => true | id_record_extract,id_locker2_map_claim_lock => true | id_record_extract,id_ok => true | id_record_extract,id_case5 => true | id_record_extract,id_tuple => true | id_record_extract,id_member => true | id_record_extract,id_s => true | id_record_extract,id_delete => true | id_record_extract,id_T => true | id_record_extract,id_case9 => true | id_record_extract,id_record_extract => true | id_record_extract,id_excl => false | id_record_extract,id_case2 => false | id_record_extract,id_nil => false | id_record_extract,id_eqc => false | id_record_extract,id_case0 => false | id_record_extract,id_request => false | id_record_extract,id_locker2_check_available => false | id_record_extract,id_not => false | id_record_extract,id_pushs => false | id_record_extract,id_locker2_promote_pending => false | id_record_extract,id_mcrlrecord => false | id_record_extract,id_locker2_obtainables => false | id_record_extract,id_cons => false | id_record_extract,id_push1 => false | id_record_extract,id_case1 => false | id_record_extract,id_element => false | id_record_extract,id_locker2_adduniq => false | id_record_extract,id_and => false | id_record_extract,id_empty => false | id_record_extract,id_record_updates => false | id_record_extract,id_lock => false | id_record_extract,id_excllock => false | id_record_extract,id_pid => false | id_record_extract,id_calls => false | id_record_extract,id_subtract => false | id_record_extract,id_tag => false | id_record_extract,id_equal => false | id_record_extract,id_eq => false | id_record_extract,id_tops => false | id_record_extract,id_locker2_claim_lock => false | id_record_extract,id_pending => false | id_record_extract,id_andt => false | id_record_extract,id_tuplenil => false | id_record_extract,id_append => false | id_record_extract,id_0 => false | id_record_extract,id_case8 => false | id_excl,id_or => true | id_excl,id_gen_tag => true | id_excl,id_record_new => true | id_excl,id_a => true | id_excl,id_locker2_release_lock => true | id_excl,id_eqt => true | id_excl,id_istops => true | id_excl,id_locker2_map_add_pending => true | id_excl,id_release => true | id_excl,id_locker2_obtainable => true | id_excl,id_imp => true | id_excl,id_stack => true | id_excl,id_locker2_map_promote_pending => true | id_excl,id_locker => true | id_excl,id_case4 => true | id_excl,id_int => true | id_excl,id_push => true | id_excl,id_locker2_add_pending => true | id_excl,id_true => true | id_excl,id_locker2_check_availables => true | id_excl,id_F => true | id_excl,id_eqs => true | id_excl,id_record_update => true | id_excl,id_false => true | id_excl,id_gen_modtageq => true | id_excl,id_undefined => true | id_excl,id_nocalls => true | id_excl,id_locker2_remove_pending => true | id_excl,id_resource => true | id_excl,id_case6 => true | id_excl,id_if => true | id_excl,id_pops => true | id_excl,id_locker2_map_claim_lock => true | id_excl,id_ok => true | id_excl,id_case5 => true | id_excl,id_tuple => true | id_excl,id_member => true | id_excl,id_s => true | id_excl,id_delete => true | id_excl,id_T => true | id_excl,id_case9 => true | id_excl,id_record_extract => true | id_excl,id_excl => true | id_excl,id_case2 => false | id_excl,id_nil => false | id_excl,id_eqc => false | id_excl,id_case0 => false | id_excl,id_request => false | id_excl,id_locker2_check_available => false | id_excl,id_not => false | id_excl,id_pushs => false | id_excl,id_locker2_promote_pending => false | id_excl,id_mcrlrecord => false | id_excl,id_locker2_obtainables => false | id_excl,id_cons => false | id_excl,id_push1 => false | id_excl,id_case1 => false | id_excl,id_element => false | id_excl,id_locker2_adduniq => false | id_excl,id_and => false | id_excl,id_empty => false | id_excl,id_record_updates => false | id_excl,id_lock => false | id_excl,id_excllock => false | id_excl,id_pid => false | id_excl,id_calls => false | id_excl,id_subtract => false | id_excl,id_tag => false | id_excl,id_equal => false | id_excl,id_eq => false | id_excl,id_tops => false | id_excl,id_locker2_claim_lock => false | id_excl,id_pending => false | id_excl,id_andt => false | id_excl,id_tuplenil => false | id_excl,id_append => false | id_excl,id_0 => false | id_excl,id_case8 => false | id_case2,id_or => true | id_case2,id_gen_tag => true | id_case2,id_record_new => true | id_case2,id_a => true | id_case2,id_locker2_release_lock => true | id_case2,id_eqt => true | id_case2,id_istops => true | id_case2,id_locker2_map_add_pending => true | id_case2,id_release => true | id_case2,id_locker2_obtainable => true | id_case2,id_imp => true | id_case2,id_stack => true | id_case2,id_locker2_map_promote_pending => true | id_case2,id_locker => true | id_case2,id_case4 => true | id_case2,id_int => true | id_case2,id_push => true | id_case2,id_locker2_add_pending => true | id_case2,id_true => true | id_case2,id_locker2_check_availables => true | id_case2,id_F => true | id_case2,id_eqs => true | id_case2,id_record_update => true | id_case2,id_false => true | id_case2,id_gen_modtageq => true | id_case2,id_undefined => true | id_case2,id_nocalls => true | id_case2,id_locker2_remove_pending => true | id_case2,id_resource => true | id_case2,id_case6 => true | id_case2,id_if => true | id_case2,id_pops => true | id_case2,id_locker2_map_claim_lock => true | id_case2,id_ok => true | id_case2,id_case5 => true | id_case2,id_tuple => true | id_case2,id_member => true | id_case2,id_s => true | id_case2,id_delete => true | id_case2,id_T => true | id_case2,id_case9 => true | id_case2,id_record_extract => true | id_case2,id_excl => true | id_case2,id_case2 => true | id_case2,id_nil => false | id_case2,id_eqc => false | id_case2,id_case0 => false | id_case2,id_request => false | id_case2,id_locker2_check_available => false | id_case2,id_not => false | id_case2,id_pushs => false | id_case2,id_locker2_promote_pending => false | id_case2,id_mcrlrecord => false | id_case2,id_locker2_obtainables => false | id_case2,id_cons => false | id_case2,id_push1 => false | id_case2,id_case1 => false | id_case2,id_element => false | id_case2,id_locker2_adduniq => false | id_case2,id_and => false | id_case2,id_empty => false | id_case2,id_record_updates => false | id_case2,id_lock => false | id_case2,id_excllock => false | id_case2,id_pid => false | id_case2,id_calls => false | id_case2,id_subtract => false | id_case2,id_tag => false | id_case2,id_equal => false | id_case2,id_eq => false | id_case2,id_tops => false | id_case2,id_locker2_claim_lock => false | id_case2,id_pending => false | id_case2,id_andt => false | id_case2,id_tuplenil => false | id_case2,id_append => false | id_case2,id_0 => false | id_case2,id_case8 => false | id_nil,id_or => true | id_nil,id_gen_tag => true | id_nil,id_record_new => true | id_nil,id_a => true | id_nil,id_locker2_release_lock => true | id_nil,id_eqt => true | id_nil,id_istops => true | id_nil,id_locker2_map_add_pending => true | id_nil,id_release => true | id_nil,id_locker2_obtainable => true | id_nil,id_imp => true | id_nil,id_stack => true | id_nil,id_locker2_map_promote_pending => true | id_nil,id_locker => true | id_nil,id_case4 => true | id_nil,id_int => true | id_nil,id_push => true | id_nil,id_locker2_add_pending => true | id_nil,id_true => true | id_nil,id_locker2_check_availables => true | id_nil,id_F => true | id_nil,id_eqs => true | id_nil,id_record_update => true | id_nil,id_false => true | id_nil,id_gen_modtageq => true | id_nil,id_undefined => true | id_nil,id_nocalls => true | id_nil,id_locker2_remove_pending => true | id_nil,id_resource => true | id_nil,id_case6 => true | id_nil,id_if => true | id_nil,id_pops => true | id_nil,id_locker2_map_claim_lock => true | id_nil,id_ok => true | id_nil,id_case5 => true | id_nil,id_tuple => true | id_nil,id_member => true | id_nil,id_s => true | id_nil,id_delete => true | id_nil,id_T => true | id_nil,id_case9 => true | id_nil,id_record_extract => true | id_nil,id_excl => true | id_nil,id_case2 => true | id_nil,id_nil => true | id_nil,id_eqc => false | id_nil,id_case0 => false | id_nil,id_request => false | id_nil,id_locker2_check_available => false | id_nil,id_not => false | id_nil,id_pushs => false | id_nil,id_locker2_promote_pending => false | id_nil,id_mcrlrecord => false | id_nil,id_locker2_obtainables => false | id_nil,id_cons => false | id_nil,id_push1 => false | id_nil,id_case1 => false | id_nil,id_element => false | id_nil,id_locker2_adduniq => false | id_nil,id_and => false | id_nil,id_empty => false | id_nil,id_record_updates => false | id_nil,id_lock => false | id_nil,id_excllock => false | id_nil,id_pid => false | id_nil,id_calls => false | id_nil,id_subtract => false | id_nil,id_tag => false | id_nil,id_equal => false | id_nil,id_eq => false | id_nil,id_tops => false | id_nil,id_locker2_claim_lock => false | id_nil,id_pending => false | id_nil,id_andt => false | id_nil,id_tuplenil => false | id_nil,id_append => false | id_nil,id_0 => false | id_nil,id_case8 => false | id_eqc,id_or => true | id_eqc,id_gen_tag => true | id_eqc,id_record_new => true | id_eqc,id_a => true | id_eqc,id_locker2_release_lock => true | id_eqc,id_eqt => true | id_eqc,id_istops => true | id_eqc,id_locker2_map_add_pending => true | id_eqc,id_release => true | id_eqc,id_locker2_obtainable => true | id_eqc,id_imp => true | id_eqc,id_stack => true | id_eqc,id_locker2_map_promote_pending => true | id_eqc,id_locker => true | id_eqc,id_case4 => true | id_eqc,id_int => true | id_eqc,id_push => true | id_eqc,id_locker2_add_pending => true | id_eqc,id_true => true | id_eqc,id_locker2_check_availables => true | id_eqc,id_F => true | id_eqc,id_eqs => true | id_eqc,id_record_update => true | id_eqc,id_false => true | id_eqc,id_gen_modtageq => true | id_eqc,id_undefined => true | id_eqc,id_nocalls => true | id_eqc,id_locker2_remove_pending => true | id_eqc,id_resource => true | id_eqc,id_case6 => true | id_eqc,id_if => true | id_eqc,id_pops => true | id_eqc,id_locker2_map_claim_lock => true | id_eqc,id_ok => true | id_eqc,id_case5 => true | id_eqc,id_tuple => true | id_eqc,id_member => true | id_eqc,id_s => true | id_eqc,id_delete => true | id_eqc,id_T => true | id_eqc,id_case9 => true | id_eqc,id_record_extract => true | id_eqc,id_excl => true | id_eqc,id_case2 => true | id_eqc,id_nil => true | id_eqc,id_eqc => true | id_eqc,id_case0 => false | id_eqc,id_request => false | id_eqc,id_locker2_check_available => false | id_eqc,id_not => false | id_eqc,id_pushs => false | id_eqc,id_locker2_promote_pending => false | id_eqc,id_mcrlrecord => false | id_eqc,id_locker2_obtainables => false | id_eqc,id_cons => false | id_eqc,id_push1 => false | id_eqc,id_case1 => false | id_eqc,id_element => false | id_eqc,id_locker2_adduniq => false | id_eqc,id_and => false | id_eqc,id_empty => false | id_eqc,id_record_updates => false | id_eqc,id_lock => false | id_eqc,id_excllock => false | id_eqc,id_pid => false | id_eqc,id_calls => false | id_eqc,id_subtract => false | id_eqc,id_tag => false | id_eqc,id_equal => false | id_eqc,id_eq => false | id_eqc,id_tops => false | id_eqc,id_locker2_claim_lock => false | id_eqc,id_pending => false | id_eqc,id_andt => false | id_eqc,id_tuplenil => false | id_eqc,id_append => false | id_eqc,id_0 => false | id_eqc,id_case8 => false | id_case0,id_or => true | id_case0,id_gen_tag => true | id_case0,id_record_new => true | id_case0,id_a => true | id_case0,id_locker2_release_lock => true | id_case0,id_eqt => true | id_case0,id_istops => true | id_case0,id_locker2_map_add_pending => true | id_case0,id_release => true | id_case0,id_locker2_obtainable => true | id_case0,id_imp => true | id_case0,id_stack => true | id_case0,id_locker2_map_promote_pending => true | id_case0,id_locker => true | id_case0,id_case4 => true | id_case0,id_int => true | id_case0,id_push => true | id_case0,id_locker2_add_pending => true | id_case0,id_true => true | id_case0,id_locker2_check_availables => true | id_case0,id_F => true | id_case0,id_eqs => true | id_case0,id_record_update => true | id_case0,id_false => true | id_case0,id_gen_modtageq => true | id_case0,id_undefined => true | id_case0,id_nocalls => true | id_case0,id_locker2_remove_pending => true | id_case0,id_resource => true | id_case0,id_case6 => true | id_case0,id_if => true | id_case0,id_pops => true | id_case0,id_locker2_map_claim_lock => true | id_case0,id_ok => true | id_case0,id_case5 => true | id_case0,id_tuple => true | id_case0,id_member => true | id_case0,id_s => true | id_case0,id_delete => true | id_case0,id_T => true | id_case0,id_case9 => true | id_case0,id_record_extract => true | id_case0,id_excl => true | id_case0,id_case2 => true | id_case0,id_nil => true | id_case0,id_eqc => true | id_case0,id_case0 => true | id_case0,id_request => false | id_case0,id_locker2_check_available => false | id_case0,id_not => false | id_case0,id_pushs => false | id_case0,id_locker2_promote_pending => false | id_case0,id_mcrlrecord => false | id_case0,id_locker2_obtainables => false | id_case0,id_cons => false | id_case0,id_push1 => false | id_case0,id_case1 => false | id_case0,id_element => false | id_case0,id_locker2_adduniq => false | id_case0,id_and => false | id_case0,id_empty => false | id_case0,id_record_updates => false | id_case0,id_lock => false | id_case0,id_excllock => false | id_case0,id_pid => false | id_case0,id_calls => false | id_case0,id_subtract => false | id_case0,id_tag => false | id_case0,id_equal => false | id_case0,id_eq => false | id_case0,id_tops => false | id_case0,id_locker2_claim_lock => false | id_case0,id_pending => false | id_case0,id_andt => false | id_case0,id_tuplenil => false | id_case0,id_append => false | id_case0,id_0 => false | id_case0,id_case8 => false | id_request,id_or => true | id_request,id_gen_tag => true | id_request,id_record_new => true | id_request,id_a => true | id_request,id_locker2_release_lock => true | id_request,id_eqt => true | id_request,id_istops => true | id_request,id_locker2_map_add_pending => true | id_request,id_release => true | id_request,id_locker2_obtainable => true | id_request,id_imp => true | id_request,id_stack => true | id_request,id_locker2_map_promote_pending => true | id_request,id_locker => true | id_request,id_case4 => true | id_request,id_int => true | id_request,id_push => true | id_request,id_locker2_add_pending => true | id_request,id_true => true | id_request,id_locker2_check_availables => true | id_request,id_F => true | id_request,id_eqs => true | id_request,id_record_update => true | id_request,id_false => true | id_request,id_gen_modtageq => true | id_request,id_undefined => true | id_request,id_nocalls => true | id_request,id_locker2_remove_pending => true | id_request,id_resource => true | id_request,id_case6 => true | id_request,id_if => true | id_request,id_pops => true | id_request,id_locker2_map_claim_lock => true | id_request,id_ok => true | id_request,id_case5 => true | id_request,id_tuple => true | id_request,id_member => true | id_request,id_s => true | id_request,id_delete => true | id_request,id_T => true | id_request,id_case9 => true | id_request,id_record_extract => true | id_request,id_excl => true | id_request,id_case2 => true | id_request,id_nil => true | id_request,id_eqc => true | id_request,id_case0 => true | id_request,id_request => true | id_request,id_locker2_check_available => false | id_request,id_not => false | id_request,id_pushs => false | id_request,id_locker2_promote_pending => false | id_request,id_mcrlrecord => false | id_request,id_locker2_obtainables => false | id_request,id_cons => false | id_request,id_push1 => false | id_request,id_case1 => false | id_request,id_element => false | id_request,id_locker2_adduniq => false | id_request,id_and => false | id_request,id_empty => false | id_request,id_record_updates => false | id_request,id_lock => false | id_request,id_excllock => false | id_request,id_pid => false | id_request,id_calls => false | id_request,id_subtract => false | id_request,id_tag => false | id_request,id_equal => false | id_request,id_eq => false | id_request,id_tops => false | id_request,id_locker2_claim_lock => false | id_request,id_pending => false | id_request,id_andt => false | id_request,id_tuplenil => false | id_request,id_append => false | id_request,id_0 => false | id_request,id_case8 => false | id_locker2_check_available,id_or => true | id_locker2_check_available,id_gen_tag => true | id_locker2_check_available,id_record_new => true | id_locker2_check_available,id_a => true | id_locker2_check_available,id_locker2_release_lock => true | id_locker2_check_available,id_eqt => true | id_locker2_check_available,id_istops => true | id_locker2_check_available,id_locker2_map_add_pending => true | id_locker2_check_available,id_release => true | id_locker2_check_available,id_locker2_obtainable => true | id_locker2_check_available,id_imp => true | id_locker2_check_available,id_stack => true | id_locker2_check_available,id_locker2_map_promote_pending => true | id_locker2_check_available,id_locker => true | id_locker2_check_available,id_case4 => true | id_locker2_check_available,id_int => true | id_locker2_check_available,id_push => true | id_locker2_check_available,id_locker2_add_pending => true | id_locker2_check_available,id_true => true | id_locker2_check_available,id_locker2_check_availables => true | id_locker2_check_available,id_F => true | id_locker2_check_available,id_eqs => true | id_locker2_check_available,id_record_update => true | id_locker2_check_available,id_false => true | id_locker2_check_available,id_gen_modtageq => true | id_locker2_check_available,id_undefined => true | id_locker2_check_available,id_nocalls => true | id_locker2_check_available,id_locker2_remove_pending => true | id_locker2_check_available,id_resource => true | id_locker2_check_available,id_case6 => true | id_locker2_check_available,id_if => true | id_locker2_check_available,id_pops => true | id_locker2_check_available,id_locker2_map_claim_lock => true | id_locker2_check_available,id_ok => true | id_locker2_check_available,id_case5 => true | id_locker2_check_available,id_tuple => true | id_locker2_check_available,id_member => true | id_locker2_check_available,id_s => true | id_locker2_check_available,id_delete => true | id_locker2_check_available,id_T => true | id_locker2_check_available,id_case9 => true | id_locker2_check_available,id_record_extract => true | id_locker2_check_available,id_excl => true | id_locker2_check_available,id_case2 => true | id_locker2_check_available,id_nil => true | id_locker2_check_available,id_eqc => true | id_locker2_check_available,id_case0 => true | id_locker2_check_available,id_request => true | id_locker2_check_available,id_locker2_check_available => true | id_locker2_check_available,id_not => false | id_locker2_check_available,id_pushs => false | id_locker2_check_available,id_locker2_promote_pending => false | id_locker2_check_available,id_mcrlrecord => false | id_locker2_check_available,id_locker2_obtainables => false | id_locker2_check_available,id_cons => false | id_locker2_check_available,id_push1 => false | id_locker2_check_available,id_case1 => false | id_locker2_check_available,id_element => false | id_locker2_check_available,id_locker2_adduniq => false | id_locker2_check_available,id_and => false | id_locker2_check_available,id_empty => false | id_locker2_check_available,id_record_updates => false | id_locker2_check_available,id_lock => false | id_locker2_check_available,id_excllock => false | id_locker2_check_available,id_pid => false | id_locker2_check_available,id_calls => false | id_locker2_check_available,id_subtract => false | id_locker2_check_available,id_tag => false | id_locker2_check_available,id_equal => false | id_locker2_check_available,id_eq => false | id_locker2_check_available,id_tops => false | id_locker2_check_available,id_locker2_claim_lock => false | id_locker2_check_available,id_pending => false | id_locker2_check_available,id_andt => false | id_locker2_check_available,id_tuplenil => false | id_locker2_check_available,id_append => false | id_locker2_check_available,id_0 => false | id_locker2_check_available,id_case8 => false | id_not,id_or => true | id_not,id_gen_tag => true | id_not,id_record_new => true | id_not,id_a => true | id_not,id_locker2_release_lock => true | id_not,id_eqt => true | id_not,id_istops => true | id_not,id_locker2_map_add_pending => true | id_not,id_release => true | id_not,id_locker2_obtainable => true | id_not,id_imp => true | id_not,id_stack => true | id_not,id_locker2_map_promote_pending => true | id_not,id_locker => true | id_not,id_case4 => true | id_not,id_int => true | id_not,id_push => true | id_not,id_locker2_add_pending => true | id_not,id_true => true | id_not,id_locker2_check_availables => true | id_not,id_F => true | id_not,id_eqs => true | id_not,id_record_update => true | id_not,id_false => true | id_not,id_gen_modtageq => true | id_not,id_undefined => true | id_not,id_nocalls => true | id_not,id_locker2_remove_pending => true | id_not,id_resource => true | id_not,id_case6 => true | id_not,id_if => true | id_not,id_pops => true | id_not,id_locker2_map_claim_lock => true | id_not,id_ok => true | id_not,id_case5 => true | id_not,id_tuple => true | id_not,id_member => true | id_not,id_s => true | id_not,id_delete => true | id_not,id_T => true | id_not,id_case9 => true | id_not,id_record_extract => true | id_not,id_excl => true | id_not,id_case2 => true | id_not,id_nil => true | id_not,id_eqc => true | id_not,id_case0 => true | id_not,id_request => true | id_not,id_locker2_check_available => true | id_not,id_not => true | id_not,id_pushs => false | id_not,id_locker2_promote_pending => false | id_not,id_mcrlrecord => false | id_not,id_locker2_obtainables => false | id_not,id_cons => false | id_not,id_push1 => false | id_not,id_case1 => false | id_not,id_element => false | id_not,id_locker2_adduniq => false | id_not,id_and => false | id_not,id_empty => false | id_not,id_record_updates => false | id_not,id_lock => false | id_not,id_excllock => false | id_not,id_pid => false | id_not,id_calls => false | id_not,id_subtract => false | id_not,id_tag => false | id_not,id_equal => false | id_not,id_eq => false | id_not,id_tops => false | id_not,id_locker2_claim_lock => false | id_not,id_pending => false | id_not,id_andt => false | id_not,id_tuplenil => false | id_not,id_append => false | id_not,id_0 => false | id_not,id_case8 => false | id_pushs,id_or => true | id_pushs,id_gen_tag => true | id_pushs,id_record_new => true | id_pushs,id_a => true | id_pushs,id_locker2_release_lock => true | id_pushs,id_eqt => true | id_pushs,id_istops => true | id_pushs,id_locker2_map_add_pending => true | id_pushs,id_release => true | id_pushs,id_locker2_obtainable => true | id_pushs,id_imp => true | id_pushs,id_stack => true | id_pushs,id_locker2_map_promote_pending => true | id_pushs,id_locker => true | id_pushs,id_case4 => true | id_pushs,id_int => true | id_pushs,id_push => true | id_pushs,id_locker2_add_pending => true | id_pushs,id_true => true | id_pushs,id_locker2_check_availables => true | id_pushs,id_F => true | id_pushs,id_eqs => true | id_pushs,id_record_update => true | id_pushs,id_false => true | id_pushs,id_gen_modtageq => true | id_pushs,id_undefined => true | id_pushs,id_nocalls => true | id_pushs,id_locker2_remove_pending => true | id_pushs,id_resource => true | id_pushs,id_case6 => true | id_pushs,id_if => true | id_pushs,id_pops => true | id_pushs,id_locker2_map_claim_lock => true | id_pushs,id_ok => true | id_pushs,id_case5 => true | id_pushs,id_tuple => true | id_pushs,id_member => true | id_pushs,id_s => true | id_pushs,id_delete => true | id_pushs,id_T => true | id_pushs,id_case9 => true | id_pushs,id_record_extract => true | id_pushs,id_excl => true | id_pushs,id_case2 => true | id_pushs,id_nil => true | id_pushs,id_eqc => true | id_pushs,id_case0 => true | id_pushs,id_request => true | id_pushs,id_locker2_check_available => true | id_pushs,id_not => true | id_pushs,id_pushs => true | id_pushs,id_locker2_promote_pending => false | id_pushs,id_mcrlrecord => false | id_pushs,id_locker2_obtainables => false | id_pushs,id_cons => false | id_pushs,id_push1 => false | id_pushs,id_case1 => false | id_pushs,id_element => false | id_pushs,id_locker2_adduniq => false | id_pushs,id_and => false | id_pushs,id_empty => false | id_pushs,id_record_updates => false | id_pushs,id_lock => false | id_pushs,id_excllock => false | id_pushs,id_pid => false | id_pushs,id_calls => false | id_pushs,id_subtract => false | id_pushs,id_tag => false | id_pushs,id_equal => false | id_pushs,id_eq => false | id_pushs,id_tops => false | id_pushs,id_locker2_claim_lock => false | id_pushs,id_pending => false | id_pushs,id_andt => false | id_pushs,id_tuplenil => false | id_pushs,id_append => false | id_pushs,id_0 => false | id_pushs,id_case8 => false | id_locker2_promote_pending,id_or => true | id_locker2_promote_pending,id_gen_tag => true | id_locker2_promote_pending,id_record_new => true | id_locker2_promote_pending,id_a => true | id_locker2_promote_pending,id_locker2_release_lock => true | id_locker2_promote_pending,id_eqt => true | id_locker2_promote_pending,id_istops => true | id_locker2_promote_pending,id_locker2_map_add_pending => true | id_locker2_promote_pending,id_release => true | id_locker2_promote_pending,id_locker2_obtainable => true | id_locker2_promote_pending,id_imp => true | id_locker2_promote_pending,id_stack => true | id_locker2_promote_pending,id_locker2_map_promote_pending => true | id_locker2_promote_pending,id_locker => true | id_locker2_promote_pending,id_case4 => true | id_locker2_promote_pending,id_int => true | id_locker2_promote_pending,id_push => true | id_locker2_promote_pending,id_locker2_add_pending => true | id_locker2_promote_pending,id_true => true | id_locker2_promote_pending,id_locker2_check_availables => true | id_locker2_promote_pending,id_F => true | id_locker2_promote_pending,id_eqs => true | id_locker2_promote_pending,id_record_update => true | id_locker2_promote_pending,id_false => true | id_locker2_promote_pending,id_gen_modtageq => true | id_locker2_promote_pending,id_undefined => true | id_locker2_promote_pending,id_nocalls => true | id_locker2_promote_pending,id_locker2_remove_pending => true | id_locker2_promote_pending,id_resource => true | id_locker2_promote_pending,id_case6 => true | id_locker2_promote_pending,id_if => true | id_locker2_promote_pending,id_pops => true | id_locker2_promote_pending,id_locker2_map_claim_lock => true | id_locker2_promote_pending,id_ok => true | id_locker2_promote_pending,id_case5 => true | id_locker2_promote_pending,id_tuple => true | id_locker2_promote_pending,id_member => true | id_locker2_promote_pending,id_s => true | id_locker2_promote_pending,id_delete => true | id_locker2_promote_pending,id_T => true | id_locker2_promote_pending,id_case9 => true | id_locker2_promote_pending,id_record_extract => true | id_locker2_promote_pending,id_excl => true | id_locker2_promote_pending,id_case2 => true | id_locker2_promote_pending,id_nil => true | id_locker2_promote_pending,id_eqc => true | id_locker2_promote_pending,id_case0 => true | id_locker2_promote_pending,id_request => true | id_locker2_promote_pending,id_locker2_check_available => true | id_locker2_promote_pending,id_not => true | id_locker2_promote_pending,id_pushs => true | id_locker2_promote_pending,id_locker2_promote_pending => true | id_locker2_promote_pending,id_mcrlrecord => false | id_locker2_promote_pending,id_locker2_obtainables => false | id_locker2_promote_pending,id_cons => false | id_locker2_promote_pending,id_push1 => false | id_locker2_promote_pending,id_case1 => false | id_locker2_promote_pending,id_element => false | id_locker2_promote_pending,id_locker2_adduniq => false | id_locker2_promote_pending,id_and => false | id_locker2_promote_pending,id_empty => false | id_locker2_promote_pending,id_record_updates => false | id_locker2_promote_pending,id_lock => false | id_locker2_promote_pending,id_excllock => false | id_locker2_promote_pending,id_pid => false | id_locker2_promote_pending,id_calls => false | id_locker2_promote_pending,id_subtract => false | id_locker2_promote_pending,id_tag => false | id_locker2_promote_pending,id_equal => false | id_locker2_promote_pending,id_eq => false | id_locker2_promote_pending,id_tops => false | id_locker2_promote_pending,id_locker2_claim_lock => false | id_locker2_promote_pending,id_pending => false | id_locker2_promote_pending,id_andt => false | id_locker2_promote_pending,id_tuplenil => false | id_locker2_promote_pending,id_append => false | id_locker2_promote_pending,id_0 => false | id_locker2_promote_pending,id_case8 => false | id_mcrlrecord,id_or => true | id_mcrlrecord,id_gen_tag => true | id_mcrlrecord,id_record_new => true | id_mcrlrecord,id_a => true | id_mcrlrecord,id_locker2_release_lock => true | id_mcrlrecord,id_eqt => true | id_mcrlrecord,id_istops => true | id_mcrlrecord,id_locker2_map_add_pending => true | id_mcrlrecord,id_release => true | id_mcrlrecord,id_locker2_obtainable => true | id_mcrlrecord,id_imp => true | id_mcrlrecord,id_stack => true | id_mcrlrecord,id_locker2_map_promote_pending => true | id_mcrlrecord,id_locker => true | id_mcrlrecord,id_case4 => true | id_mcrlrecord,id_int => true | id_mcrlrecord,id_push => true | id_mcrlrecord,id_locker2_add_pending => true | id_mcrlrecord,id_true => true | id_mcrlrecord,id_locker2_check_availables => true | id_mcrlrecord,id_F => true | id_mcrlrecord,id_eqs => true | id_mcrlrecord,id_record_update => true | id_mcrlrecord,id_false => true | id_mcrlrecord,id_gen_modtageq => true | id_mcrlrecord,id_undefined => true | id_mcrlrecord,id_nocalls => true | id_mcrlrecord,id_locker2_remove_pending => true | id_mcrlrecord,id_resource => true | id_mcrlrecord,id_case6 => true | id_mcrlrecord,id_if => true | id_mcrlrecord,id_pops => true | id_mcrlrecord,id_locker2_map_claim_lock => true | id_mcrlrecord,id_ok => true | id_mcrlrecord,id_case5 => true | id_mcrlrecord,id_tuple => true | id_mcrlrecord,id_member => true | id_mcrlrecord,id_s => true | id_mcrlrecord,id_delete => true | id_mcrlrecord,id_T => true | id_mcrlrecord,id_case9 => true | id_mcrlrecord,id_record_extract => true | id_mcrlrecord,id_excl => true | id_mcrlrecord,id_case2 => true | id_mcrlrecord,id_nil => true | id_mcrlrecord,id_eqc => true | id_mcrlrecord,id_case0 => true | id_mcrlrecord,id_request => true | id_mcrlrecord,id_locker2_check_available => true | id_mcrlrecord,id_not => true | id_mcrlrecord,id_pushs => true | id_mcrlrecord,id_locker2_promote_pending => true | id_mcrlrecord,id_mcrlrecord => true | id_mcrlrecord,id_locker2_obtainables => false | id_mcrlrecord,id_cons => false | id_mcrlrecord,id_push1 => false | id_mcrlrecord,id_case1 => false | id_mcrlrecord,id_element => false | id_mcrlrecord,id_locker2_adduniq => false | id_mcrlrecord,id_and => false | id_mcrlrecord,id_empty => false | id_mcrlrecord,id_record_updates => false | id_mcrlrecord,id_lock => false | id_mcrlrecord,id_excllock => false | id_mcrlrecord,id_pid => false | id_mcrlrecord,id_calls => false | id_mcrlrecord,id_subtract => false | id_mcrlrecord,id_tag => false | id_mcrlrecord,id_equal => false | id_mcrlrecord,id_eq => false | id_mcrlrecord,id_tops => false | id_mcrlrecord,id_locker2_claim_lock => false | id_mcrlrecord,id_pending => false | id_mcrlrecord,id_andt => false | id_mcrlrecord,id_tuplenil => false | id_mcrlrecord,id_append => false | id_mcrlrecord,id_0 => false | id_mcrlrecord,id_case8 => false | id_locker2_obtainables,id_or => true | id_locker2_obtainables,id_gen_tag => true | id_locker2_obtainables,id_record_new => true | id_locker2_obtainables,id_a => true | id_locker2_obtainables,id_locker2_release_lock => true | id_locker2_obtainables,id_eqt => true | id_locker2_obtainables,id_istops => true | id_locker2_obtainables,id_locker2_map_add_pending => true | id_locker2_obtainables,id_release => true | id_locker2_obtainables,id_locker2_obtainable => true | id_locker2_obtainables,id_imp => true | id_locker2_obtainables,id_stack => true | id_locker2_obtainables,id_locker2_map_promote_pending => true | id_locker2_obtainables,id_locker => true | id_locker2_obtainables,id_case4 => true | id_locker2_obtainables,id_int => true | id_locker2_obtainables,id_push => true | id_locker2_obtainables,id_locker2_add_pending => true | id_locker2_obtainables,id_true => true | id_locker2_obtainables,id_locker2_check_availables => true | id_locker2_obtainables,id_F => true | id_locker2_obtainables,id_eqs => true | id_locker2_obtainables,id_record_update => true | id_locker2_obtainables,id_false => true | id_locker2_obtainables,id_gen_modtageq => true | id_locker2_obtainables,id_undefined => true | id_locker2_obtainables,id_nocalls => true | id_locker2_obtainables,id_locker2_remove_pending => true | id_locker2_obtainables,id_resource => true | id_locker2_obtainables,id_case6 => true | id_locker2_obtainables,id_if => true | id_locker2_obtainables,id_pops => true | id_locker2_obtainables,id_locker2_map_claim_lock => true | id_locker2_obtainables,id_ok => true | id_locker2_obtainables,id_case5 => true | id_locker2_obtainables,id_tuple => true | id_locker2_obtainables,id_member => true | id_locker2_obtainables,id_s => true | id_locker2_obtainables,id_delete => true | id_locker2_obtainables,id_T => true | id_locker2_obtainables,id_case9 => true | id_locker2_obtainables,id_record_extract => true | id_locker2_obtainables,id_excl => true | id_locker2_obtainables,id_case2 => true | id_locker2_obtainables,id_nil => true | id_locker2_obtainables,id_eqc => true | id_locker2_obtainables,id_case0 => true | id_locker2_obtainables,id_request => true | id_locker2_obtainables,id_locker2_check_available => true | id_locker2_obtainables,id_not => true | id_locker2_obtainables,id_pushs => true | id_locker2_obtainables,id_locker2_promote_pending => true | id_locker2_obtainables,id_mcrlrecord => true | id_locker2_obtainables,id_locker2_obtainables => true | id_locker2_obtainables,id_cons => false | id_locker2_obtainables,id_push1 => false | id_locker2_obtainables,id_case1 => false | id_locker2_obtainables,id_element => false | id_locker2_obtainables,id_locker2_adduniq => false | id_locker2_obtainables,id_and => false | id_locker2_obtainables,id_empty => false | id_locker2_obtainables,id_record_updates => false | id_locker2_obtainables,id_lock => false | id_locker2_obtainables,id_excllock => false | id_locker2_obtainables,id_pid => false | id_locker2_obtainables,id_calls => false | id_locker2_obtainables,id_subtract => false | id_locker2_obtainables,id_tag => false | id_locker2_obtainables,id_equal => false | id_locker2_obtainables,id_eq => false | id_locker2_obtainables,id_tops => false | id_locker2_obtainables,id_locker2_claim_lock => false | id_locker2_obtainables,id_pending => false | id_locker2_obtainables,id_andt => false | id_locker2_obtainables,id_tuplenil => false | id_locker2_obtainables,id_append => false | id_locker2_obtainables,id_0 => false | id_locker2_obtainables,id_case8 => false | id_cons,id_or => true | id_cons,id_gen_tag => true | id_cons,id_record_new => true | id_cons,id_a => true | id_cons,id_locker2_release_lock => true | id_cons,id_eqt => true | id_cons,id_istops => true | id_cons,id_locker2_map_add_pending => true | id_cons,id_release => true | id_cons,id_locker2_obtainable => true | id_cons,id_imp => true | id_cons,id_stack => true | id_cons,id_locker2_map_promote_pending => true | id_cons,id_locker => true | id_cons,id_case4 => true | id_cons,id_int => true | id_cons,id_push => true | id_cons,id_locker2_add_pending => true | id_cons,id_true => true | id_cons,id_locker2_check_availables => true | id_cons,id_F => true | id_cons,id_eqs => true | id_cons,id_record_update => true | id_cons,id_false => true | id_cons,id_gen_modtageq => true | id_cons,id_undefined => true | id_cons,id_nocalls => true | id_cons,id_locker2_remove_pending => true | id_cons,id_resource => true | id_cons,id_case6 => true | id_cons,id_if => true | id_cons,id_pops => true | id_cons,id_locker2_map_claim_lock => true | id_cons,id_ok => true | id_cons,id_case5 => true | id_cons,id_tuple => true | id_cons,id_member => true | id_cons,id_s => true | id_cons,id_delete => true | id_cons,id_T => true | id_cons,id_case9 => true | id_cons,id_record_extract => true | id_cons,id_excl => true | id_cons,id_case2 => true | id_cons,id_nil => true | id_cons,id_eqc => true | id_cons,id_case0 => true | id_cons,id_request => true | id_cons,id_locker2_check_available => true | id_cons,id_not => true | id_cons,id_pushs => true | id_cons,id_locker2_promote_pending => true | id_cons,id_mcrlrecord => true | id_cons,id_locker2_obtainables => true | id_cons,id_cons => true | id_cons,id_push1 => false | id_cons,id_case1 => false | id_cons,id_element => false | id_cons,id_locker2_adduniq => false | id_cons,id_and => false | id_cons,id_empty => false | id_cons,id_record_updates => false | id_cons,id_lock => false | id_cons,id_excllock => false | id_cons,id_pid => false | id_cons,id_calls => false | id_cons,id_subtract => false | id_cons,id_tag => false | id_cons,id_equal => false | id_cons,id_eq => false | id_cons,id_tops => false | id_cons,id_locker2_claim_lock => false | id_cons,id_pending => false | id_cons,id_andt => false | id_cons,id_tuplenil => false | id_cons,id_append => false | id_cons,id_0 => false | id_cons,id_case8 => false | id_push1,id_or => true | id_push1,id_gen_tag => true | id_push1,id_record_new => true | id_push1,id_a => true | id_push1,id_locker2_release_lock => true | id_push1,id_eqt => true | id_push1,id_istops => true | id_push1,id_locker2_map_add_pending => true | id_push1,id_release => true | id_push1,id_locker2_obtainable => true | id_push1,id_imp => true | id_push1,id_stack => true | id_push1,id_locker2_map_promote_pending => true | id_push1,id_locker => true | id_push1,id_case4 => true | id_push1,id_int => true | id_push1,id_push => true | id_push1,id_locker2_add_pending => true | id_push1,id_true => true | id_push1,id_locker2_check_availables => true | id_push1,id_F => true | id_push1,id_eqs => true | id_push1,id_record_update => true | id_push1,id_false => true | id_push1,id_gen_modtageq => true | id_push1,id_undefined => true | id_push1,id_nocalls => true | id_push1,id_locker2_remove_pending => true | id_push1,id_resource => true | id_push1,id_case6 => true | id_push1,id_if => true | id_push1,id_pops => true | id_push1,id_locker2_map_claim_lock => true | id_push1,id_ok => true | id_push1,id_case5 => true | id_push1,id_tuple => true | id_push1,id_member => true | id_push1,id_s => true | id_push1,id_delete => true | id_push1,id_T => true | id_push1,id_case9 => true | id_push1,id_record_extract => true | id_push1,id_excl => true | id_push1,id_case2 => true | id_push1,id_nil => true | id_push1,id_eqc => true | id_push1,id_case0 => true | id_push1,id_request => true | id_push1,id_locker2_check_available => true | id_push1,id_not => true | id_push1,id_pushs => true | id_push1,id_locker2_promote_pending => true | id_push1,id_mcrlrecord => true | id_push1,id_locker2_obtainables => true | id_push1,id_cons => true | id_push1,id_push1 => true | id_push1,id_case1 => false | id_push1,id_element => false | id_push1,id_locker2_adduniq => false | id_push1,id_and => false | id_push1,id_empty => false | id_push1,id_record_updates => false | id_push1,id_lock => false | id_push1,id_excllock => false | id_push1,id_pid => false | id_push1,id_calls => false | id_push1,id_subtract => false | id_push1,id_tag => false | id_push1,id_equal => false | id_push1,id_eq => false | id_push1,id_tops => false | id_push1,id_locker2_claim_lock => false | id_push1,id_pending => false | id_push1,id_andt => false | id_push1,id_tuplenil => false | id_push1,id_append => false | id_push1,id_0 => false | id_push1,id_case8 => false | id_case1,id_or => true | id_case1,id_gen_tag => true | id_case1,id_record_new => true | id_case1,id_a => true | id_case1,id_locker2_release_lock => true | id_case1,id_eqt => true | id_case1,id_istops => true | id_case1,id_locker2_map_add_pending => true | id_case1,id_release => true | id_case1,id_locker2_obtainable => true | id_case1,id_imp => true | id_case1,id_stack => true | id_case1,id_locker2_map_promote_pending => true | id_case1,id_locker => true | id_case1,id_case4 => true | id_case1,id_int => true | id_case1,id_push => true | id_case1,id_locker2_add_pending => true | id_case1,id_true => true | id_case1,id_locker2_check_availables => true | id_case1,id_F => true | id_case1,id_eqs => true | id_case1,id_record_update => true | id_case1,id_false => true | id_case1,id_gen_modtageq => true | id_case1,id_undefined => true | id_case1,id_nocalls => true | id_case1,id_locker2_remove_pending => true | id_case1,id_resource => true | id_case1,id_case6 => true | id_case1,id_if => true | id_case1,id_pops => true | id_case1,id_locker2_map_claim_lock => true | id_case1,id_ok => true | id_case1,id_case5 => true | id_case1,id_tuple => true | id_case1,id_member => true | id_case1,id_s => true | id_case1,id_delete => true | id_case1,id_T => true | id_case1,id_case9 => true | id_case1,id_record_extract => true | id_case1,id_excl => true | id_case1,id_case2 => true | id_case1,id_nil => true | id_case1,id_eqc => true | id_case1,id_case0 => true | id_case1,id_request => true | id_case1,id_locker2_check_available => true | id_case1,id_not => true | id_case1,id_pushs => true | id_case1,id_locker2_promote_pending => true | id_case1,id_mcrlrecord => true | id_case1,id_locker2_obtainables => true | id_case1,id_cons => true | id_case1,id_push1 => true | id_case1,id_case1 => true | id_case1,id_element => false | id_case1,id_locker2_adduniq => false | id_case1,id_and => false | id_case1,id_empty => false | id_case1,id_record_updates => false | id_case1,id_lock => false | id_case1,id_excllock => false | id_case1,id_pid => false | id_case1,id_calls => false | id_case1,id_subtract => false | id_case1,id_tag => false | id_case1,id_equal => false | id_case1,id_eq => false | id_case1,id_tops => false | id_case1,id_locker2_claim_lock => false | id_case1,id_pending => false | id_case1,id_andt => false | id_case1,id_tuplenil => false | id_case1,id_append => false | id_case1,id_0 => false | id_case1,id_case8 => false | id_element,id_or => true | id_element,id_gen_tag => true | id_element,id_record_new => true | id_element,id_a => true | id_element,id_locker2_release_lock => true | id_element,id_eqt => true | id_element,id_istops => true | id_element,id_locker2_map_add_pending => true | id_element,id_release => true | id_element,id_locker2_obtainable => true | id_element,id_imp => true | id_element,id_stack => true | id_element,id_locker2_map_promote_pending => true | id_element,id_locker => true | id_element,id_case4 => true | id_element,id_int => true | id_element,id_push => true | id_element,id_locker2_add_pending => true | id_element,id_true => true | id_element,id_locker2_check_availables => true | id_element,id_F => true | id_element,id_eqs => true | id_element,id_record_update => true | id_element,id_false => true | id_element,id_gen_modtageq => true | id_element,id_undefined => true | id_element,id_nocalls => true | id_element,id_locker2_remove_pending => true | id_element,id_resource => true | id_element,id_case6 => true | id_element,id_if => true | id_element,id_pops => true | id_element,id_locker2_map_claim_lock => true | id_element,id_ok => true | id_element,id_case5 => true | id_element,id_tuple => true | id_element,id_member => true | id_element,id_s => true | id_element,id_delete => true | id_element,id_T => true | id_element,id_case9 => true | id_element,id_record_extract => true | id_element,id_excl => true | id_element,id_case2 => true | id_element,id_nil => true | id_element,id_eqc => true | id_element,id_case0 => true | id_element,id_request => true | id_element,id_locker2_check_available => true | id_element,id_not => true | id_element,id_pushs => true | id_element,id_locker2_promote_pending => true | id_element,id_mcrlrecord => true | id_element,id_locker2_obtainables => true | id_element,id_cons => true | id_element,id_push1 => true | id_element,id_case1 => true | id_element,id_element => true | id_element,id_locker2_adduniq => false | id_element,id_and => false | id_element,id_empty => false | id_element,id_record_updates => false | id_element,id_lock => false | id_element,id_excllock => false | id_element,id_pid => false | id_element,id_calls => false | id_element,id_subtract => false | id_element,id_tag => false | id_element,id_equal => false | id_element,id_eq => false | id_element,id_tops => false | id_element,id_locker2_claim_lock => false | id_element,id_pending => false | id_element,id_andt => false | id_element,id_tuplenil => false | id_element,id_append => false | id_element,id_0 => false | id_element,id_case8 => false | id_locker2_adduniq,id_or => true | id_locker2_adduniq,id_gen_tag => true | id_locker2_adduniq,id_record_new => true | id_locker2_adduniq,id_a => true | id_locker2_adduniq,id_locker2_release_lock => true | id_locker2_adduniq,id_eqt => true | id_locker2_adduniq,id_istops => true | id_locker2_adduniq,id_locker2_map_add_pending => true | id_locker2_adduniq,id_release => true | id_locker2_adduniq,id_locker2_obtainable => true | id_locker2_adduniq,id_imp => true | id_locker2_adduniq,id_stack => true | id_locker2_adduniq,id_locker2_map_promote_pending => true | id_locker2_adduniq,id_locker => true | id_locker2_adduniq,id_case4 => true | id_locker2_adduniq,id_int => true | id_locker2_adduniq,id_push => true | id_locker2_adduniq,id_locker2_add_pending => true | id_locker2_adduniq,id_true => true | id_locker2_adduniq,id_locker2_check_availables => true | id_locker2_adduniq,id_F => true | id_locker2_adduniq,id_eqs => true | id_locker2_adduniq,id_record_update => true | id_locker2_adduniq,id_false => true | id_locker2_adduniq,id_gen_modtageq => true | id_locker2_adduniq,id_undefined => true | id_locker2_adduniq,id_nocalls => true | id_locker2_adduniq,id_locker2_remove_pending => true | id_locker2_adduniq,id_resource => true | id_locker2_adduniq,id_case6 => true | id_locker2_adduniq,id_if => true | id_locker2_adduniq,id_pops => true | id_locker2_adduniq,id_locker2_map_claim_lock => true | id_locker2_adduniq,id_ok => true | id_locker2_adduniq,id_case5 => true | id_locker2_adduniq,id_tuple => true | id_locker2_adduniq,id_member => true | id_locker2_adduniq,id_s => true | id_locker2_adduniq,id_delete => true | id_locker2_adduniq,id_T => true | id_locker2_adduniq,id_case9 => true | id_locker2_adduniq,id_record_extract => true | id_locker2_adduniq,id_excl => true | id_locker2_adduniq,id_case2 => true | id_locker2_adduniq,id_nil => true | id_locker2_adduniq,id_eqc => true | id_locker2_adduniq,id_case0 => true | id_locker2_adduniq,id_request => true | id_locker2_adduniq,id_locker2_check_available => true | id_locker2_adduniq,id_not => true | id_locker2_adduniq,id_pushs => true | id_locker2_adduniq,id_locker2_promote_pending => true | id_locker2_adduniq,id_mcrlrecord => true | id_locker2_adduniq,id_locker2_obtainables => true | id_locker2_adduniq,id_cons => true | id_locker2_adduniq,id_push1 => true | id_locker2_adduniq,id_case1 => true | id_locker2_adduniq,id_element => true | id_locker2_adduniq,id_locker2_adduniq => true | id_locker2_adduniq,id_and => false | id_locker2_adduniq,id_empty => false | id_locker2_adduniq,id_record_updates => false | id_locker2_adduniq,id_lock => false | id_locker2_adduniq,id_excllock => false | id_locker2_adduniq,id_pid => false | id_locker2_adduniq,id_calls => false | id_locker2_adduniq,id_subtract => false | id_locker2_adduniq,id_tag => false | id_locker2_adduniq,id_equal => false | id_locker2_adduniq,id_eq => false | id_locker2_adduniq,id_tops => false | id_locker2_adduniq,id_locker2_claim_lock => false | id_locker2_adduniq,id_pending => false | id_locker2_adduniq,id_andt => false | id_locker2_adduniq,id_tuplenil => false | id_locker2_adduniq,id_append => false | id_locker2_adduniq,id_0 => false | id_locker2_adduniq,id_case8 => false | id_and,id_or => true | id_and,id_gen_tag => true | id_and,id_record_new => true | id_and,id_a => true | id_and,id_locker2_release_lock => true | id_and,id_eqt => true | id_and,id_istops => true | id_and,id_locker2_map_add_pending => true | id_and,id_release => true | id_and,id_locker2_obtainable => true | id_and,id_imp => true | id_and,id_stack => true | id_and,id_locker2_map_promote_pending => true | id_and,id_locker => true | id_and,id_case4 => true | id_and,id_int => true | id_and,id_push => true | id_and,id_locker2_add_pending => true | id_and,id_true => true | id_and,id_locker2_check_availables => true | id_and,id_F => true | id_and,id_eqs => true | id_and,id_record_update => true | id_and,id_false => true | id_and,id_gen_modtageq => true | id_and,id_undefined => true | id_and,id_nocalls => true | id_and,id_locker2_remove_pending => true | id_and,id_resource => true | id_and,id_case6 => true | id_and,id_if => true | id_and,id_pops => true | id_and,id_locker2_map_claim_lock => true | id_and,id_ok => true | id_and,id_case5 => true | id_and,id_tuple => true | id_and,id_member => true | id_and,id_s => true | id_and,id_delete => true | id_and,id_T => true | id_and,id_case9 => true | id_and,id_record_extract => true | id_and,id_excl => true | id_and,id_case2 => true | id_and,id_nil => true | id_and,id_eqc => true | id_and,id_case0 => true | id_and,id_request => true | id_and,id_locker2_check_available => true | id_and,id_not => true | id_and,id_pushs => true | id_and,id_locker2_promote_pending => true | id_and,id_mcrlrecord => true | id_and,id_locker2_obtainables => true | id_and,id_cons => true | id_and,id_push1 => true | id_and,id_case1 => true | id_and,id_element => true | id_and,id_locker2_adduniq => true | id_and,id_and => true | id_and,id_empty => false | id_and,id_record_updates => false | id_and,id_lock => false | id_and,id_excllock => false | id_and,id_pid => false | id_and,id_calls => false | id_and,id_subtract => false | id_and,id_tag => false | id_and,id_equal => false | id_and,id_eq => false | id_and,id_tops => false | id_and,id_locker2_claim_lock => false | id_and,id_pending => false | id_and,id_andt => false | id_and,id_tuplenil => false | id_and,id_append => false | id_and,id_0 => false | id_and,id_case8 => false | id_empty,id_or => true | id_empty,id_gen_tag => true | id_empty,id_record_new => true | id_empty,id_a => true | id_empty,id_locker2_release_lock => true | id_empty,id_eqt => true | id_empty,id_istops => true | id_empty,id_locker2_map_add_pending => true | id_empty,id_release => true | id_empty,id_locker2_obtainable => true | id_empty,id_imp => true | id_empty,id_stack => true | id_empty,id_locker2_map_promote_pending => true | id_empty,id_locker => true | id_empty,id_case4 => true | id_empty,id_int => true | id_empty,id_push => true | id_empty,id_locker2_add_pending => true | id_empty,id_true => true | id_empty,id_locker2_check_availables => true | id_empty,id_F => true | id_empty,id_eqs => true | id_empty,id_record_update => true | id_empty,id_false => true | id_empty,id_gen_modtageq => true | id_empty,id_undefined => true | id_empty,id_nocalls => true | id_empty,id_locker2_remove_pending => true | id_empty,id_resource => true | id_empty,id_case6 => true | id_empty,id_if => true | id_empty,id_pops => true | id_empty,id_locker2_map_claim_lock => true | id_empty,id_ok => true | id_empty,id_case5 => true | id_empty,id_tuple => true | id_empty,id_member => true | id_empty,id_s => true | id_empty,id_delete => true | id_empty,id_T => true | id_empty,id_case9 => true | id_empty,id_record_extract => true | id_empty,id_excl => true | id_empty,id_case2 => true | id_empty,id_nil => true | id_empty,id_eqc => true | id_empty,id_case0 => true | id_empty,id_request => true | id_empty,id_locker2_check_available => true | id_empty,id_not => true | id_empty,id_pushs => true | id_empty,id_locker2_promote_pending => true | id_empty,id_mcrlrecord => true | id_empty,id_locker2_obtainables => true | id_empty,id_cons => true | id_empty,id_push1 => true | id_empty,id_case1 => true | id_empty,id_element => true | id_empty,id_locker2_adduniq => true | id_empty,id_and => true | id_empty,id_empty => true | id_empty,id_record_updates => false | id_empty,id_lock => false | id_empty,id_excllock => false | id_empty,id_pid => false | id_empty,id_calls => false | id_empty,id_subtract => false | id_empty,id_tag => false | id_empty,id_equal => false | id_empty,id_eq => false | id_empty,id_tops => false | id_empty,id_locker2_claim_lock => false | id_empty,id_pending => false | id_empty,id_andt => false | id_empty,id_tuplenil => false | id_empty,id_append => false | id_empty,id_0 => false | id_empty,id_case8 => false | id_record_updates,id_or => true | id_record_updates,id_gen_tag => true | id_record_updates,id_record_new => true | id_record_updates,id_a => true | id_record_updates,id_locker2_release_lock => true | id_record_updates,id_eqt => true | id_record_updates,id_istops => true | id_record_updates,id_locker2_map_add_pending => true | id_record_updates,id_release => true | id_record_updates,id_locker2_obtainable => true | id_record_updates,id_imp => true | id_record_updates,id_stack => true | id_record_updates,id_locker2_map_promote_pending => true | id_record_updates,id_locker => true | id_record_updates,id_case4 => true | id_record_updates,id_int => true | id_record_updates,id_push => true | id_record_updates,id_locker2_add_pending => true | id_record_updates,id_true => true | id_record_updates,id_locker2_check_availables => true | id_record_updates,id_F => true | id_record_updates,id_eqs => true | id_record_updates,id_record_update => true | id_record_updates,id_false => true | id_record_updates,id_gen_modtageq => true | id_record_updates,id_undefined => true | id_record_updates,id_nocalls => true | id_record_updates,id_locker2_remove_pending => true | id_record_updates,id_resource => true | id_record_updates,id_case6 => true | id_record_updates,id_if => true | id_record_updates,id_pops => true | id_record_updates,id_locker2_map_claim_lock => true | id_record_updates,id_ok => true | id_record_updates,id_case5 => true | id_record_updates,id_tuple => true | id_record_updates,id_member => true | id_record_updates,id_s => true | id_record_updates,id_delete => true | id_record_updates,id_T => true | id_record_updates,id_case9 => true | id_record_updates,id_record_extract => true | id_record_updates,id_excl => true | id_record_updates,id_case2 => true | id_record_updates,id_nil => true | id_record_updates,id_eqc => true | id_record_updates,id_case0 => true | id_record_updates,id_request => true | id_record_updates,id_locker2_check_available => true | id_record_updates,id_not => true | id_record_updates,id_pushs => true | id_record_updates,id_locker2_promote_pending => true | id_record_updates,id_mcrlrecord => true | id_record_updates,id_locker2_obtainables => true | id_record_updates,id_cons => true | id_record_updates,id_push1 => true | id_record_updates,id_case1 => true | id_record_updates,id_element => true | id_record_updates,id_locker2_adduniq => true | id_record_updates,id_and => true | id_record_updates,id_empty => true | id_record_updates,id_record_updates => true | id_record_updates,id_lock => false | id_record_updates,id_excllock => false | id_record_updates,id_pid => false | id_record_updates,id_calls => false | id_record_updates,id_subtract => false | id_record_updates,id_tag => false | id_record_updates,id_equal => false | id_record_updates,id_eq => false | id_record_updates,id_tops => false | id_record_updates,id_locker2_claim_lock => false | id_record_updates,id_pending => false | id_record_updates,id_andt => false | id_record_updates,id_tuplenil => false | id_record_updates,id_append => false | id_record_updates,id_0 => false | id_record_updates,id_case8 => false | id_lock,id_or => true | id_lock,id_gen_tag => true | id_lock,id_record_new => true | id_lock,id_a => true | id_lock,id_locker2_release_lock => true | id_lock,id_eqt => true | id_lock,id_istops => true | id_lock,id_locker2_map_add_pending => true | id_lock,id_release => true | id_lock,id_locker2_obtainable => true | id_lock,id_imp => true | id_lock,id_stack => true | id_lock,id_locker2_map_promote_pending => true | id_lock,id_locker => true | id_lock,id_case4 => true | id_lock,id_int => true | id_lock,id_push => true | id_lock,id_locker2_add_pending => true | id_lock,id_true => true | id_lock,id_locker2_check_availables => true | id_lock,id_F => true | id_lock,id_eqs => true | id_lock,id_record_update => true | id_lock,id_false => true | id_lock,id_gen_modtageq => true | id_lock,id_undefined => true | id_lock,id_nocalls => true | id_lock,id_locker2_remove_pending => true | id_lock,id_resource => true | id_lock,id_case6 => true | id_lock,id_if => true | id_lock,id_pops => true | id_lock,id_locker2_map_claim_lock => true | id_lock,id_ok => true | id_lock,id_case5 => true | id_lock,id_tuple => true | id_lock,id_member => true | id_lock,id_s => true | id_lock,id_delete => true | id_lock,id_T => true | id_lock,id_case9 => true | id_lock,id_record_extract => true | id_lock,id_excl => true | id_lock,id_case2 => true | id_lock,id_nil => true | id_lock,id_eqc => true | id_lock,id_case0 => true | id_lock,id_request => true | id_lock,id_locker2_check_available => true | id_lock,id_not => true | id_lock,id_pushs => true | id_lock,id_locker2_promote_pending => true | id_lock,id_mcrlrecord => true | id_lock,id_locker2_obtainables => true | id_lock,id_cons => true | id_lock,id_push1 => true | id_lock,id_case1 => true | id_lock,id_element => true | id_lock,id_locker2_adduniq => true | id_lock,id_and => true | id_lock,id_empty => true | id_lock,id_record_updates => true | id_lock,id_lock => true | id_lock,id_excllock => false | id_lock,id_pid => false | id_lock,id_calls => false | id_lock,id_subtract => false | id_lock,id_tag => false | id_lock,id_equal => false | id_lock,id_eq => false | id_lock,id_tops => false | id_lock,id_locker2_claim_lock => false | id_lock,id_pending => false | id_lock,id_andt => false | id_lock,id_tuplenil => false | id_lock,id_append => false | id_lock,id_0 => false | id_lock,id_case8 => false | id_excllock,id_or => true | id_excllock,id_gen_tag => true | id_excllock,id_record_new => true | id_excllock,id_a => true | id_excllock,id_locker2_release_lock => true | id_excllock,id_eqt => true | id_excllock,id_istops => true | id_excllock,id_locker2_map_add_pending => true | id_excllock,id_release => true | id_excllock,id_locker2_obtainable => true | id_excllock,id_imp => true | id_excllock,id_stack => true | id_excllock,id_locker2_map_promote_pending => true | id_excllock,id_locker => true | id_excllock,id_case4 => true | id_excllock,id_int => true | id_excllock,id_push => true | id_excllock,id_locker2_add_pending => true | id_excllock,id_true => true | id_excllock,id_locker2_check_availables => true | id_excllock,id_F => true | id_excllock,id_eqs => true | id_excllock,id_record_update => true | id_excllock,id_false => true | id_excllock,id_gen_modtageq => true | id_excllock,id_undefined => true | id_excllock,id_nocalls => true | id_excllock,id_locker2_remove_pending => true | id_excllock,id_resource => true | id_excllock,id_case6 => true | id_excllock,id_if => true | id_excllock,id_pops => true | id_excllock,id_locker2_map_claim_lock => true | id_excllock,id_ok => true | id_excllock,id_case5 => true | id_excllock,id_tuple => true | id_excllock,id_member => true | id_excllock,id_s => true | id_excllock,id_delete => true | id_excllock,id_T => true | id_excllock,id_case9 => true | id_excllock,id_record_extract => true | id_excllock,id_excl => true | id_excllock,id_case2 => true | id_excllock,id_nil => true | id_excllock,id_eqc => true | id_excllock,id_case0 => true | id_excllock,id_request => true | id_excllock,id_locker2_check_available => true | id_excllock,id_not => true | id_excllock,id_pushs => true | id_excllock,id_locker2_promote_pending => true | id_excllock,id_mcrlrecord => true | id_excllock,id_locker2_obtainables => true | id_excllock,id_cons => true | id_excllock,id_push1 => true | id_excllock,id_case1 => true | id_excllock,id_element => true | id_excllock,id_locker2_adduniq => true | id_excllock,id_and => true | id_excllock,id_empty => true | id_excllock,id_record_updates => true | id_excllock,id_lock => true | id_excllock,id_excllock => true | id_excllock,id_pid => false | id_excllock,id_calls => false | id_excllock,id_subtract => false | id_excllock,id_tag => false | id_excllock,id_equal => false | id_excllock,id_eq => false | id_excllock,id_tops => false | id_excllock,id_locker2_claim_lock => false | id_excllock,id_pending => false | id_excllock,id_andt => false | id_excllock,id_tuplenil => false | id_excllock,id_append => false | id_excllock,id_0 => false | id_excllock,id_case8 => false | id_pid,id_or => true | id_pid,id_gen_tag => true | id_pid,id_record_new => true | id_pid,id_a => true | id_pid,id_locker2_release_lock => true | id_pid,id_eqt => true | id_pid,id_istops => true | id_pid,id_locker2_map_add_pending => true | id_pid,id_release => true | id_pid,id_locker2_obtainable => true | id_pid,id_imp => true | id_pid,id_stack => true | id_pid,id_locker2_map_promote_pending => true | id_pid,id_locker => true | id_pid,id_case4 => true | id_pid,id_int => true | id_pid,id_push => true | id_pid,id_locker2_add_pending => true | id_pid,id_true => true | id_pid,id_locker2_check_availables => true | id_pid,id_F => true | id_pid,id_eqs => true | id_pid,id_record_update => true | id_pid,id_false => true | id_pid,id_gen_modtageq => true | id_pid,id_undefined => true | id_pid,id_nocalls => true | id_pid,id_locker2_remove_pending => true | id_pid,id_resource => true | id_pid,id_case6 => true | id_pid,id_if => true | id_pid,id_pops => true | id_pid,id_locker2_map_claim_lock => true | id_pid,id_ok => true | id_pid,id_case5 => true | id_pid,id_tuple => true | id_pid,id_member => true | id_pid,id_s => true | id_pid,id_delete => true | id_pid,id_T => true | id_pid,id_case9 => true | id_pid,id_record_extract => true | id_pid,id_excl => true | id_pid,id_case2 => true | id_pid,id_nil => true | id_pid,id_eqc => true | id_pid,id_case0 => true | id_pid,id_request => true | id_pid,id_locker2_check_available => true | id_pid,id_not => true | id_pid,id_pushs => true | id_pid,id_locker2_promote_pending => true | id_pid,id_mcrlrecord => true | id_pid,id_locker2_obtainables => true | id_pid,id_cons => true | id_pid,id_push1 => true | id_pid,id_case1 => true | id_pid,id_element => true | id_pid,id_locker2_adduniq => true | id_pid,id_and => true | id_pid,id_empty => true | id_pid,id_record_updates => true | id_pid,id_lock => true | id_pid,id_excllock => true | id_pid,id_pid => true | id_pid,id_calls => false | id_pid,id_subtract => false | id_pid,id_tag => false | id_pid,id_equal => false | id_pid,id_eq => false | id_pid,id_tops => false | id_pid,id_locker2_claim_lock => false | id_pid,id_pending => false | id_pid,id_andt => false | id_pid,id_tuplenil => false | id_pid,id_append => false | id_pid,id_0 => false | id_pid,id_case8 => false | id_calls,id_or => true | id_calls,id_gen_tag => true | id_calls,id_record_new => true | id_calls,id_a => true | id_calls,id_locker2_release_lock => true | id_calls,id_eqt => true | id_calls,id_istops => true | id_calls,id_locker2_map_add_pending => true | id_calls,id_release => true | id_calls,id_locker2_obtainable => true | id_calls,id_imp => true | id_calls,id_stack => true | id_calls,id_locker2_map_promote_pending => true | id_calls,id_locker => true | id_calls,id_case4 => true | id_calls,id_int => true | id_calls,id_push => true | id_calls,id_locker2_add_pending => true | id_calls,id_true => true | id_calls,id_locker2_check_availables => true | id_calls,id_F => true | id_calls,id_eqs => true | id_calls,id_record_update => true | id_calls,id_false => true | id_calls,id_gen_modtageq => true | id_calls,id_undefined => true | id_calls,id_nocalls => true | id_calls,id_locker2_remove_pending => true | id_calls,id_resource => true | id_calls,id_case6 => true | id_calls,id_if => true | id_calls,id_pops => true | id_calls,id_locker2_map_claim_lock => true | id_calls,id_ok => true | id_calls,id_case5 => true | id_calls,id_tuple => true | id_calls,id_member => true | id_calls,id_s => true | id_calls,id_delete => true | id_calls,id_T => true | id_calls,id_case9 => true | id_calls,id_record_extract => true | id_calls,id_excl => true | id_calls,id_case2 => true | id_calls,id_nil => true | id_calls,id_eqc => true | id_calls,id_case0 => true | id_calls,id_request => true | id_calls,id_locker2_check_available => true | id_calls,id_not => true | id_calls,id_pushs => true | id_calls,id_locker2_promote_pending => true | id_calls,id_mcrlrecord => true | id_calls,id_locker2_obtainables => true | id_calls,id_cons => true | id_calls,id_push1 => true | id_calls,id_case1 => true | id_calls,id_element => true | id_calls,id_locker2_adduniq => true | id_calls,id_and => true | id_calls,id_empty => true | id_calls,id_record_updates => true | id_calls,id_lock => true | id_calls,id_excllock => true | id_calls,id_pid => true | id_calls,id_calls => true | id_calls,id_subtract => false | id_calls,id_tag => false | id_calls,id_equal => false | id_calls,id_eq => false | id_calls,id_tops => false | id_calls,id_locker2_claim_lock => false | id_calls,id_pending => false | id_calls,id_andt => false | id_calls,id_tuplenil => false | id_calls,id_append => false | id_calls,id_0 => false | id_calls,id_case8 => false | id_subtract,id_or => true | id_subtract,id_gen_tag => true | id_subtract,id_record_new => true | id_subtract,id_a => true | id_subtract,id_locker2_release_lock => true | id_subtract,id_eqt => true | id_subtract,id_istops => true | id_subtract,id_locker2_map_add_pending => true | id_subtract,id_release => true | id_subtract,id_locker2_obtainable => true | id_subtract,id_imp => true | id_subtract,id_stack => true | id_subtract,id_locker2_map_promote_pending => true | id_subtract,id_locker => true | id_subtract,id_case4 => true | id_subtract,id_int => true | id_subtract,id_push => true | id_subtract,id_locker2_add_pending => true | id_subtract,id_true => true | id_subtract,id_locker2_check_availables => true | id_subtract,id_F => true | id_subtract,id_eqs => true | id_subtract,id_record_update => true | id_subtract,id_false => true | id_subtract,id_gen_modtageq => true | id_subtract,id_undefined => true | id_subtract,id_nocalls => true | id_subtract,id_locker2_remove_pending => true | id_subtract,id_resource => true | id_subtract,id_case6 => true | id_subtract,id_if => true | id_subtract,id_pops => true | id_subtract,id_locker2_map_claim_lock => true | id_subtract,id_ok => true | id_subtract,id_case5 => true | id_subtract,id_tuple => true | id_subtract,id_member => true | id_subtract,id_s => true | id_subtract,id_delete => true | id_subtract,id_T => true | id_subtract,id_case9 => true | id_subtract,id_record_extract => true | id_subtract,id_excl => true | id_subtract,id_case2 => true | id_subtract,id_nil => true | id_subtract,id_eqc => true | id_subtract,id_case0 => true | id_subtract,id_request => true | id_subtract,id_locker2_check_available => true | id_subtract,id_not => true | id_subtract,id_pushs => true | id_subtract,id_locker2_promote_pending => true | id_subtract,id_mcrlrecord => true | id_subtract,id_locker2_obtainables => true | id_subtract,id_cons => true | id_subtract,id_push1 => true | id_subtract,id_case1 => true | id_subtract,id_element => true | id_subtract,id_locker2_adduniq => true | id_subtract,id_and => true | id_subtract,id_empty => true | id_subtract,id_record_updates => true | id_subtract,id_lock => true | id_subtract,id_excllock => true | id_subtract,id_pid => true | id_subtract,id_calls => true | id_subtract,id_subtract => true | id_subtract,id_tag => false | id_subtract,id_equal => false | id_subtract,id_eq => false | id_subtract,id_tops => false | id_subtract,id_locker2_claim_lock => false | id_subtract,id_pending => false | id_subtract,id_andt => false | id_subtract,id_tuplenil => false | id_subtract,id_append => false | id_subtract,id_0 => false | id_subtract,id_case8 => false | id_tag,id_or => true | id_tag,id_gen_tag => true | id_tag,id_record_new => true | id_tag,id_a => true | id_tag,id_locker2_release_lock => true | id_tag,id_eqt => true | id_tag,id_istops => true | id_tag,id_locker2_map_add_pending => true | id_tag,id_release => true | id_tag,id_locker2_obtainable => true | id_tag,id_imp => true | id_tag,id_stack => true | id_tag,id_locker2_map_promote_pending => true | id_tag,id_locker => true | id_tag,id_case4 => true | id_tag,id_int => true | id_tag,id_push => true | id_tag,id_locker2_add_pending => true | id_tag,id_true => true | id_tag,id_locker2_check_availables => true | id_tag,id_F => true | id_tag,id_eqs => true | id_tag,id_record_update => true | id_tag,id_false => true | id_tag,id_gen_modtageq => true | id_tag,id_undefined => true | id_tag,id_nocalls => true | id_tag,id_locker2_remove_pending => true | id_tag,id_resource => true | id_tag,id_case6 => true | id_tag,id_if => true | id_tag,id_pops => true | id_tag,id_locker2_map_claim_lock => true | id_tag,id_ok => true | id_tag,id_case5 => true | id_tag,id_tuple => true | id_tag,id_member => true | id_tag,id_s => true | id_tag,id_delete => true | id_tag,id_T => true | id_tag,id_case9 => true | id_tag,id_record_extract => true | id_tag,id_excl => true | id_tag,id_case2 => true | id_tag,id_nil => true | id_tag,id_eqc => true | id_tag,id_case0 => true | id_tag,id_request => true | id_tag,id_locker2_check_available => true | id_tag,id_not => true | id_tag,id_pushs => true | id_tag,id_locker2_promote_pending => true | id_tag,id_mcrlrecord => true | id_tag,id_locker2_obtainables => true | id_tag,id_cons => true | id_tag,id_push1 => true | id_tag,id_case1 => true | id_tag,id_element => true | id_tag,id_locker2_adduniq => true | id_tag,id_and => true | id_tag,id_empty => true | id_tag,id_record_updates => true | id_tag,id_lock => true | id_tag,id_excllock => true | id_tag,id_pid => true | id_tag,id_calls => true | id_tag,id_subtract => true | id_tag,id_tag => true | id_tag,id_equal => false | id_tag,id_eq => false | id_tag,id_tops => false | id_tag,id_locker2_claim_lock => false | id_tag,id_pending => false | id_tag,id_andt => false | id_tag,id_tuplenil => false | id_tag,id_append => false | id_tag,id_0 => false | id_tag,id_case8 => false | id_equal,id_or => true | id_equal,id_gen_tag => true | id_equal,id_record_new => true | id_equal,id_a => true | id_equal,id_locker2_release_lock => true | id_equal,id_eqt => true | id_equal,id_istops => true | id_equal,id_locker2_map_add_pending => true | id_equal,id_release => true | id_equal,id_locker2_obtainable => true | id_equal,id_imp => true | id_equal,id_stack => true | id_equal,id_locker2_map_promote_pending => true | id_equal,id_locker => true | id_equal,id_case4 => true | id_equal,id_int => true | id_equal,id_push => true | id_equal,id_locker2_add_pending => true | id_equal,id_true => true | id_equal,id_locker2_check_availables => true | id_equal,id_F => true | id_equal,id_eqs => true | id_equal,id_record_update => true | id_equal,id_false => true | id_equal,id_gen_modtageq => true | id_equal,id_undefined => true | id_equal,id_nocalls => true | id_equal,id_locker2_remove_pending => true | id_equal,id_resource => true | id_equal,id_case6 => true | id_equal,id_if => true | id_equal,id_pops => true | id_equal,id_locker2_map_claim_lock => true | id_equal,id_ok => true | id_equal,id_case5 => true | id_equal,id_tuple => true | id_equal,id_member => true | id_equal,id_s => true | id_equal,id_delete => true | id_equal,id_T => true | id_equal,id_case9 => true | id_equal,id_record_extract => true | id_equal,id_excl => true | id_equal,id_case2 => true | id_equal,id_nil => true | id_equal,id_eqc => true | id_equal,id_case0 => true | id_equal,id_request => true | id_equal,id_locker2_check_available => true | id_equal,id_not => true | id_equal,id_pushs => true | id_equal,id_locker2_promote_pending => true | id_equal,id_mcrlrecord => true | id_equal,id_locker2_obtainables => true | id_equal,id_cons => true | id_equal,id_push1 => true | id_equal,id_case1 => true | id_equal,id_element => true | id_equal,id_locker2_adduniq => true | id_equal,id_and => true | id_equal,id_empty => true | id_equal,id_record_updates => true | id_equal,id_lock => true | id_equal,id_excllock => true | id_equal,id_pid => true | id_equal,id_calls => true | id_equal,id_subtract => true | id_equal,id_tag => true | id_equal,id_equal => true | id_equal,id_eq => false | id_equal,id_tops => false | id_equal,id_locker2_claim_lock => false | id_equal,id_pending => false | id_equal,id_andt => false | id_equal,id_tuplenil => false | id_equal,id_append => false | id_equal,id_0 => false | id_equal,id_case8 => false | id_eq,id_or => true | id_eq,id_gen_tag => true | id_eq,id_record_new => true | id_eq,id_a => true | id_eq,id_locker2_release_lock => true | id_eq,id_eqt => true | id_eq,id_istops => true | id_eq,id_locker2_map_add_pending => true | id_eq,id_release => true | id_eq,id_locker2_obtainable => true | id_eq,id_imp => true | id_eq,id_stack => true | id_eq,id_locker2_map_promote_pending => true | id_eq,id_locker => true | id_eq,id_case4 => true | id_eq,id_int => true | id_eq,id_push => true | id_eq,id_locker2_add_pending => true | id_eq,id_true => true | id_eq,id_locker2_check_availables => true | id_eq,id_F => true | id_eq,id_eqs => true | id_eq,id_record_update => true | id_eq,id_false => true | id_eq,id_gen_modtageq => true | id_eq,id_undefined => true | id_eq,id_nocalls => true | id_eq,id_locker2_remove_pending => true | id_eq,id_resource => true | id_eq,id_case6 => true | id_eq,id_if => true | id_eq,id_pops => true | id_eq,id_locker2_map_claim_lock => true | id_eq,id_ok => true | id_eq,id_case5 => true | id_eq,id_tuple => true | id_eq,id_member => true | id_eq,id_s => true | id_eq,id_delete => true | id_eq,id_T => true | id_eq,id_case9 => true | id_eq,id_record_extract => true | id_eq,id_excl => true | id_eq,id_case2 => true | id_eq,id_nil => true | id_eq,id_eqc => true | id_eq,id_case0 => true | id_eq,id_request => true | id_eq,id_locker2_check_available => true | id_eq,id_not => true | id_eq,id_pushs => true | id_eq,id_locker2_promote_pending => true | id_eq,id_mcrlrecord => true | id_eq,id_locker2_obtainables => true | id_eq,id_cons => true | id_eq,id_push1 => true | id_eq,id_case1 => true | id_eq,id_element => true | id_eq,id_locker2_adduniq => true | id_eq,id_and => true | id_eq,id_empty => true | id_eq,id_record_updates => true | id_eq,id_lock => true | id_eq,id_excllock => true | id_eq,id_pid => true | id_eq,id_calls => true | id_eq,id_subtract => true | id_eq,id_tag => true | id_eq,id_equal => true | id_eq,id_eq => true | id_eq,id_tops => false | id_eq,id_locker2_claim_lock => false | id_eq,id_pending => false | id_eq,id_andt => false | id_eq,id_tuplenil => false | id_eq,id_append => false | id_eq,id_0 => false | id_eq,id_case8 => false | id_tops,id_or => true | id_tops,id_gen_tag => true | id_tops,id_record_new => true | id_tops,id_a => true | id_tops,id_locker2_release_lock => true | id_tops,id_eqt => true | id_tops,id_istops => true | id_tops,id_locker2_map_add_pending => true | id_tops,id_release => true | id_tops,id_locker2_obtainable => true | id_tops,id_imp => true | id_tops,id_stack => true | id_tops,id_locker2_map_promote_pending => true | id_tops,id_locker => true | id_tops,id_case4 => true | id_tops,id_int => true | id_tops,id_push => true | id_tops,id_locker2_add_pending => true | id_tops,id_true => true | id_tops,id_locker2_check_availables => true | id_tops,id_F => true | id_tops,id_eqs => true | id_tops,id_record_update => true | id_tops,id_false => true | id_tops,id_gen_modtageq => true | id_tops,id_undefined => true | id_tops,id_nocalls => true | id_tops,id_locker2_remove_pending => true | id_tops,id_resource => true | id_tops,id_case6 => true | id_tops,id_if => true | id_tops,id_pops => true | id_tops,id_locker2_map_claim_lock => true | id_tops,id_ok => true | id_tops,id_case5 => true | id_tops,id_tuple => true | id_tops,id_member => true | id_tops,id_s => true | id_tops,id_delete => true | id_tops,id_T => true | id_tops,id_case9 => true | id_tops,id_record_extract => true | id_tops,id_excl => true | id_tops,id_case2 => true | id_tops,id_nil => true | id_tops,id_eqc => true | id_tops,id_case0 => true | id_tops,id_request => true | id_tops,id_locker2_check_available => true | id_tops,id_not => true | id_tops,id_pushs => true | id_tops,id_locker2_promote_pending => true | id_tops,id_mcrlrecord => true | id_tops,id_locker2_obtainables => true | id_tops,id_cons => true | id_tops,id_push1 => true | id_tops,id_case1 => true | id_tops,id_element => true | id_tops,id_locker2_adduniq => true | id_tops,id_and => true | id_tops,id_empty => true | id_tops,id_record_updates => true | id_tops,id_lock => true | id_tops,id_excllock => true | id_tops,id_pid => true | id_tops,id_calls => true | id_tops,id_subtract => true | id_tops,id_tag => true | id_tops,id_equal => true | id_tops,id_eq => true | id_tops,id_tops => true | id_tops,id_locker2_claim_lock => false | id_tops,id_pending => false | id_tops,id_andt => false | id_tops,id_tuplenil => false | id_tops,id_append => false | id_tops,id_0 => false | id_tops,id_case8 => false | id_locker2_claim_lock,id_or => true | id_locker2_claim_lock,id_gen_tag => true | id_locker2_claim_lock,id_record_new => true | id_locker2_claim_lock,id_a => true | id_locker2_claim_lock,id_locker2_release_lock => true | id_locker2_claim_lock,id_eqt => true | id_locker2_claim_lock,id_istops => true | id_locker2_claim_lock,id_locker2_map_add_pending => true | id_locker2_claim_lock,id_release => true | id_locker2_claim_lock,id_locker2_obtainable => true | id_locker2_claim_lock,id_imp => true | id_locker2_claim_lock,id_stack => true | id_locker2_claim_lock,id_locker2_map_promote_pending => true | id_locker2_claim_lock,id_locker => true | id_locker2_claim_lock,id_case4 => true | id_locker2_claim_lock,id_int => true | id_locker2_claim_lock,id_push => true | id_locker2_claim_lock,id_locker2_add_pending => true | id_locker2_claim_lock,id_true => true | id_locker2_claim_lock,id_locker2_check_availables => true | id_locker2_claim_lock,id_F => true | id_locker2_claim_lock,id_eqs => true | id_locker2_claim_lock,id_record_update => true | id_locker2_claim_lock,id_false => true | id_locker2_claim_lock,id_gen_modtageq => true | id_locker2_claim_lock,id_undefined => true | id_locker2_claim_lock,id_nocalls => true | id_locker2_claim_lock,id_locker2_remove_pending => true | id_locker2_claim_lock,id_resource => true | id_locker2_claim_lock,id_case6 => true | id_locker2_claim_lock,id_if => true | id_locker2_claim_lock,id_pops => true | id_locker2_claim_lock,id_locker2_map_claim_lock => true | id_locker2_claim_lock,id_ok => true | id_locker2_claim_lock,id_case5 => true | id_locker2_claim_lock,id_tuple => true | id_locker2_claim_lock,id_member => true | id_locker2_claim_lock,id_s => true | id_locker2_claim_lock,id_delete => true | id_locker2_claim_lock,id_T => true | id_locker2_claim_lock,id_case9 => true | id_locker2_claim_lock,id_record_extract => true | id_locker2_claim_lock,id_excl => true | id_locker2_claim_lock,id_case2 => true | id_locker2_claim_lock,id_nil => true | id_locker2_claim_lock,id_eqc => true | id_locker2_claim_lock,id_case0 => true | id_locker2_claim_lock,id_request => true | id_locker2_claim_lock,id_locker2_check_available => true | id_locker2_claim_lock,id_not => true | id_locker2_claim_lock,id_pushs => true | id_locker2_claim_lock,id_locker2_promote_pending => true | id_locker2_claim_lock,id_mcrlrecord => true | id_locker2_claim_lock,id_locker2_obtainables => true | id_locker2_claim_lock,id_cons => true | id_locker2_claim_lock,id_push1 => true | id_locker2_claim_lock,id_case1 => true | id_locker2_claim_lock,id_element => true | id_locker2_claim_lock,id_locker2_adduniq => true | id_locker2_claim_lock,id_and => true | id_locker2_claim_lock,id_empty => true | id_locker2_claim_lock,id_record_updates => true | id_locker2_claim_lock,id_lock => true | id_locker2_claim_lock,id_excllock => true | id_locker2_claim_lock,id_pid => true | id_locker2_claim_lock,id_calls => true | id_locker2_claim_lock,id_subtract => true | id_locker2_claim_lock,id_tag => true | id_locker2_claim_lock,id_equal => true | id_locker2_claim_lock,id_eq => true | id_locker2_claim_lock,id_tops => true | id_locker2_claim_lock,id_locker2_claim_lock => true | id_locker2_claim_lock,id_pending => false | id_locker2_claim_lock,id_andt => false | id_locker2_claim_lock,id_tuplenil => false | id_locker2_claim_lock,id_append => false | id_locker2_claim_lock,id_0 => false | id_locker2_claim_lock,id_case8 => false | id_pending,id_or => true | id_pending,id_gen_tag => true | id_pending,id_record_new => true | id_pending,id_a => true | id_pending,id_locker2_release_lock => true | id_pending,id_eqt => true | id_pending,id_istops => true | id_pending,id_locker2_map_add_pending => true | id_pending,id_release => true | id_pending,id_locker2_obtainable => true | id_pending,id_imp => true | id_pending,id_stack => true | id_pending,id_locker2_map_promote_pending => true | id_pending,id_locker => true | id_pending,id_case4 => true | id_pending,id_int => true | id_pending,id_push => true | id_pending,id_locker2_add_pending => true | id_pending,id_true => true | id_pending,id_locker2_check_availables => true | id_pending,id_F => true | id_pending,id_eqs => true | id_pending,id_record_update => true | id_pending,id_false => true | id_pending,id_gen_modtageq => true | id_pending,id_undefined => true | id_pending,id_nocalls => true | id_pending,id_locker2_remove_pending => true | id_pending,id_resource => true | id_pending,id_case6 => true | id_pending,id_if => true | id_pending,id_pops => true | id_pending,id_locker2_map_claim_lock => true | id_pending,id_ok => true | id_pending,id_case5 => true | id_pending,id_tuple => true | id_pending,id_member => true | id_pending,id_s => true | id_pending,id_delete => true | id_pending,id_T => true | id_pending,id_case9 => true | id_pending,id_record_extract => true | id_pending,id_excl => true | id_pending,id_case2 => true | id_pending,id_nil => true | id_pending,id_eqc => true | id_pending,id_case0 => true | id_pending,id_request => true | id_pending,id_locker2_check_available => true | id_pending,id_not => true | id_pending,id_pushs => true | id_pending,id_locker2_promote_pending => true | id_pending,id_mcrlrecord => true | id_pending,id_locker2_obtainables => true | id_pending,id_cons => true | id_pending,id_push1 => true | id_pending,id_case1 => true | id_pending,id_element => true | id_pending,id_locker2_adduniq => true | id_pending,id_and => true | id_pending,id_empty => true | id_pending,id_record_updates => true | id_pending,id_lock => true | id_pending,id_excllock => true | id_pending,id_pid => true | id_pending,id_calls => true | id_pending,id_subtract => true | id_pending,id_tag => true | id_pending,id_equal => true | id_pending,id_eq => true | id_pending,id_tops => true | id_pending,id_locker2_claim_lock => true | id_pending,id_pending => true | id_pending,id_andt => false | id_pending,id_tuplenil => false | id_pending,id_append => false | id_pending,id_0 => false | id_pending,id_case8 => false | id_andt,id_or => true | id_andt,id_gen_tag => true | id_andt,id_record_new => true | id_andt,id_a => true | id_andt,id_locker2_release_lock => true | id_andt,id_eqt => true | id_andt,id_istops => true | id_andt,id_locker2_map_add_pending => true | id_andt,id_release => true | id_andt,id_locker2_obtainable => true | id_andt,id_imp => true | id_andt,id_stack => true | id_andt,id_locker2_map_promote_pending => true | id_andt,id_locker => true | id_andt,id_case4 => true | id_andt,id_int => true | id_andt,id_push => true | id_andt,id_locker2_add_pending => true | id_andt,id_true => true | id_andt,id_locker2_check_availables => true | id_andt,id_F => true | id_andt,id_eqs => true | id_andt,id_record_update => true | id_andt,id_false => true | id_andt,id_gen_modtageq => true | id_andt,id_undefined => true | id_andt,id_nocalls => true | id_andt,id_locker2_remove_pending => true | id_andt,id_resource => true | id_andt,id_case6 => true | id_andt,id_if => true | id_andt,id_pops => true | id_andt,id_locker2_map_claim_lock => true | id_andt,id_ok => true | id_andt,id_case5 => true | id_andt,id_tuple => true | id_andt,id_member => true | id_andt,id_s => true | id_andt,id_delete => true | id_andt,id_T => true | id_andt,id_case9 => true | id_andt,id_record_extract => true | id_andt,id_excl => true | id_andt,id_case2 => true | id_andt,id_nil => true | id_andt,id_eqc => true | id_andt,id_case0 => true | id_andt,id_request => true | id_andt,id_locker2_check_available => true | id_andt,id_not => true | id_andt,id_pushs => true | id_andt,id_locker2_promote_pending => true | id_andt,id_mcrlrecord => true | id_andt,id_locker2_obtainables => true | id_andt,id_cons => true | id_andt,id_push1 => true | id_andt,id_case1 => true | id_andt,id_element => true | id_andt,id_locker2_adduniq => true | id_andt,id_and => true | id_andt,id_empty => true | id_andt,id_record_updates => true | id_andt,id_lock => true | id_andt,id_excllock => true | id_andt,id_pid => true | id_andt,id_calls => true | id_andt,id_subtract => true | id_andt,id_tag => true | id_andt,id_equal => true | id_andt,id_eq => true | id_andt,id_tops => true | id_andt,id_locker2_claim_lock => true | id_andt,id_pending => true | id_andt,id_andt => true | id_andt,id_tuplenil => false | id_andt,id_append => false | id_andt,id_0 => false | id_andt,id_case8 => false | id_tuplenil,id_or => true | id_tuplenil,id_gen_tag => true | id_tuplenil,id_record_new => true | id_tuplenil,id_a => true | id_tuplenil,id_locker2_release_lock => true | id_tuplenil,id_eqt => true | id_tuplenil,id_istops => true | id_tuplenil,id_locker2_map_add_pending => true | id_tuplenil,id_release => true | id_tuplenil,id_locker2_obtainable => true | id_tuplenil,id_imp => true | id_tuplenil,id_stack => true | id_tuplenil,id_locker2_map_promote_pending => true | id_tuplenil,id_locker => true | id_tuplenil,id_case4 => true | id_tuplenil,id_int => true | id_tuplenil,id_push => true | id_tuplenil,id_locker2_add_pending => true | id_tuplenil,id_true => true | id_tuplenil,id_locker2_check_availables => true | id_tuplenil,id_F => true | id_tuplenil,id_eqs => true | id_tuplenil,id_record_update => true | id_tuplenil,id_false => true | id_tuplenil,id_gen_modtageq => true | id_tuplenil,id_undefined => true | id_tuplenil,id_nocalls => true | id_tuplenil,id_locker2_remove_pending => true | id_tuplenil,id_resource => true | id_tuplenil,id_case6 => true | id_tuplenil,id_if => true | id_tuplenil,id_pops => true | id_tuplenil,id_locker2_map_claim_lock => true | id_tuplenil,id_ok => true | id_tuplenil,id_case5 => true | id_tuplenil,id_tuple => true | id_tuplenil,id_member => true | id_tuplenil,id_s => true | id_tuplenil,id_delete => true | id_tuplenil,id_T => true | id_tuplenil,id_case9 => true | id_tuplenil,id_record_extract => true | id_tuplenil,id_excl => true | id_tuplenil,id_case2 => true | id_tuplenil,id_nil => true | id_tuplenil,id_eqc => true | id_tuplenil,id_case0 => true | id_tuplenil,id_request => true | id_tuplenil,id_locker2_check_available => true | id_tuplenil,id_not => true | id_tuplenil,id_pushs => true | id_tuplenil,id_locker2_promote_pending => true | id_tuplenil,id_mcrlrecord => true | id_tuplenil,id_locker2_obtainables => true | id_tuplenil,id_cons => true | id_tuplenil,id_push1 => true | id_tuplenil,id_case1 => true | id_tuplenil,id_element => true | id_tuplenil,id_locker2_adduniq => true | id_tuplenil,id_and => true | id_tuplenil,id_empty => true | id_tuplenil,id_record_updates => true | id_tuplenil,id_lock => true | id_tuplenil,id_excllock => true | id_tuplenil,id_pid => true | id_tuplenil,id_calls => true | id_tuplenil,id_subtract => true | id_tuplenil,id_tag => true | id_tuplenil,id_equal => true | id_tuplenil,id_eq => true | id_tuplenil,id_tops => true | id_tuplenil,id_locker2_claim_lock => true | id_tuplenil,id_pending => true | id_tuplenil,id_andt => true | id_tuplenil,id_tuplenil => true | id_tuplenil,id_append => false | id_tuplenil,id_0 => false | id_tuplenil,id_case8 => false | id_append,id_or => true | id_append,id_gen_tag => true | id_append,id_record_new => true | id_append,id_a => true | id_append,id_locker2_release_lock => true | id_append,id_eqt => true | id_append,id_istops => true | id_append,id_locker2_map_add_pending => true | id_append,id_release => true | id_append,id_locker2_obtainable => true | id_append,id_imp => true | id_append,id_stack => true | id_append,id_locker2_map_promote_pending => true | id_append,id_locker => true | id_append,id_case4 => true | id_append,id_int => true | id_append,id_push => true | id_append,id_locker2_add_pending => true | id_append,id_true => true | id_append,id_locker2_check_availables => true | id_append,id_F => true | id_append,id_eqs => true | id_append,id_record_update => true | id_append,id_false => true | id_append,id_gen_modtageq => true | id_append,id_undefined => true | id_append,id_nocalls => true | id_append,id_locker2_remove_pending => true | id_append,id_resource => true | id_append,id_case6 => true | id_append,id_if => true | id_append,id_pops => true | id_append,id_locker2_map_claim_lock => true | id_append,id_ok => true | id_append,id_case5 => true | id_append,id_tuple => true | id_append,id_member => true | id_append,id_s => true | id_append,id_delete => true | id_append,id_T => true | id_append,id_case9 => true | id_append,id_record_extract => true | id_append,id_excl => true | id_append,id_case2 => true | id_append,id_nil => true | id_append,id_eqc => true | id_append,id_case0 => true | id_append,id_request => true | id_append,id_locker2_check_available => true | id_append,id_not => true | id_append,id_pushs => true | id_append,id_locker2_promote_pending => true | id_append,id_mcrlrecord => true | id_append,id_locker2_obtainables => true | id_append,id_cons => true | id_append,id_push1 => true | id_append,id_case1 => true | id_append,id_element => true | id_append,id_locker2_adduniq => true | id_append,id_and => true | id_append,id_empty => true | id_append,id_record_updates => true | id_append,id_lock => true | id_append,id_excllock => true | id_append,id_pid => true | id_append,id_calls => true | id_append,id_subtract => true | id_append,id_tag => true | id_append,id_equal => true | id_append,id_eq => true | id_append,id_tops => true | id_append,id_locker2_claim_lock => true | id_append,id_pending => true | id_append,id_andt => true | id_append,id_tuplenil => true | id_append,id_append => true | id_append,id_0 => false | id_append,id_case8 => false | id_0,id_or => true | id_0,id_gen_tag => true | id_0,id_record_new => true | id_0,id_a => true | id_0,id_locker2_release_lock => true | id_0,id_eqt => true | id_0,id_istops => true | id_0,id_locker2_map_add_pending => true | id_0,id_release => true | id_0,id_locker2_obtainable => true | id_0,id_imp => true | id_0,id_stack => true | id_0,id_locker2_map_promote_pending => true | id_0,id_locker => true | id_0,id_case4 => true | id_0,id_int => true | id_0,id_push => true | id_0,id_locker2_add_pending => true | id_0,id_true => true | id_0,id_locker2_check_availables => true | id_0,id_F => true | id_0,id_eqs => true | id_0,id_record_update => true | id_0,id_false => true | id_0,id_gen_modtageq => true | id_0,id_undefined => true | id_0,id_nocalls => true | id_0,id_locker2_remove_pending => true | id_0,id_resource => true | id_0,id_case6 => true | id_0,id_if => true | id_0,id_pops => true | id_0,id_locker2_map_claim_lock => true | id_0,id_ok => true | id_0,id_case5 => true | id_0,id_tuple => true | id_0,id_member => true | id_0,id_s => true | id_0,id_delete => true | id_0,id_T => true | id_0,id_case9 => true | id_0,id_record_extract => true | id_0,id_excl => true | id_0,id_case2 => true | id_0,id_nil => true | id_0,id_eqc => true | id_0,id_case0 => true | id_0,id_request => true | id_0,id_locker2_check_available => true | id_0,id_not => true | id_0,id_pushs => true | id_0,id_locker2_promote_pending => true | id_0,id_mcrlrecord => true | id_0,id_locker2_obtainables => true | id_0,id_cons => true | id_0,id_push1 => true | id_0,id_case1 => true | id_0,id_element => true | id_0,id_locker2_adduniq => true | id_0,id_and => true | id_0,id_empty => true | id_0,id_record_updates => true | id_0,id_lock => true | id_0,id_excllock => true | id_0,id_pid => true | id_0,id_calls => true | id_0,id_subtract => true | id_0,id_tag => true | id_0,id_equal => true | id_0,id_eq => true | id_0,id_tops => true | id_0,id_locker2_claim_lock => true | id_0,id_pending => true | id_0,id_andt => true | id_0,id_tuplenil => true | id_0,id_append => true | id_0,id_0 => true | id_0,id_case8 => false | id_case8,id_or => true | id_case8,id_gen_tag => true | id_case8,id_record_new => true | id_case8,id_a => true | id_case8,id_locker2_release_lock => true | id_case8,id_eqt => true | id_case8,id_istops => true | id_case8,id_locker2_map_add_pending => true | id_case8,id_release => true | id_case8,id_locker2_obtainable => true | id_case8,id_imp => true | id_case8,id_stack => true | id_case8,id_locker2_map_promote_pending => true | id_case8,id_locker => true | id_case8,id_case4 => true | id_case8,id_int => true | id_case8,id_push => true | id_case8,id_locker2_add_pending => true | id_case8,id_true => true | id_case8,id_locker2_check_availables => true | id_case8,id_F => true | id_case8,id_eqs => true | id_case8,id_record_update => true | id_case8,id_false => true | id_case8,id_gen_modtageq => true | id_case8,id_undefined => true | id_case8,id_nocalls => true | id_case8,id_locker2_remove_pending => true | id_case8,id_resource => true | id_case8,id_case6 => true | id_case8,id_if => true | id_case8,id_pops => true | id_case8,id_locker2_map_claim_lock => true | id_case8,id_ok => true | id_case8,id_case5 => true | id_case8,id_tuple => true | id_case8,id_member => true | id_case8,id_s => true | id_case8,id_delete => true | id_case8,id_T => true | id_case8,id_case9 => true | id_case8,id_record_extract => true | id_case8,id_excl => true | id_case8,id_case2 => true | id_case8,id_nil => true | id_case8,id_eqc => true | id_case8,id_case0 => true | id_case8,id_request => true | id_case8,id_locker2_check_available => true | id_case8,id_not => true | id_case8,id_pushs => true | id_case8,id_locker2_promote_pending => true | id_case8,id_mcrlrecord => true | id_case8,id_locker2_obtainables => true | id_case8,id_cons => true | id_case8,id_push1 => true | id_case8,id_case1 => true | id_case8,id_element => true | id_case8,id_locker2_adduniq => true | id_case8,id_and => true | id_case8,id_empty => true | id_case8,id_record_updates => true | id_case8,id_lock => true | id_case8,id_excllock => true | id_case8,id_pid => true | id_case8,id_calls => true | id_case8,id_subtract => true | id_case8,id_tag => true | id_case8,id_equal => true | id_case8,id_eq => true | id_case8,id_tops => true | id_case8,id_locker2_claim_lock => true | id_case8,id_pending => true | id_case8,id_andt => true | id_case8,id_tuplenil => true | id_case8,id_append => true | id_case8,id_0 => true | id_case8,id_case8 => 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 := (* or(T,T) -> T *) | R_xml_0_rule_0 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_or ((algebra.Alg.Term algebra.F.id_T nil)::(algebra.Alg.Term algebra.F.id_T nil)::nil)) (* or(F,T) -> T *) | R_xml_0_rule_1 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_or ((algebra.Alg.Term algebra.F.id_F nil)::(algebra.Alg.Term algebra.F.id_T nil)::nil)) (* or(T,F) -> T *) | R_xml_0_rule_2 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_or ((algebra.Alg.Term algebra.F.id_T nil)::(algebra.Alg.Term algebra.F.id_F nil)::nil)) (* or(F,F) -> F *) | R_xml_0_rule_3 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_or ((algebra.Alg.Term algebra.F.id_F nil)::(algebra.Alg.Term algebra.F.id_F nil)::nil)) (* and(T,B_) -> B_ *) | R_xml_0_rule_4 : R_xml_0_rules (algebra.Alg.Var 1) (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_T nil)::(algebra.Alg.Var 1)::nil)) (* and(B_,T) -> B_ *) | R_xml_0_rule_5 : 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_T nil)::nil)) (* and(F,B_) -> F *) | R_xml_0_rule_6 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_F nil)::(algebra.Alg.Var 1)::nil)) (* and(B_,F) -> F *) | R_xml_0_rule_7 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id_F nil)::nil)) (* imp(T,B_) -> B_ *) | R_xml_0_rule_8 : R_xml_0_rules (algebra.Alg.Var 1) (algebra.Alg.Term algebra.F.id_imp ((algebra.Alg.Term algebra.F.id_T nil)::(algebra.Alg.Var 1)::nil)) (* imp(F,B_) -> T *) | R_xml_0_rule_9 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_imp ((algebra.Alg.Term algebra.F.id_F nil)::(algebra.Alg.Var 1)::nil)) (* not(T) -> F *) | R_xml_0_rule_10 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_not ((algebra.Alg.Term algebra.F.id_T nil)::nil)) (* not(F) -> T *) | R_xml_0_rule_11 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_not ((algebra.Alg.Term algebra.F.id_F nil)::nil)) (* if(T,B1_,B2_) -> B1_ *) | R_xml_0_rule_12 : R_xml_0_rules (algebra.Alg.Var 2) (algebra.Alg.Term algebra.F.id_if ((algebra.Alg.Term algebra.F.id_T nil)::(algebra.Alg.Var 2)::(algebra.Alg.Var 3)::nil)) (* if(F,B1_,B2_) -> B2_ *) | R_xml_0_rule_13 : R_xml_0_rules (algebra.Alg.Var 3) (algebra.Alg.Term algebra.F.id_if ((algebra.Alg.Term algebra.F.id_F nil)::(algebra.Alg.Var 2)::(algebra.Alg.Var 3)::nil)) (* eq(T,T) -> T *) | R_xml_0_rule_14 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_eq ((algebra.Alg.Term algebra.F.id_T nil)::(algebra.Alg.Term algebra.F.id_T nil)::nil)) (* eq(F,F) -> T *) | R_xml_0_rule_15 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_eq ((algebra.Alg.Term algebra.F.id_F nil)::(algebra.Alg.Term algebra.F.id_F nil)::nil)) (* eq(T,F) -> F *) | R_xml_0_rule_16 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eq ((algebra.Alg.Term algebra.F.id_T nil)::(algebra.Alg.Term algebra.F.id_F nil)::nil)) (* eq(F,T) -> F *) | R_xml_0_rule_17 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eq ((algebra.Alg.Term algebra.F.id_F nil)::(algebra.Alg.Term algebra.F.id_T nil)::nil)) (* eqt(nil,undefined) -> F *) | R_xml_0_rule_18 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(nil,pid(N2_)) -> F *) | R_xml_0_rule_19 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(nil,int(N2_)) -> F *) | R_xml_0_rule_20 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(nil,cons(H2_,T2_)) -> F *) | R_xml_0_rule_21 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(nil,tuple(H2_,T2_)) -> F *) | R_xml_0_rule_22 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(nil,tuplenil(H2_)) -> F *) | R_xml_0_rule_23 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* eqt(a,nil) -> F *) | R_xml_0_rule_24 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* eqt(a,a) -> T *) | R_xml_0_rule_25 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)) (* eqt(a,excl) -> F *) | R_xml_0_rule_26 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)) (* eqt(a,false) -> F *) | R_xml_0_rule_27 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* eqt(a,lock) -> F *) | R_xml_0_rule_28 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)) (* eqt(a,locker) -> F *) | R_xml_0_rule_29 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)) (* eqt(a,mcrlrecord) -> F *) | R_xml_0_rule_30 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)) (* eqt(a,ok) -> F *) | R_xml_0_rule_31 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)) (* eqt(a,pending) -> F *) | R_xml_0_rule_32 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (* eqt(a,release) -> F *) | R_xml_0_rule_33 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)) (* eqt(a,request) -> F *) | R_xml_0_rule_34 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)) (* eqt(a,resource) -> F *) | R_xml_0_rule_35 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(a,tag) -> F *) | R_xml_0_rule_36 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)) (* eqt(a,true) -> F *) | R_xml_0_rule_37 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* eqt(a,undefined) -> F *) | R_xml_0_rule_38 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(a,pid(N2_)) -> F *) | R_xml_0_rule_39 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(a,int(N2_)) -> F *) | R_xml_0_rule_40 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(a,cons(H2_,T2_)) -> F *) | R_xml_0_rule_41 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(a,tuple(H2_,T2_)) -> F *) | R_xml_0_rule_42 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(a,tuplenil(H2_)) -> F *) | R_xml_0_rule_43 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* eqt(excl,nil) -> F *) | R_xml_0_rule_44 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* eqt(excl,a) -> F *) | R_xml_0_rule_45 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)) (* eqt(excl,excl) -> T *) | R_xml_0_rule_46 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)) (* eqt(excl,false) -> F *) | R_xml_0_rule_47 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* eqt(excl,lock) -> F *) | R_xml_0_rule_48 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)) (* eqt(excl,locker) -> F *) | R_xml_0_rule_49 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)) (* eqt(excl,mcrlrecord) -> F *) | R_xml_0_rule_50 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)) (* eqt(excl,ok) -> F *) | R_xml_0_rule_51 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)) (* eqt(excl,pending) -> F *) | R_xml_0_rule_52 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (* eqt(excl,release) -> F *) | R_xml_0_rule_53 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)) (* eqt(excl,request) -> F *) | R_xml_0_rule_54 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)) (* eqt(excl,resource) -> F *) | R_xml_0_rule_55 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(excl,tag) -> F *) | R_xml_0_rule_56 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)) (* eqt(excl,true) -> F *) | R_xml_0_rule_57 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* eqt(excl,undefined) -> F *) | R_xml_0_rule_58 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(excl,pid(N2_)) -> F *) | R_xml_0_rule_59 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(excl,eqt(false,int(N2_))) -> F *) | R_xml_0_rule_60 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_false nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil))::nil)) (* eqt(false,cons(H2_,T2_)) -> F *) | R_xml_0_rule_61 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_false nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* eqt(false,tuple(H2_,T2_)) -> F *) | R_xml_0_rule_62 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_false nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* eqt(false,tuplenil(H2_)) -> F *) | R_xml_0_rule_63 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_false nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* eqt(lock,nil) -> F *) | R_xml_0_rule_64 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* eqt(lock,a) -> F *) | R_xml_0_rule_65 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)) (* eqt(lock,excl) -> F *) | R_xml_0_rule_66 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)) (* eqt(lock,false) -> F *) | R_xml_0_rule_67 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* eqt(lock,lock) -> T *) | R_xml_0_rule_68 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)) (* eqt(lock,locker) -> F *) | R_xml_0_rule_69 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)) (* eqt(lock,mcrlrecord) -> F *) | R_xml_0_rule_70 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)) (* eqt(lock,ok) -> F *) | R_xml_0_rule_71 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)) (* eqt(lock,pending) -> F *) | R_xml_0_rule_72 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (* eqt(lock,release) -> F *) | R_xml_0_rule_73 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)) (* eqt(lock,request) -> F *) | R_xml_0_rule_74 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)) (* eqt(lock,resource) -> F *) | R_xml_0_rule_75 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(lock,tag) -> F *) | R_xml_0_rule_76 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)) (* eqt(lock,true) -> F *) | R_xml_0_rule_77 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* eqt(lock,undefined) -> F *) | R_xml_0_rule_78 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(lock,pid(N2_)) -> F *) | R_xml_0_rule_79 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(lock,int(N2_)) -> F *) | R_xml_0_rule_80 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(lock,cons(H2_,T2_)) -> F *) | R_xml_0_rule_81 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(lock,tuple(H2_,T2_)) -> F *) | R_xml_0_rule_82 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(lock,tuplenil(H2_)) -> F *) | R_xml_0_rule_83 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* eqt(locker,nil) -> F *) | R_xml_0_rule_84 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* eqt(locker,a) -> F *) | R_xml_0_rule_85 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)) (* eqt(locker,excl) -> F *) | R_xml_0_rule_86 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)) (* eqt(locker,false) -> F *) | R_xml_0_rule_87 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* eqt(locker,lock) -> F *) | R_xml_0_rule_88 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)) (* eqt(locker,locker) -> T *) | R_xml_0_rule_89 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)) (* eqt(locker,mcrlrecord) -> F *) | R_xml_0_rule_90 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)) (* eqt(locker,ok) -> F *) | R_xml_0_rule_91 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)) (* eqt(locker,pending) -> F *) | R_xml_0_rule_92 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (* eqt(locker,release) -> F *) | R_xml_0_rule_93 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)) (* eqt(locker,request) -> F *) | R_xml_0_rule_94 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)) (* eqt(locker,resource) -> F *) | R_xml_0_rule_95 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(locker,tag) -> F *) | R_xml_0_rule_96 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)) (* eqt(locker,true) -> F *) | R_xml_0_rule_97 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* eqt(locker,undefined) -> F *) | R_xml_0_rule_98 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(locker,pid(N2_)) -> F *) | R_xml_0_rule_99 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(locker,int(N2_)) -> F *) | R_xml_0_rule_100 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(locker,cons(H2_,T2_)) -> F *) | R_xml_0_rule_101 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* eqt(locker,tuple(H2_,T2_)) -> F *) | R_xml_0_rule_102 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* eqt(locker,tuplenil(H2_)) -> F *) | R_xml_0_rule_103 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* eqt(mcrlrecord,nil) -> F *) | R_xml_0_rule_104 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* eqt(mcrlrecord,a) -> F *) | R_xml_0_rule_105 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)) (* eqt(mcrlrecord,excl) -> F *) | R_xml_0_rule_106 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)) (* eqt(mcrlrecord,false) -> F *) | R_xml_0_rule_107 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* eqt(mcrlrecord,lock) -> F *) | R_xml_0_rule_108 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)) (* eqt(mcrlrecord,locker) -> F *) | R_xml_0_rule_109 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)) (* eqt(mcrlrecord,mcrlrecord) -> T *) | R_xml_0_rule_110 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)) (* eqt(mcrlrecord,ok) -> F *) | R_xml_0_rule_111 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)) (* eqt(mcrlrecord,pending) -> F *) | R_xml_0_rule_112 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (* eqt(mcrlrecord,release) -> F *) | R_xml_0_rule_113 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)) (* eqt(mcrlrecord,request) -> F *) | R_xml_0_rule_114 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)) (* eqt(mcrlrecord,resource) -> F *) | R_xml_0_rule_115 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(ok,resource) -> F *) | R_xml_0_rule_116 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(ok,tag) -> F *) | R_xml_0_rule_117 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)) (* eqt(ok,true) -> F *) | R_xml_0_rule_118 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* eqt(ok,undefined) -> F *) | R_xml_0_rule_119 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(ok,pid(N2_)) -> F *) | R_xml_0_rule_120 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(ok,int(N2_)) -> F *) | R_xml_0_rule_121 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(ok,cons(H2_,T2_)) -> F *) | R_xml_0_rule_122 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(ok,tuple(H2_,T2_)) -> F *) | R_xml_0_rule_123 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(ok,tuplenil(H2_)) -> F *) | R_xml_0_rule_124 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* eqt(pending,nil) -> F *) | R_xml_0_rule_125 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* eqt(pending,a) -> F *) | R_xml_0_rule_126 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)) (* eqt(pending,excl) -> F *) | R_xml_0_rule_127 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)) (* eqt(pending,false) -> F *) | R_xml_0_rule_128 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* eqt(pending,lock) -> F *) | R_xml_0_rule_129 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)) (* eqt(pending,locker) -> F *) | R_xml_0_rule_130 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)) (* eqt(pending,mcrlrecord) -> F *) | R_xml_0_rule_131 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)) (* eqt(pending,ok) -> F *) | R_xml_0_rule_132 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)) (* eqt(pending,pending) -> T *) | R_xml_0_rule_133 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (* eqt(pending,release) -> F *) | R_xml_0_rule_134 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)) (* eqt(pending,request) -> F *) | R_xml_0_rule_135 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)) (* eqt(pending,resource) -> F *) | R_xml_0_rule_136 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(pending,tag) -> F *) | R_xml_0_rule_137 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)) (* eqt(pending,true) -> F *) | R_xml_0_rule_138 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* eqt(pending,undefined) -> F *) | R_xml_0_rule_139 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(pending,pid(N2_)) -> F *) | R_xml_0_rule_140 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(pending,int(N2_)) -> F *) | R_xml_0_rule_141 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(pending,cons(H2_,T2_)) -> F *) | R_xml_0_rule_142 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* eqt(pending,tuple(H2_,T2_)) -> F *) | R_xml_0_rule_143 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* eqt(pending,tuplenil(H2_)) -> F *) | R_xml_0_rule_144 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* eqt(release,nil) -> F *) | R_xml_0_rule_145 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* eqt(release,a) -> F *) | R_xml_0_rule_146 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)) (* eqt(release,excl) -> F *) | R_xml_0_rule_147 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)) (* eqt(release,false) -> F *) | R_xml_0_rule_148 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* eqt(release,lock) -> F *) | R_xml_0_rule_149 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)) (* eqt(release,locker) -> F *) | R_xml_0_rule_150 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)) (* eqt(release,mcrlrecord) -> F *) | R_xml_0_rule_151 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)) (* eqt(release,ok) -> F *) | R_xml_0_rule_152 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)) (* eqt(request,mcrlrecord) -> F *) | R_xml_0_rule_153 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)) (* eqt(request,ok) -> F *) | R_xml_0_rule_154 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)) (* eqt(request,pending) -> F *) | R_xml_0_rule_155 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (* eqt(request,release) -> F *) | R_xml_0_rule_156 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)) (* eqt(request,request) -> T *) | R_xml_0_rule_157 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)) (* eqt(request,resource) -> F *) | R_xml_0_rule_158 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(request,tag) -> F *) | R_xml_0_rule_159 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)) (* eqt(request,true) -> F *) | R_xml_0_rule_160 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* eqt(request,undefined) -> F *) | R_xml_0_rule_161 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(request,pid(N2_)) -> F *) | R_xml_0_rule_162 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(request,int(N2_)) -> F *) | R_xml_0_rule_163 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(request,cons(H2_,T2_)) -> F *) | R_xml_0_rule_164 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* eqt(request,tuple(H2_,T2_)) -> F *) | R_xml_0_rule_165 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* eqt(request,tuplenil(H2_)) -> F *) | R_xml_0_rule_166 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* eqt(resource,nil) -> F *) | R_xml_0_rule_167 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* eqt(resource,a) -> F *) | R_xml_0_rule_168 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)) (* eqt(resource,excl) -> F *) | R_xml_0_rule_169 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)) (* eqt(resource,false) -> F *) | R_xml_0_rule_170 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* eqt(resource,lock) -> F *) | R_xml_0_rule_171 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)) (* eqt(resource,locker) -> F *) | R_xml_0_rule_172 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)) (* eqt(resource,mcrlrecord) -> F *) | R_xml_0_rule_173 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)) (* eqt(resource,ok) -> F *) | R_xml_0_rule_174 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)) (* eqt(resource,pending) -> F *) | R_xml_0_rule_175 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (* eqt(resource,release) -> F *) | R_xml_0_rule_176 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)) (* eqt(resource,request) -> F *) | R_xml_0_rule_177 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)) (* eqt(resource,resource) -> T *) | R_xml_0_rule_178 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(resource,tag) -> F *) | R_xml_0_rule_179 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)) (* eqt(resource,true) -> F *) | R_xml_0_rule_180 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* eqt(resource,undefined) -> F *) | R_xml_0_rule_181 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(resource,pid(N2_)) -> F *) | R_xml_0_rule_182 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(resource,int(N2_)) -> F *) | R_xml_0_rule_183 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(resource,cons(H2_,T2_)) -> F *) | R_xml_0_rule_184 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* eqt(resource,tuple(H2_,T2_)) -> F *) | R_xml_0_rule_185 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* eqt(resource,tuplenil(H2_)) -> F *) | R_xml_0_rule_186 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* eqt(tag,nil) -> F *) | R_xml_0_rule_187 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* eqt(tag,a) -> F *) | R_xml_0_rule_188 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)) (* eqt(tag,excl) -> F *) | R_xml_0_rule_189 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)) (* eqt(tag,false) -> F *) | R_xml_0_rule_190 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* eqt(tag,lock) -> F *) | R_xml_0_rule_191 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)) (* eqt(tag,locker) -> F *) | R_xml_0_rule_192 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)) (* eqt(tag,mcrlrecord) -> F *) | R_xml_0_rule_193 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)) (* eqt(tag,ok) -> F *) | R_xml_0_rule_194 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)) (* eqt(tag,pending) -> F *) | R_xml_0_rule_195 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (* eqt(tag,release) -> F *) | R_xml_0_rule_196 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)) (* eqt(tag,request) -> F *) | R_xml_0_rule_197 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)) (* eqt(tag,resource) -> F *) | R_xml_0_rule_198 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(tag,tag) -> T *) | R_xml_0_rule_199 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)) (* eqt(tag,true) -> F *) | R_xml_0_rule_200 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* eqt(tag,undefined) -> F *) | R_xml_0_rule_201 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(tag,pid(N2_)) -> F *) | R_xml_0_rule_202 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(tag,int(N2_)) -> F *) | R_xml_0_rule_203 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(tag,cons(H2_,T2_)) -> F *) | R_xml_0_rule_204 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(tag,tuple(H2_,T2_)) -> F *) | R_xml_0_rule_205 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(tag,tuplenil(H2_)) -> F *) | R_xml_0_rule_206 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* eqt(true,nil) -> F *) | R_xml_0_rule_207 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* eqt(true,a) -> F *) | R_xml_0_rule_208 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)) (* eqt(true,excl) -> F *) | R_xml_0_rule_209 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)) (* eqt(true,false) -> F *) | R_xml_0_rule_210 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* eqt(true,lock) -> F *) | R_xml_0_rule_211 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)) (* eqt(true,locker) -> F *) | R_xml_0_rule_212 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)) (* eqt(true,mcrlrecord) -> F *) | R_xml_0_rule_213 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)) (* eqt(true,ok) -> F *) | R_xml_0_rule_214 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)) (* eqt(true,pending) -> F *) | R_xml_0_rule_215 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (* eqt(true,release) -> F *) | R_xml_0_rule_216 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)) (* eqt(true,request) -> F *) | R_xml_0_rule_217 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)) (* eqt(true,resource) -> F *) | R_xml_0_rule_218 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(true,tag) -> F *) | R_xml_0_rule_219 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)) (* eqt(true,true) -> T *) | R_xml_0_rule_220 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* eqt(true,undefined) -> F *) | R_xml_0_rule_221 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(true,pid(N2_)) -> F *) | R_xml_0_rule_222 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(true,int(N2_)) -> F *) | R_xml_0_rule_223 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(true,cons(H2_,T2_)) -> F *) | R_xml_0_rule_224 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(true,tuple(H2_,T2_)) -> F *) | R_xml_0_rule_225 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(true,tuplenil(H2_)) -> F *) | R_xml_0_rule_226 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* eqt(undefined,nil) -> F *) | R_xml_0_rule_227 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_undefined nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* eqt(undefined,a) -> F *) | R_xml_0_rule_228 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_undefined nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)) (* eqt(undefined,tuplenil(H2_)) -> F *) | R_xml_0_rule_229 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_undefined nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* eqt(pid(N1_),nil) -> F *) | R_xml_0_rule_230 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* eqt(pid(N1_),a) -> F *) | R_xml_0_rule_231 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_a nil)::nil)) (* eqt(pid(N1_),excl) -> F *) | R_xml_0_rule_232 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_excl nil)::nil)) (* eqt(pid(N1_),false) -> F *) | R_xml_0_rule_233 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* eqt(pid(N1_),lock) -> F *) | R_xml_0_rule_234 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_lock nil)::nil)) (* eqt(pid(N1_),locker) -> F *) | R_xml_0_rule_235 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_locker nil)::nil)) (* eqt(pid(N1_),mcrlrecord) -> F *) | R_xml_0_rule_236 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)) (* eqt(pid(N1_),ok) -> F *) | R_xml_0_rule_237 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_ok nil)::nil)) (* eqt(pid(N1_),pending) -> F *) | R_xml_0_rule_238 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (* eqt(pid(N1_),release) -> F *) | R_xml_0_rule_239 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_release nil)::nil)) (* eqt(pid(N1_),request) -> F *) | R_xml_0_rule_240 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_request nil)::nil)) (* eqt(pid(N1_),resource) -> F *) | R_xml_0_rule_241 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(pid(N1_),tag) -> F *) | R_xml_0_rule_242 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_tag nil)::nil)) (* eqt(pid(N1_),true) -> F *) | R_xml_0_rule_243 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* eqt(pid(N1_),undefined) -> F *) | R_xml_0_rule_244 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(pid(N1_),pid(N2_)) -> eqt(N1_,N2_) *) | R_xml_0_rule_245 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 7):: (algebra.Alg.Var 4)::nil)) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(pid(N1_),int(N2_)) -> F *) | R_xml_0_rule_246 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(pid(N1_),cons(H2_,T2_)) -> F *) | R_xml_0_rule_247 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* eqt(pid(N1_),tuple(H2_,T2_)) -> F *) | R_xml_0_rule_248 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)) (* eqt(pid(N1_),tuplenil(H2_)) -> F *) | R_xml_0_rule_249 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* eqt(int(N1_),nil) -> F *) | R_xml_0_rule_250 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* eqt(int(N1_),a) -> F *) | R_xml_0_rule_251 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_a nil)::nil)) (* eqt(int(N1_),excl) -> F *) | R_xml_0_rule_252 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_excl nil)::nil)) (* eqt(int(N1_),false) -> F *) | R_xml_0_rule_253 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* eqt(int(N1_),lock) -> F *) | R_xml_0_rule_254 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_lock nil)::nil)) (* eqt(int(N1_),locker) -> F *) | R_xml_0_rule_255 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_locker nil)::nil)) (* eqt(int(N1_),mcrlrecord) -> F *) | R_xml_0_rule_256 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)) (* eqt(int(N1_),ok) -> F *) | R_xml_0_rule_257 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_ok nil)::nil)) (* eqt(int(N1_),pending) -> F *) | R_xml_0_rule_258 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (* eqt(int(N1_),release) -> F *) | R_xml_0_rule_259 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_release nil)::nil)) (* eqt(int(N1_),request) -> F *) | R_xml_0_rule_260 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_request nil)::nil)) (* eqt(int(N1_),resource) -> F *) | R_xml_0_rule_261 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(int(N1_),tag) -> F *) | R_xml_0_rule_262 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_tag nil)::nil)) (* eqt(int(N1_),true) -> F *) | R_xml_0_rule_263 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* eqt(int(N1_),undefined) -> F *) | R_xml_0_rule_264 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(cons(H1_,T1_),resource) -> F *) | R_xml_0_rule_265 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(cons(H1_,T1_),tag) -> F *) | R_xml_0_rule_266 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_tag nil)::nil)) (* eqt(cons(H1_,T1_),true) -> F *) | R_xml_0_rule_267 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* eqt(cons(H1_,T1_),undefined) -> F *) | R_xml_0_rule_268 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(cons(H1_,T1_),pid(N2_)) -> F *) | R_xml_0_rule_269 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(cons(H1_,T1_),int(N2_)) -> F *) | R_xml_0_rule_270 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(cons(H1_,T1_),cons(H2_,T2_)) -> and(eqt(H1_,H2_),eqt(T1_,T2_)) *) | R_xml_0_rule_271 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 8):: (algebra.Alg.Var 5)::nil))::(algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 9):: (algebra.Alg.Var 6)::nil))::nil)) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(cons(H1_,T1_),tuple(H2_,T2_)) -> F *) | R_xml_0_rule_272 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(cons(H1_,T1_),tuplenil(H2_)) -> F *) | R_xml_0_rule_273 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* eqt(tuple(H1_,T1_),nil) -> F *) | R_xml_0_rule_274 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* eqt(tuple(H1_,T1_),a) -> F *) | R_xml_0_rule_275 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_a nil)::nil)) (* eqt(tuple(H1_,T1_),excl) -> F *) | R_xml_0_rule_276 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_excl nil)::nil)) (* eqt(tuple(H1_,T1_),false) -> F *) | R_xml_0_rule_277 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_false nil)::nil)) (* eqt(tuple(H1_,T1_),lock) -> F *) | R_xml_0_rule_278 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_lock nil)::nil)) (* eqt(tuple(H1_,T1_),locker) -> F *) | R_xml_0_rule_279 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_locker nil)::nil)) (* eqt(tuple(H1_,T1_),mcrlrecord) -> F *) | R_xml_0_rule_280 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)) (* eqt(tuple(H1_,T1_),ok) -> F *) | R_xml_0_rule_281 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_ok nil)::nil)) (* eqt(tuple(H1_,T1_),pending) -> F *) | R_xml_0_rule_282 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_pending nil)::nil)) (* eqt(tuple(H1_,T1_),release) -> F *) | R_xml_0_rule_283 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_release nil)::nil)) (* eqt(tuple(H1_,T1_),request) -> F *) | R_xml_0_rule_284 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_request nil)::nil)) (* eqt(tuple(H1_,T1_),resource) -> F *) | R_xml_0_rule_285 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(tuple(H1_,T1_),tag) -> F *) | R_xml_0_rule_286 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_tag nil)::nil)) (* eqt(tuple(H1_,T1_),true) -> F *) | R_xml_0_rule_287 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_true nil)::nil)) (* eqt(tuple(H1_,T1_),undefined) -> F *) | R_xml_0_rule_288 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(tuple(H1_,T1_),pid(N2_)) -> F *) | R_xml_0_rule_289 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(tuple(H1_,T1_),int(N2_)) -> F *) | R_xml_0_rule_290 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(tuple(H1_,T1_),cons(H2_,T2_)) -> F *) | R_xml_0_rule_291 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(tuple(H1_,T1_),tuple(H2_,T2_)) -> and(eqt(H1_,H2_),eqt(T1_,T2_)) *) | R_xml_0_rule_292 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 8):: (algebra.Alg.Var 5)::nil))::(algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 9):: (algebra.Alg.Var 6)::nil))::nil)) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(tuple(H1_,T1_),tuplenil(H2_)) -> F *) | R_xml_0_rule_293 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil)):: (algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* eqt(tuplenil(H1_),nil) -> F *) | R_xml_0_rule_294 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* eqt(tuplenil(H1_),a) -> F *) | R_xml_0_rule_295 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_a nil)::nil)) (* eqt(tuplenil(H1_),excl) -> F *) | R_xml_0_rule_296 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_excl nil)::nil)) (* eqt(tuplenil(H1_),false) -> F *) | R_xml_0_rule_297 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* eqt(tuplenil(H1_),lock) -> F *) | R_xml_0_rule_298 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_lock nil)::nil)) (* eqt(tuplenil(H1_),locker) -> F *) | R_xml_0_rule_299 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_locker nil)::nil)) (* eqt(tuplenil(H1_),mcrlrecord) -> F *) | R_xml_0_rule_300 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)) (* eqt(tuplenil(H1_),ok) -> F *) | R_xml_0_rule_301 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_ok nil)::nil)) (* eqt(tuplenil(H1_),pending) -> F *) | R_xml_0_rule_302 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (* eqt(tuplenil(H1_),release) -> F *) | R_xml_0_rule_303 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_release nil)::nil)) (* eqt(tuplenil(H1_),request) -> F *) | R_xml_0_rule_304 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_request nil)::nil)) (* eqt(tuplenil(H1_),resource) -> F *) | R_xml_0_rule_305 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* eqt(tuplenil(H1_),tag) -> F *) | R_xml_0_rule_306 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_tag nil)::nil)) (* eqt(tuplenil(H1_),true) -> F *) | R_xml_0_rule_307 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* eqt(tuplenil(H1_),undefined) -> F *) | R_xml_0_rule_308 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)) (* eqt(tuplenil(H1_),pid(N2_)) -> F *) | R_xml_0_rule_309 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(tuplenil(H1_),int(N2_)) -> F *) | R_xml_0_rule_310 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)) (* eqt(tuplenil(H1_),cons(H2_,T2_)) -> F *) | R_xml_0_rule_311 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(tuplenil(H1_),tuple(H2_,T2_)) -> F *) | R_xml_0_rule_312 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)) (* eqt(tuplenil(H1_),tuplenil(H2_)) -> eqt(H1_,H2_) *) | R_xml_0_rule_313 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 8):: (algebra.Alg.Var 5)::nil)) (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)) (* element(int(s(0)),tuplenil(T1_)) -> T1_ *) | R_xml_0_rule_314 : R_xml_0_rules (algebra.Alg.Var 9) (algebra.Alg.Term algebra.F.id_element ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_0 nil)::nil))::nil))::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 9)::nil))::nil)) (* element(int(s(0)),tuple(T1_,T2_)) -> T1_ *) | R_xml_0_rule_315 : R_xml_0_rules (algebra.Alg.Var 9) (algebra.Alg.Term algebra.F.id_element ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_0 nil)::nil))::nil))::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 9):: (algebra.Alg.Var 6)::nil))::nil)) (* element(int(s(s(N1_))),tuple(T1_,T2_)) -> element(int(s(N1_)),T2_) *) | R_xml_0_rule_316 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_element ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Var 7)::nil))::nil)):: (algebra.Alg.Var 6)::nil)) (algebra.Alg.Term algebra.F.id_element ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Var 7)::nil))::nil))::nil)):: (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 9):: (algebra.Alg.Var 6)::nil))::nil)) (* record_new(lock) -> tuple(mcrlrecord,tuple(lock,tuple(undefined,tuple(nil,tuplenil(nil))))) *) | R_xml_0_rule_317 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_undefined nil):: (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil))::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_record_new ((algebra.Alg.Term algebra.F.id_lock nil)::nil)) (* record_extract(tuple(mcrlrecord, tuple(lock,tuple(F0_,tuple(F1_,tuplenil(F2_))))),lock, resource) -> tuple(mcrlrecord,tuple(lock,tuple(F0_,tuple(F1_,tuplenil(F2_))))) *) | R_xml_0_rule_318 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 10)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 11):: (algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 12)::nil))::nil))::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_mcrlrecord nil):: (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 10)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 11)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 12)::nil))::nil))::nil))::nil))::nil)):: (algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (* record_update(tuple(mcrlrecord, tuple(lock,tuple(F0_,tuple(F1_,tuplenil(F2_))))),lock, pending,NewF_) -> tuple(mcrlrecord,tuple(lock,tuple(F0_,tuple(F1_,tuplenil(NewF_))))) *) | R_xml_0_rule_319 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 10)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 11):: (algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 13)::nil))::nil))::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_record_update ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_mcrlrecord nil):: (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 10)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 11)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 12)::nil))::nil))::nil))::nil))::nil)):: (algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Var 13)::nil)) (* record_updates(Record_,Name_,nil) -> Record_ *) | R_xml_0_rule_320 : R_xml_0_rules (algebra.Alg.Var 14) (algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Var 14):: (algebra.Alg.Var 15)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* record_updates(Record_,Name_,cons(tuple(Field_,tuplenil(NewF_)),Fields_)) -> record_updates(record_update(Record_,Name_,Field_,NewF_),Name_,Fields_) *) | R_xml_0_rule_321 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Term algebra.F.id_record_update ((algebra.Alg.Var 14)::(algebra.Alg.Var 15):: (algebra.Alg.Var 16)::(algebra.Alg.Var 13)::nil)):: (algebra.Alg.Var 15)::(algebra.Alg.Var 17)::nil)) (algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Var 14):: (algebra.Alg.Var 15)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 16):: (algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 13)::nil))::nil))::(algebra.Alg.Var 17)::nil))::nil)) (* locker2_map_promote_pending(nil,Pending_) -> nil *) | R_xml_0_rule_322 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_nil nil) (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Var 18)::nil)) (* locker2_map_promote_pending(cons(Lock_,Locks_),Pending_) -> cons(locker2_promote_pending(Lock_,Pending_), locker2_map_promote_pending(Locks_,Pending_)) *) | R_xml_0_rule_323 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_locker2_promote_pending ((algebra.Alg.Var 19)::(algebra.Alg.Var 18)::nil)):: (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending ((algebra.Alg.Var 20)::(algebra.Alg.Var 18)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 19):: (algebra.Alg.Var 20)::nil))::(algebra.Alg.Var 18)::nil)) (* locker2_map_claim_lock(nil,Resources_,Client_) -> nil *) | R_xml_0_rule_324 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_nil nil) (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Var 21):: (algebra.Alg.Var 22)::nil)) (* locker2_map_claim_lock(cons(Lock_,Locks_),Resources_,Client_) -> cons(locker2_claim_lock(Lock_,Resources_,Client_), locker2_map_claim_lock(Locks_,Resources_,Client_)) *) | R_xml_0_rule_325 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_locker2_claim_lock ((algebra.Alg.Var 19):: (algebra.Alg.Var 21)::(algebra.Alg.Var 22)::nil)):: (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock ((algebra.Alg.Var 20)::(algebra.Alg.Var 21):: (algebra.Alg.Var 22)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 19):: (algebra.Alg.Var 20)::nil))::(algebra.Alg.Var 21):: (algebra.Alg.Var 22)::nil)) (* locker2_map_add_pending(nil,Resources_,Client_) -> nil *) | R_xml_0_rule_326 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_nil nil) (algebra.Alg.Term algebra.F.id_locker2_map_add_pending ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Var 21):: (algebra.Alg.Var 22)::nil)) (* locker2_promote_pending(Lock_,Client_) -> case0(Client_,Lock_,record_extract(Lock_,lock,pending)) *) | R_xml_0_rule_327 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_case0 ((algebra.Alg.Var 22)::(algebra.Alg.Var 19):: (algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_promote_pending ((algebra.Alg.Var 19)::(algebra.Alg.Var 22)::nil)) (* case0(Client_,Lock_,cons(Client_,Pendings_)) -> record_updates(Lock_,lock, cons(tuple(excl,tuplenil(Client_)), cons(tuple(pending,tuplenil(Pendings_)),nil))) *) | R_xml_0_rule_328 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_excl nil):: (algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 22)::nil))::nil))::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_pending nil):: (algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 23)::nil))::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_case0 ((algebra.Alg.Var 22):: (algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 22)::(algebra.Alg.Var 23)::nil))::nil)) (* case0(Client_,Lock_,MCRLFree0_) -> Lock_ *) | R_xml_0_rule_329 : R_xml_0_rules (algebra.Alg.Var 19) (algebra.Alg.Term algebra.F.id_case0 ((algebra.Alg.Var 22):: (algebra.Alg.Var 19)::(algebra.Alg.Var 24)::nil)) (* locker2_remove_pending(Lock_,Client_) -> record_updates(Lock_,lock, cons(tuple(pending, tuplenil(subtract(record_extract(Lock_,lock,pending), cons(Client_,nil)))),nil)) *) | R_xml_0_rule_330 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_pending nil):: (algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Term algebra.F.id_subtract ((algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil))::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 22):: (algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil))::nil))::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_remove_pending ((algebra.Alg.Var 19)::(algebra.Alg.Var 22)::nil)) (* locker2_add_pending(Lock_,Resources_,Client_) -> case1(Client_,Resources_,Lock_, member(record_extract(Lock_,lock,resource),Resources_)) *) | R_xml_0_rule_331 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_case1 ((algebra.Alg.Var 22)::(algebra.Alg.Var 21):: (algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_member ((algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19):: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_resource nil)::nil)):: (algebra.Alg.Var 21)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_add_pending ((algebra.Alg.Var 19)::(algebra.Alg.Var 21):: (algebra.Alg.Var 22)::nil)) (* case1(Client_,Resources_,Lock_,true) -> record_updates(Lock_,lock, cons(tuple(pending, tuplenil(append(record_extract(Lock_,lock,pending),cons(Client_,nil)))), nil)) *) | R_xml_0_rule_332 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_pending nil):: (algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19):: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil)):: (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 22)::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil))::nil))::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_case1 ((algebra.Alg.Var 22):: (algebra.Alg.Var 21)::(algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* case1(Client_,Resources_,Lock_,false) -> Lock_ *) | R_xml_0_rule_333 : R_xml_0_rules (algebra.Alg.Var 19) (algebra.Alg.Term algebra.F.id_case1 ((algebra.Alg.Var 22):: (algebra.Alg.Var 21)::(algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* locker2_release_lock(Lock_,Client_) -> case2(Client_,Lock_,gen_modtageq(Client_,record_extract(Lock_,lock,excl))) *) | R_xml_0_rule_334 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_case2 ((algebra.Alg.Var 22)::(algebra.Alg.Var 19):: (algebra.Alg.Term algebra.F.id_gen_modtageq ((algebra.Alg.Var 22)::(algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19):: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_excl nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_release_lock ((algebra.Alg.Var 19)::(algebra.Alg.Var 22)::nil)) (* case2(Client_,Lock_,true) -> record_updates(Lock_,lock,cons(tuple(excllock,excl),nil)) *) | R_xml_0_rule_335 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_excllock nil):: (algebra.Alg.Term algebra.F.id_excl nil)::nil)):: (algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_case2 ((algebra.Alg.Var 22):: (algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* case4(Client_,Lock_,MCRLFree1_) -> false *) | R_xml_0_rule_336 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_false nil) (algebra.Alg.Term algebra.F.id_case4 ((algebra.Alg.Var 22):: (algebra.Alg.Var 19)::(algebra.Alg.Var 25)::nil)) (* locker2_obtainables(nil,Client_) -> true *) | R_xml_0_rule_337 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_true nil) (algebra.Alg.Term algebra.F.id_locker2_obtainables ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Var 22)::nil)) (* locker2_obtainables(cons(Lock_,Locks_),Client_) -> case5(Client_,Locks_,Lock_, member(Client_,record_extract(Lock_,lock,pending))) *) | R_xml_0_rule_338 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_case5 ((algebra.Alg.Var 22)::(algebra.Alg.Var 20):: (algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_member ((algebra.Alg.Var 22):: (algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_obtainables ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 19)::(algebra.Alg.Var 20)::nil)):: (algebra.Alg.Var 22)::nil)) (* case5(Client_,Locks_,Lock_,true) -> andt(locker2_obtainable(Lock_,Client_),locker2_obtainables(Locks_,Client_)) *) | R_xml_0_rule_339 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_andt ((algebra.Alg.Term algebra.F.id_locker2_obtainable ((algebra.Alg.Var 19):: (algebra.Alg.Var 22)::nil))::(algebra.Alg.Term algebra.F.id_locker2_obtainables ((algebra.Alg.Var 20):: (algebra.Alg.Var 22)::nil))::nil)) (algebra.Alg.Term algebra.F.id_case5 ((algebra.Alg.Var 22):: (algebra.Alg.Var 20)::(algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* case5(Client_,Locks_,Lock_,false) -> locker2_obtainables(Locks_,Client_) *) | R_xml_0_rule_340 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_locker2_obtainables ((algebra.Alg.Var 20)::(algebra.Alg.Var 22)::nil)) (algebra.Alg.Term algebra.F.id_case5 ((algebra.Alg.Var 22):: (algebra.Alg.Var 20)::(algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* locker2_check_available(Resource_,nil) -> false *) | R_xml_0_rule_341 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_false nil) (algebra.Alg.Term algebra.F.id_locker2_check_available ((algebra.Alg.Var 26)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* locker2_check_available(Resource_,cons(Lock_,Locks_)) -> case6(Locks_,Lock_,Resource_, equal(Resource_,record_extract(Lock_,lock,resource))) *) | R_xml_0_rule_342 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_case6 ((algebra.Alg.Var 20)::(algebra.Alg.Var 19):: (algebra.Alg.Var 26)::(algebra.Alg.Term algebra.F.id_equal ((algebra.Alg.Var 26):: (algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_check_available ((algebra.Alg.Var 26)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 19)::(algebra.Alg.Var 20)::nil))::nil)) (* case6(Locks_,Lock_,Resource_,true) -> andt(equal(record_extract(Lock_,lock,excl),nil), equal(record_extract(Lock_,lock,pending),nil)) *) | R_xml_0_rule_343 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_andt ((algebra.Alg.Term algebra.F.id_equal ((algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19):: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_excl nil)::nil)):: (algebra.Alg.Term algebra.F.id_nil nil)::nil)):: (algebra.Alg.Term algebra.F.id_equal ((algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19):: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil)):: (algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_case6 ((algebra.Alg.Var 20):: (algebra.Alg.Var 19)::(algebra.Alg.Var 26)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* case6(Locks_,Lock_,Resource_,false) -> locker2_check_available(Resource_,Locks_) *) | R_xml_0_rule_344 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_locker2_check_available ((algebra.Alg.Var 26)::(algebra.Alg.Var 20)::nil)) (algebra.Alg.Term algebra.F.id_case6 ((algebra.Alg.Var 20):: (algebra.Alg.Var 19)::(algebra.Alg.Var 26)::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* locker2_check_availables(nil,Locks_) -> true *) | R_xml_0_rule_345 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_true nil) (algebra.Alg.Term algebra.F.id_locker2_check_availables ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Var 20)::nil)) (* locker2_check_availables(cons(Resource_,Resources_),Locks_) -> andt(locker2_check_available(Resource_,Locks_), locker2_check_availables(Resources_,Locks_)) *) | R_xml_0_rule_346 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_andt ((algebra.Alg.Term algebra.F.id_locker2_check_available ((algebra.Alg.Var 26)::(algebra.Alg.Var 20)::nil)):: (algebra.Alg.Term algebra.F.id_locker2_check_availables ((algebra.Alg.Var 21)::(algebra.Alg.Var 20)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_check_availables ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 26):: (algebra.Alg.Var 21)::nil))::(algebra.Alg.Var 20)::nil)) (* locker2_adduniq(nil,List_) -> List_ *) | R_xml_0_rule_347 : R_xml_0_rules (algebra.Alg.Var 27) (algebra.Alg.Term algebra.F.id_locker2_adduniq ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Var 27)::nil)) (* append(cons(Head_,Tail_),List_) -> cons(Head_,append(Tail_,List_)) *) | R_xml_0_rule_348 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 28):: (algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Var 29)::(algebra.Alg.Var 27)::nil))::nil)) (algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 28)::(algebra.Alg.Var 29)::nil)):: (algebra.Alg.Var 27)::nil)) (* subtract(List_,nil) -> List_ *) | R_xml_0_rule_349 : R_xml_0_rules (algebra.Alg.Var 27) (algebra.Alg.Term algebra.F.id_subtract ((algebra.Alg.Var 27):: (algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* subtract(List_,cons(Head_,Tail_)) -> subtract(delete(Head_,List_),Tail_) *) | R_xml_0_rule_350 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_subtract ((algebra.Alg.Term algebra.F.id_delete ((algebra.Alg.Var 28):: (algebra.Alg.Var 27)::nil))::(algebra.Alg.Var 29)::nil)) (algebra.Alg.Term algebra.F.id_subtract ((algebra.Alg.Var 27):: (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 28):: (algebra.Alg.Var 29)::nil))::nil)) (* delete(E_,nil) -> nil *) | R_xml_0_rule_351 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_nil nil) (algebra.Alg.Term algebra.F.id_delete ((algebra.Alg.Var 30):: (algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* delete(E_,cons(Head_,Tail_)) -> case8(Tail_,Head_,E_,equal(E_,Head_)) *) | R_xml_0_rule_352 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_case8 ((algebra.Alg.Var 29)::(algebra.Alg.Var 28):: (algebra.Alg.Var 30)::(algebra.Alg.Term algebra.F.id_equal ((algebra.Alg.Var 30):: (algebra.Alg.Var 28)::nil))::nil)) (algebra.Alg.Term algebra.F.id_delete ((algebra.Alg.Var 30):: (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 28):: (algebra.Alg.Var 29)::nil))::nil)) (* case8(Tail_,Head_,E_,true) -> Tail_ *) | R_xml_0_rule_353 : R_xml_0_rules (algebra.Alg.Var 29) (algebra.Alg.Term algebra.F.id_case8 ((algebra.Alg.Var 29):: (algebra.Alg.Var 28)::(algebra.Alg.Var 30)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* case8(Tail_,Head_,E_,false) -> cons(Head_,delete(E_,Tail_)) *) | R_xml_0_rule_354 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 28):: (algebra.Alg.Term algebra.F.id_delete ((algebra.Alg.Var 30)::(algebra.Alg.Var 29)::nil))::nil)) (algebra.Alg.Term algebra.F.id_case8 ((algebra.Alg.Var 29):: (algebra.Alg.Var 28)::(algebra.Alg.Var 30)::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* gen_tag(Pid_) -> tuple(Pid_,tuplenil(tag)) *) | R_xml_0_rule_355 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 31)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Term algebra.F.id_tag nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_gen_tag ((algebra.Alg.Var 31)::nil)) (* gen_modtageq(Client1_,Client2_) -> equal(Client1_,Client2_) *) | R_xml_0_rule_356 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_equal ((algebra.Alg.Var 32)::(algebra.Alg.Var 33)::nil)) (algebra.Alg.Term algebra.F.id_gen_modtageq ((algebra.Alg.Var 32):: (algebra.Alg.Var 33)::nil)) (* member(E_,nil) -> false *) | R_xml_0_rule_357 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_false nil) (algebra.Alg.Term algebra.F.id_member ((algebra.Alg.Var 30):: (algebra.Alg.Term algebra.F.id_nil nil)::nil)) (* member(E_,cons(Head_,Tail_)) -> case9(Tail_,Head_,E_,equal(E_,Head_)) *) | R_xml_0_rule_358 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_case9 ((algebra.Alg.Var 29)::(algebra.Alg.Var 28):: (algebra.Alg.Var 30)::(algebra.Alg.Term algebra.F.id_equal ((algebra.Alg.Var 30):: (algebra.Alg.Var 28)::nil))::nil)) (algebra.Alg.Term algebra.F.id_member ((algebra.Alg.Var 30):: (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 28):: (algebra.Alg.Var 29)::nil))::nil)) (* case9(Tail_,Head_,E_,true) -> true *) | R_xml_0_rule_359 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_true nil) (algebra.Alg.Term algebra.F.id_case9 ((algebra.Alg.Var 29):: (algebra.Alg.Var 28)::(algebra.Alg.Var 30)::(algebra.Alg.Term algebra.F.id_true nil)::nil)) (* case9(Tail_,Head_,E_,false) -> member(E_,Tail_) *) | R_xml_0_rule_360 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_member ((algebra.Alg.Var 30)::(algebra.Alg.Var 29)::nil)) (algebra.Alg.Term algebra.F.id_case9 ((algebra.Alg.Var 29):: (algebra.Alg.Var 28)::(algebra.Alg.Var 30)::(algebra.Alg.Term algebra.F.id_false nil)::nil)) (* eqs(empty,empty) -> T *) | R_xml_0_rule_361 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_eqs ((algebra.Alg.Term algebra.F.id_empty nil)::(algebra.Alg.Term algebra.F.id_empty nil)::nil)) (* eqs(empty,stack(E2_,S2_)) -> F *) | R_xml_0_rule_362 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqs ((algebra.Alg.Term algebra.F.id_empty nil)::(algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 34)::(algebra.Alg.Var 35)::nil))::nil)) (* eqs(stack(E1_,S1_),empty) -> F *) | R_xml_0_rule_363 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqs ((algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 36)::(algebra.Alg.Var 37)::nil)):: (algebra.Alg.Term algebra.F.id_empty nil)::nil)) (* eqs(stack(E1_,S1_),stack(E2_,S2_)) -> and(eqt(E1_,E2_),eqs(S1_,S2_)) *) | R_xml_0_rule_364 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 36):: (algebra.Alg.Var 34)::nil))::(algebra.Alg.Term algebra.F.id_eqs ((algebra.Alg.Var 37):: (algebra.Alg.Var 35)::nil))::nil)) (algebra.Alg.Term algebra.F.id_eqs ((algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 36)::(algebra.Alg.Var 37)::nil)):: (algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 34):: (algebra.Alg.Var 35)::nil))::nil)) (* pushs(E1_,S1_) -> stack(E1_,S1_) *) | R_xml_0_rule_365 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 36)::(algebra.Alg.Var 37)::nil)) (algebra.Alg.Term algebra.F.id_pushs ((algebra.Alg.Var 36):: (algebra.Alg.Var 37)::nil)) (* pops(stack(E1_,S1_)) -> S1_ *) | R_xml_0_rule_366 : R_xml_0_rules (algebra.Alg.Var 37) (algebra.Alg.Term algebra.F.id_pops ((algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 36):: (algebra.Alg.Var 37)::nil))::nil)) (* tops(stack(E1_,S1_)) -> E1_ *) | R_xml_0_rule_367 : R_xml_0_rules (algebra.Alg.Var 36) (algebra.Alg.Term algebra.F.id_tops ((algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 36):: (algebra.Alg.Var 37)::nil))::nil)) (* istops(E1_,empty) -> F *) | R_xml_0_rule_368 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_istops ((algebra.Alg.Var 36):: (algebra.Alg.Term algebra.F.id_empty nil)::nil)) (* istops(E1_,stack(E2_,S1_)) -> eqt(E1_,E2_) *) | R_xml_0_rule_369 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 36):: (algebra.Alg.Var 34)::nil)) (algebra.Alg.Term algebra.F.id_istops ((algebra.Alg.Var 36):: (algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 34):: (algebra.Alg.Var 37)::nil))::nil)) (* eqc(nocalls,nocalls) -> T *) | R_xml_0_rule_370 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_T nil) (algebra.Alg.Term algebra.F.id_eqc ((algebra.Alg.Term algebra.F.id_nocalls nil)::(algebra.Alg.Term algebra.F.id_nocalls nil)::nil)) (* eqc(nocalls,calls(E2_,S2_,CS2_)) -> F *) | R_xml_0_rule_371 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqc ((algebra.Alg.Term algebra.F.id_nocalls nil)::(algebra.Alg.Term algebra.F.id_calls ((algebra.Alg.Var 34)::(algebra.Alg.Var 35):: (algebra.Alg.Var 38)::nil))::nil)) (* eqc(calls(E1_,S1_,CS1_),nocalls) -> F *) | R_xml_0_rule_372 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_F nil) (algebra.Alg.Term algebra.F.id_eqc ((algebra.Alg.Term algebra.F.id_calls ((algebra.Alg.Var 36)::(algebra.Alg.Var 37):: (algebra.Alg.Var 39)::nil))::(algebra.Alg.Term algebra.F.id_nocalls nil)::nil)) (* eqc(calls(E1_,S1_,CS1_),calls(E2_,S2_,CS2_)) -> and(eqt(E1_,E2_),and(eqs(S1_,S2_),eqc(CS1_,CS2_))) *) | R_xml_0_rule_373 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 36):: (algebra.Alg.Var 34)::nil))::(algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqs ((algebra.Alg.Var 37)::(algebra.Alg.Var 35)::nil)):: (algebra.Alg.Term algebra.F.id_eqc ((algebra.Alg.Var 39):: (algebra.Alg.Var 38)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_eqc ((algebra.Alg.Term algebra.F.id_calls ((algebra.Alg.Var 36)::(algebra.Alg.Var 37):: (algebra.Alg.Var 39)::nil))::(algebra.Alg.Term algebra.F.id_calls ((algebra.Alg.Var 34)::(algebra.Alg.Var 35):: (algebra.Alg.Var 38)::nil))::nil)) (* push(E1_,E2_,nocalls) -> calls(E1_,stack(E2_,empty),nocalls) *) | R_xml_0_rule_374 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_calls ((algebra.Alg.Var 36)::(algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 34):: (algebra.Alg.Term algebra.F.id_empty nil)::nil)):: (algebra.Alg.Term algebra.F.id_nocalls nil)::nil)) (algebra.Alg.Term algebra.F.id_push ((algebra.Alg.Var 36):: (algebra.Alg.Var 34)::(algebra.Alg.Term algebra.F.id_nocalls nil)::nil)) (* push(E1_,E2_,calls(E3_,S1_,CS1_)) -> push1(E1_,E2_,E3_,S1_,CS1_,eqt(E1_,E3_)) *) | R_xml_0_rule_375 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_push1 ((algebra.Alg.Var 36)::(algebra.Alg.Var 34):: (algebra.Alg.Var 40)::(algebra.Alg.Var 37):: (algebra.Alg.Var 39)::(algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 36)::(algebra.Alg.Var 40)::nil))::nil)) (algebra.Alg.Term algebra.F.id_push ((algebra.Alg.Var 36):: (algebra.Alg.Var 34)::(algebra.Alg.Term algebra.F.id_calls ((algebra.Alg.Var 40)::(algebra.Alg.Var 37):: (algebra.Alg.Var 39)::nil))::nil)) (* push1(E1_,E2_,E3_,S1_,CS1_,T) -> calls(E3_,pushs(E2_,S1_),CS1_) *) | R_xml_0_rule_376 : R_xml_0_rules (algebra.Alg.Term algebra.F.id_calls ((algebra.Alg.Var 40)::(algebra.Alg.Term algebra.F.id_pushs ((algebra.Alg.Var 34):: (algebra.Alg.Var 37)::nil))::(algebra.Alg.Var 39)::nil)) (algebra.Alg.Term algebra.F.id_push1 ((algebra.Alg.Var 36):: (algebra.Alg.Var 34)::(algebra.Alg.Var 40)::(algebra.Alg.Var 37):: (algebra.Alg.Var 39)::(algebra.Alg.Term algebra.F.id_T nil)::nil)) . Definition R_xml_0_rule_as_list_0 := ((algebra.Alg.Term algebra.F.id_or ((algebra.Alg.Term algebra.F.id_T nil)::(algebra.Alg.Term algebra.F.id_T nil)::nil)), (algebra.Alg.Term algebra.F.id_T nil))::nil. Definition R_xml_0_rule_as_list_1 := ((algebra.Alg.Term algebra.F.id_or ((algebra.Alg.Term algebra.F.id_F nil)::(algebra.Alg.Term algebra.F.id_T nil)::nil)), (algebra.Alg.Term algebra.F.id_T nil))::R_xml_0_rule_as_list_0. Definition R_xml_0_rule_as_list_2 := ((algebra.Alg.Term algebra.F.id_or ((algebra.Alg.Term algebra.F.id_T nil)::(algebra.Alg.Term algebra.F.id_F nil)::nil)), (algebra.Alg.Term algebra.F.id_T nil))::R_xml_0_rule_as_list_1. Definition R_xml_0_rule_as_list_3 := ((algebra.Alg.Term algebra.F.id_or ((algebra.Alg.Term algebra.F.id_F nil)::(algebra.Alg.Term algebra.F.id_F nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_2. Definition R_xml_0_rule_as_list_4 := ((algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_T nil)::(algebra.Alg.Var 1)::nil)),(algebra.Alg.Var 1)):: R_xml_0_rule_as_list_3. Definition R_xml_0_rule_as_list_5 := ((algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id_T nil)::nil)),(algebra.Alg.Var 1)):: R_xml_0_rule_as_list_4. Definition R_xml_0_rule_as_list_6 := ((algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_F nil)::(algebra.Alg.Var 1)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_5. Definition R_xml_0_rule_as_list_7 := ((algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Var 1):: (algebra.Alg.Term algebra.F.id_F nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_6. Definition R_xml_0_rule_as_list_8 := ((algebra.Alg.Term algebra.F.id_imp ((algebra.Alg.Term algebra.F.id_T nil)::(algebra.Alg.Var 1)::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_imp ((algebra.Alg.Term algebra.F.id_F nil)::(algebra.Alg.Var 1)::nil)),(algebra.Alg.Term algebra.F.id_T nil)):: R_xml_0_rule_as_list_8. Definition R_xml_0_rule_as_list_10 := ((algebra.Alg.Term algebra.F.id_not ((algebra.Alg.Term algebra.F.id_T nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_9. Definition R_xml_0_rule_as_list_11 := ((algebra.Alg.Term algebra.F.id_not ((algebra.Alg.Term algebra.F.id_F nil)::nil)),(algebra.Alg.Term algebra.F.id_T nil)):: R_xml_0_rule_as_list_10. Definition R_xml_0_rule_as_list_12 := ((algebra.Alg.Term algebra.F.id_if ((algebra.Alg.Term algebra.F.id_T nil)::(algebra.Alg.Var 2)::(algebra.Alg.Var 3)::nil)), (algebra.Alg.Var 2))::R_xml_0_rule_as_list_11. Definition R_xml_0_rule_as_list_13 := ((algebra.Alg.Term algebra.F.id_if ((algebra.Alg.Term algebra.F.id_F nil)::(algebra.Alg.Var 2)::(algebra.Alg.Var 3)::nil)), (algebra.Alg.Var 3))::R_xml_0_rule_as_list_12. Definition R_xml_0_rule_as_list_14 := ((algebra.Alg.Term algebra.F.id_eq ((algebra.Alg.Term algebra.F.id_T nil)::(algebra.Alg.Term algebra.F.id_T nil)::nil)), (algebra.Alg.Term algebra.F.id_T nil))::R_xml_0_rule_as_list_13. Definition R_xml_0_rule_as_list_15 := ((algebra.Alg.Term algebra.F.id_eq ((algebra.Alg.Term algebra.F.id_F nil)::(algebra.Alg.Term algebra.F.id_F nil)::nil)), (algebra.Alg.Term algebra.F.id_T nil))::R_xml_0_rule_as_list_14. Definition R_xml_0_rule_as_list_16 := ((algebra.Alg.Term algebra.F.id_eq ((algebra.Alg.Term algebra.F.id_T nil)::(algebra.Alg.Term algebra.F.id_F nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_15. Definition R_xml_0_rule_as_list_17 := ((algebra.Alg.Term algebra.F.id_eq ((algebra.Alg.Term algebra.F.id_F nil)::(algebra.Alg.Term algebra.F.id_T nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_16. Definition R_xml_0_rule_as_list_18 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_17. Definition R_xml_0_rule_as_list_19 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_18. Definition R_xml_0_rule_as_list_20 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_19. Definition R_xml_0_rule_as_list_21 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_20. Definition R_xml_0_rule_as_list_22 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_21. Definition R_xml_0_rule_as_list_23 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_22. Definition R_xml_0_rule_as_list_24 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_23. Definition R_xml_0_rule_as_list_25 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)), (algebra.Alg.Term algebra.F.id_T nil))::R_xml_0_rule_as_list_24. Definition R_xml_0_rule_as_list_26 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_25. Definition R_xml_0_rule_as_list_27 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_26. Definition R_xml_0_rule_as_list_28 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_27. Definition R_xml_0_rule_as_list_29 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_28. Definition R_xml_0_rule_as_list_30 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_29. Definition R_xml_0_rule_as_list_31 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_30. Definition R_xml_0_rule_as_list_32 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_31. Definition R_xml_0_rule_as_list_33 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_32. Definition R_xml_0_rule_as_list_34 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_33. Definition R_xml_0_rule_as_list_35 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_34. Definition R_xml_0_rule_as_list_36 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_35. Definition R_xml_0_rule_as_list_37 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_36. Definition R_xml_0_rule_as_list_38 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_37. Definition R_xml_0_rule_as_list_39 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_38. Definition R_xml_0_rule_as_list_40 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_39. Definition R_xml_0_rule_as_list_41 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_40. Definition R_xml_0_rule_as_list_42 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_41. Definition R_xml_0_rule_as_list_43 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_a nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_42. Definition R_xml_0_rule_as_list_44 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_43. Definition R_xml_0_rule_as_list_45 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_44. Definition R_xml_0_rule_as_list_46 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)), (algebra.Alg.Term algebra.F.id_T nil))::R_xml_0_rule_as_list_45. Definition R_xml_0_rule_as_list_47 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_46. Definition R_xml_0_rule_as_list_48 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_47. Definition R_xml_0_rule_as_list_49 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_48. Definition R_xml_0_rule_as_list_50 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_49. Definition R_xml_0_rule_as_list_51 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_50. Definition R_xml_0_rule_as_list_52 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_51. Definition R_xml_0_rule_as_list_53 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_52. Definition R_xml_0_rule_as_list_54 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_53. Definition R_xml_0_rule_as_list_55 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_54. Definition R_xml_0_rule_as_list_56 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_55. Definition R_xml_0_rule_as_list_57 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_56. Definition R_xml_0_rule_as_list_58 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_57. Definition R_xml_0_rule_as_list_59 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_58. Definition R_xml_0_rule_as_list_60 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_false nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_59. Definition R_xml_0_rule_as_list_61 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_false nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_60. Definition R_xml_0_rule_as_list_62 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_false nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_61. Definition R_xml_0_rule_as_list_63 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_false nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_62. Definition R_xml_0_rule_as_list_64 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_63. Definition R_xml_0_rule_as_list_65 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_64. Definition R_xml_0_rule_as_list_66 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_65. Definition R_xml_0_rule_as_list_67 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_66. Definition R_xml_0_rule_as_list_68 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)), (algebra.Alg.Term algebra.F.id_T nil))::R_xml_0_rule_as_list_67. Definition R_xml_0_rule_as_list_69 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_68. Definition R_xml_0_rule_as_list_70 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_69. Definition R_xml_0_rule_as_list_71 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_70. Definition R_xml_0_rule_as_list_72 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_71. Definition R_xml_0_rule_as_list_73 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_72. Definition R_xml_0_rule_as_list_74 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_73. Definition R_xml_0_rule_as_list_75 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_74. Definition R_xml_0_rule_as_list_76 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_75. Definition R_xml_0_rule_as_list_77 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_76. Definition R_xml_0_rule_as_list_78 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_77. Definition R_xml_0_rule_as_list_79 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_78. Definition R_xml_0_rule_as_list_80 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_79. Definition R_xml_0_rule_as_list_81 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_80. Definition R_xml_0_rule_as_list_82 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_81. Definition R_xml_0_rule_as_list_83 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_82. Definition R_xml_0_rule_as_list_84 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_83. Definition R_xml_0_rule_as_list_85 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_84. Definition R_xml_0_rule_as_list_86 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_85. Definition R_xml_0_rule_as_list_87 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_86. Definition R_xml_0_rule_as_list_88 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_87. Definition R_xml_0_rule_as_list_89 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)),(algebra.Alg.Term algebra.F.id_T nil)):: R_xml_0_rule_as_list_88. Definition R_xml_0_rule_as_list_90 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_89. Definition R_xml_0_rule_as_list_91 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_90. Definition R_xml_0_rule_as_list_92 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_91. Definition R_xml_0_rule_as_list_93 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_92. Definition R_xml_0_rule_as_list_94 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_93. Definition R_xml_0_rule_as_list_95 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_94. Definition R_xml_0_rule_as_list_96 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_95. Definition R_xml_0_rule_as_list_97 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_96. Definition R_xml_0_rule_as_list_98 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_97. Definition R_xml_0_rule_as_list_99 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_98. Definition R_xml_0_rule_as_list_100 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_99. Definition R_xml_0_rule_as_list_101 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_100. Definition R_xml_0_rule_as_list_102 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_101. Definition R_xml_0_rule_as_list_103 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_locker nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_102. Definition R_xml_0_rule_as_list_104 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_103. Definition R_xml_0_rule_as_list_105 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_104. Definition R_xml_0_rule_as_list_106 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_105. Definition R_xml_0_rule_as_list_107 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_106. Definition R_xml_0_rule_as_list_108 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_107. Definition R_xml_0_rule_as_list_109 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_108. Definition R_xml_0_rule_as_list_110 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)),(algebra.Alg.Term algebra.F.id_T nil)):: R_xml_0_rule_as_list_109. Definition R_xml_0_rule_as_list_111 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_110. Definition R_xml_0_rule_as_list_112 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_111. Definition R_xml_0_rule_as_list_113 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_112. Definition R_xml_0_rule_as_list_114 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_113. Definition R_xml_0_rule_as_list_115 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_114. Definition R_xml_0_rule_as_list_116 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_115. Definition R_xml_0_rule_as_list_117 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_116. Definition R_xml_0_rule_as_list_118 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_117. Definition R_xml_0_rule_as_list_119 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_118. Definition R_xml_0_rule_as_list_120 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_119. Definition R_xml_0_rule_as_list_121 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_120. Definition R_xml_0_rule_as_list_122 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_121. Definition R_xml_0_rule_as_list_123 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_122. Definition R_xml_0_rule_as_list_124 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_ok nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_123. Definition R_xml_0_rule_as_list_125 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_124. Definition R_xml_0_rule_as_list_126 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_125. Definition R_xml_0_rule_as_list_127 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_126. Definition R_xml_0_rule_as_list_128 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_127. Definition R_xml_0_rule_as_list_129 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_128. Definition R_xml_0_rule_as_list_130 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_129. Definition R_xml_0_rule_as_list_131 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_130. Definition R_xml_0_rule_as_list_132 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_131. Definition R_xml_0_rule_as_list_133 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)),(algebra.Alg.Term algebra.F.id_T nil)):: R_xml_0_rule_as_list_132. Definition R_xml_0_rule_as_list_134 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_133. Definition R_xml_0_rule_as_list_135 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_134. Definition R_xml_0_rule_as_list_136 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_135. Definition R_xml_0_rule_as_list_137 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_136. Definition R_xml_0_rule_as_list_138 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_137. Definition R_xml_0_rule_as_list_139 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_138. Definition R_xml_0_rule_as_list_140 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_139. Definition R_xml_0_rule_as_list_141 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_140. Definition R_xml_0_rule_as_list_142 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_141. Definition R_xml_0_rule_as_list_143 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_142. Definition R_xml_0_rule_as_list_144 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_143. Definition R_xml_0_rule_as_list_145 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_144. Definition R_xml_0_rule_as_list_146 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_145. Definition R_xml_0_rule_as_list_147 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_146. Definition R_xml_0_rule_as_list_148 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_147. Definition R_xml_0_rule_as_list_149 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_148. Definition R_xml_0_rule_as_list_150 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_149. Definition R_xml_0_rule_as_list_151 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_150. Definition R_xml_0_rule_as_list_152 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_release nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_151. Definition R_xml_0_rule_as_list_153 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_152. Definition R_xml_0_rule_as_list_154 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_153. Definition R_xml_0_rule_as_list_155 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_154. Definition R_xml_0_rule_as_list_156 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_155. Definition R_xml_0_rule_as_list_157 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)),(algebra.Alg.Term algebra.F.id_T nil)):: R_xml_0_rule_as_list_156. Definition R_xml_0_rule_as_list_158 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_157. Definition R_xml_0_rule_as_list_159 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_158. Definition R_xml_0_rule_as_list_160 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_159. Definition R_xml_0_rule_as_list_161 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_160. Definition R_xml_0_rule_as_list_162 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_161. Definition R_xml_0_rule_as_list_163 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_162. Definition R_xml_0_rule_as_list_164 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_163. Definition R_xml_0_rule_as_list_165 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_164. Definition R_xml_0_rule_as_list_166 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_request nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_165. Definition R_xml_0_rule_as_list_167 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_166. Definition R_xml_0_rule_as_list_168 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_167. Definition R_xml_0_rule_as_list_169 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_168. Definition R_xml_0_rule_as_list_170 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_169. Definition R_xml_0_rule_as_list_171 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_170. Definition R_xml_0_rule_as_list_172 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_171. Definition R_xml_0_rule_as_list_173 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_172. Definition R_xml_0_rule_as_list_174 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_173. Definition R_xml_0_rule_as_list_175 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_174. Definition R_xml_0_rule_as_list_176 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_175. Definition R_xml_0_rule_as_list_177 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_176. Definition R_xml_0_rule_as_list_178 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)),(algebra.Alg.Term algebra.F.id_T nil)):: R_xml_0_rule_as_list_177. Definition R_xml_0_rule_as_list_179 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_178. Definition R_xml_0_rule_as_list_180 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_179. Definition R_xml_0_rule_as_list_181 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_180. Definition R_xml_0_rule_as_list_182 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_181. Definition R_xml_0_rule_as_list_183 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_182. Definition R_xml_0_rule_as_list_184 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_183. Definition R_xml_0_rule_as_list_185 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_184. Definition R_xml_0_rule_as_list_186 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_resource nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_185. Definition R_xml_0_rule_as_list_187 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_186. Definition R_xml_0_rule_as_list_188 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_187. Definition R_xml_0_rule_as_list_189 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_188. Definition R_xml_0_rule_as_list_190 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_189. Definition R_xml_0_rule_as_list_191 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_190. Definition R_xml_0_rule_as_list_192 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_191. Definition R_xml_0_rule_as_list_193 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_192. Definition R_xml_0_rule_as_list_194 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_193. Definition R_xml_0_rule_as_list_195 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_194. Definition R_xml_0_rule_as_list_196 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_195. Definition R_xml_0_rule_as_list_197 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_196. Definition R_xml_0_rule_as_list_198 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_197. Definition R_xml_0_rule_as_list_199 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)), (algebra.Alg.Term algebra.F.id_T nil))::R_xml_0_rule_as_list_198. Definition R_xml_0_rule_as_list_200 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_199. Definition R_xml_0_rule_as_list_201 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_200. Definition R_xml_0_rule_as_list_202 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_201. Definition R_xml_0_rule_as_list_203 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_202. Definition R_xml_0_rule_as_list_204 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_203. Definition R_xml_0_rule_as_list_205 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_204. Definition R_xml_0_rule_as_list_206 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tag nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_205. Definition R_xml_0_rule_as_list_207 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_206. Definition R_xml_0_rule_as_list_208 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_207. Definition R_xml_0_rule_as_list_209 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_208. Definition R_xml_0_rule_as_list_210 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_false nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_209. Definition R_xml_0_rule_as_list_211 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_lock nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_210. Definition R_xml_0_rule_as_list_212 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_locker nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_211. Definition R_xml_0_rule_as_list_213 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_212. Definition R_xml_0_rule_as_list_214 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_ok nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_213. Definition R_xml_0_rule_as_list_215 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_214. Definition R_xml_0_rule_as_list_216 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_release nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_215. Definition R_xml_0_rule_as_list_217 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_request nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_216. Definition R_xml_0_rule_as_list_218 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_217. Definition R_xml_0_rule_as_list_219 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_tag nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_218. Definition R_xml_0_rule_as_list_220 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_true nil)::nil)), (algebra.Alg.Term algebra.F.id_T nil))::R_xml_0_rule_as_list_219. Definition R_xml_0_rule_as_list_221 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_220. Definition R_xml_0_rule_as_list_222 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_221. Definition R_xml_0_rule_as_list_223 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_222. Definition R_xml_0_rule_as_list_224 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_223. Definition R_xml_0_rule_as_list_225 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_224. Definition R_xml_0_rule_as_list_226 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_true nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_225. Definition R_xml_0_rule_as_list_227 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_undefined nil)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_226. Definition R_xml_0_rule_as_list_228 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_undefined nil)::(algebra.Alg.Term algebra.F.id_a nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_227. Definition R_xml_0_rule_as_list_229 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_undefined nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_228. Definition R_xml_0_rule_as_list_230 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_229. Definition R_xml_0_rule_as_list_231 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_a nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_230. Definition R_xml_0_rule_as_list_232 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_excl nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_231. Definition R_xml_0_rule_as_list_233 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_false nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_232. Definition R_xml_0_rule_as_list_234 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_lock nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_233. Definition R_xml_0_rule_as_list_235 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_locker nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_234. Definition R_xml_0_rule_as_list_236 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_235. Definition R_xml_0_rule_as_list_237 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_ok nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_236. Definition R_xml_0_rule_as_list_238 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_pending nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_237. Definition R_xml_0_rule_as_list_239 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_release nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_238. Definition R_xml_0_rule_as_list_240 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_request nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_239. Definition R_xml_0_rule_as_list_241 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_resource nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_240. Definition R_xml_0_rule_as_list_242 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_tag nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_241. Definition R_xml_0_rule_as_list_243 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_true nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_242. Definition R_xml_0_rule_as_list_244 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_243. Definition R_xml_0_rule_as_list_245 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 7):: (algebra.Alg.Var 4)::nil)))::R_xml_0_rule_as_list_244. Definition R_xml_0_rule_as_list_246 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_245. Definition R_xml_0_rule_as_list_247 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_246. Definition R_xml_0_rule_as_list_248 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5)::(algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_247. Definition R_xml_0_rule_as_list_249 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_248. Definition R_xml_0_rule_as_list_250 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_249. Definition R_xml_0_rule_as_list_251 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_a nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_250. Definition R_xml_0_rule_as_list_252 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_excl nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_251. Definition R_xml_0_rule_as_list_253 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_false nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_252. Definition R_xml_0_rule_as_list_254 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_lock nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_253. Definition R_xml_0_rule_as_list_255 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_locker nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_254. Definition R_xml_0_rule_as_list_256 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_255. Definition R_xml_0_rule_as_list_257 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_ok nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_256. Definition R_xml_0_rule_as_list_258 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_pending nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_257. Definition R_xml_0_rule_as_list_259 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_release nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_258. Definition R_xml_0_rule_as_list_260 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_request nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_259. Definition R_xml_0_rule_as_list_261 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_resource nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_260. Definition R_xml_0_rule_as_list_262 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_tag nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_261. Definition R_xml_0_rule_as_list_263 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_true nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_262. Definition R_xml_0_rule_as_list_264 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 7)::nil))::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_263. Definition R_xml_0_rule_as_list_265 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_resource nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_264. Definition R_xml_0_rule_as_list_266 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_tag nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_265. Definition R_xml_0_rule_as_list_267 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_true nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_266. Definition R_xml_0_rule_as_list_268 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_267. Definition R_xml_0_rule_as_list_269 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_268. Definition R_xml_0_rule_as_list_270 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_269. Definition R_xml_0_rule_as_list_271 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 8)::(algebra.Alg.Var 5)::nil))::(algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 9)::(algebra.Alg.Var 6)::nil))::nil))):: R_xml_0_rule_as_list_270. Definition R_xml_0_rule_as_list_272 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_271. Definition R_xml_0_rule_as_list_273 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_272. Definition R_xml_0_rule_as_list_274 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_273. Definition R_xml_0_rule_as_list_275 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_a nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_274. Definition R_xml_0_rule_as_list_276 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_excl nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_275. Definition R_xml_0_rule_as_list_277 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_false nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_276. Definition R_xml_0_rule_as_list_278 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_lock nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_277. Definition R_xml_0_rule_as_list_279 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_locker nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_278. Definition R_xml_0_rule_as_list_280 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_279. Definition R_xml_0_rule_as_list_281 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_ok nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_280. Definition R_xml_0_rule_as_list_282 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_pending nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_281. Definition R_xml_0_rule_as_list_283 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_release nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_282. Definition R_xml_0_rule_as_list_284 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_request nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_283. Definition R_xml_0_rule_as_list_285 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_resource nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_284. Definition R_xml_0_rule_as_list_286 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_tag nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_285. Definition R_xml_0_rule_as_list_287 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_true nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_286. Definition R_xml_0_rule_as_list_288 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_287. Definition R_xml_0_rule_as_list_289 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_288. Definition R_xml_0_rule_as_list_290 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_289. Definition R_xml_0_rule_as_list_291 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_290. Definition R_xml_0_rule_as_list_292 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 8)::(algebra.Alg.Var 5)::nil))::(algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 9)::(algebra.Alg.Var 6)::nil))::nil))):: R_xml_0_rule_as_list_291. Definition R_xml_0_rule_as_list_293 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 8)::(algebra.Alg.Var 9)::nil))::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_292. Definition R_xml_0_rule_as_list_294 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_293. Definition R_xml_0_rule_as_list_295 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_a nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_294. Definition R_xml_0_rule_as_list_296 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_excl nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_295. Definition R_xml_0_rule_as_list_297 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_false nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_296. Definition R_xml_0_rule_as_list_298 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_lock nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_297. Definition R_xml_0_rule_as_list_299 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_locker nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_298. Definition R_xml_0_rule_as_list_300 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_mcrlrecord nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_299. Definition R_xml_0_rule_as_list_301 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_ok nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_300. Definition R_xml_0_rule_as_list_302 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_pending nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_301. Definition R_xml_0_rule_as_list_303 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_release nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_302. Definition R_xml_0_rule_as_list_304 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_request nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_303. Definition R_xml_0_rule_as_list_305 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_resource nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_304. Definition R_xml_0_rule_as_list_306 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_tag nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_305. Definition R_xml_0_rule_as_list_307 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_true nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_306. Definition R_xml_0_rule_as_list_308 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_undefined nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_307. Definition R_xml_0_rule_as_list_309 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_pid ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_308. Definition R_xml_0_rule_as_list_310 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Var 4)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_309. Definition R_xml_0_rule_as_list_311 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_310. Definition R_xml_0_rule_as_list_312 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 5):: (algebra.Alg.Var 6)::nil))::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_311. Definition R_xml_0_rule_as_list_313 := ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 8)::nil))::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 5)::nil))::nil)), (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 8):: (algebra.Alg.Var 5)::nil)))::R_xml_0_rule_as_list_312. Definition R_xml_0_rule_as_list_314 := ((algebra.Alg.Term algebra.F.id_element ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_0 nil)::nil))::nil))::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 9)::nil))::nil)), (algebra.Alg.Var 9))::R_xml_0_rule_as_list_313. Definition R_xml_0_rule_as_list_315 := ((algebra.Alg.Term algebra.F.id_element ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_0 nil)::nil))::nil))::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 9)::(algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Var 9))::R_xml_0_rule_as_list_314. Definition R_xml_0_rule_as_list_316 := ((algebra.Alg.Term algebra.F.id_element ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Var 7)::nil))::nil))::nil)):: (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 9):: (algebra.Alg.Var 6)::nil))::nil)), (algebra.Alg.Term algebra.F.id_element ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Var 7)::nil))::nil))::(algebra.Alg.Var 6)::nil))):: R_xml_0_rule_as_list_315. Definition R_xml_0_rule_as_list_317 := ((algebra.Alg.Term algebra.F.id_record_new ((algebra.Alg.Term algebra.F.id_lock nil)::nil)), (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_undefined nil):: (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil))::nil))::nil))::nil)))::R_xml_0_rule_as_list_316. Definition R_xml_0_rule_as_list_318 := ((algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_mcrlrecord nil):: (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 10)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 11)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 12)::nil))::nil))::nil))::nil))::nil)):: (algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)), (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 10)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 11)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 12)::nil))::nil))::nil))::nil))::nil))):: R_xml_0_rule_as_list_317. Definition R_xml_0_rule_as_list_319 := ((algebra.Alg.Term algebra.F.id_record_update ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_mcrlrecord nil):: (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 10)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 11)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 12)::nil))::nil))::nil))::nil))::nil)):: (algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Var 13)::nil)), (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_mcrlrecord nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 10)::(algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 11)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 13)::nil))::nil))::nil))::nil))::nil))):: R_xml_0_rule_as_list_318. Definition R_xml_0_rule_as_list_320 := ((algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Var 14):: (algebra.Alg.Var 15)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)), (algebra.Alg.Var 14))::R_xml_0_rule_as_list_319. Definition R_xml_0_rule_as_list_321 := ((algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Var 14):: (algebra.Alg.Var 15)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 16):: (algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 13)::nil))::nil))::(algebra.Alg.Var 17)::nil))::nil)), (algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Term algebra.F.id_record_update ((algebra.Alg.Var 14)::(algebra.Alg.Var 15):: (algebra.Alg.Var 16)::(algebra.Alg.Var 13)::nil)):: (algebra.Alg.Var 15)::(algebra.Alg.Var 17)::nil))):: R_xml_0_rule_as_list_320. Definition R_xml_0_rule_as_list_322 := ((algebra.Alg.Term algebra.F.id_locker2_map_promote_pending ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Var 18)::nil)), (algebra.Alg.Term algebra.F.id_nil nil))::R_xml_0_rule_as_list_321. Definition R_xml_0_rule_as_list_323 := ((algebra.Alg.Term algebra.F.id_locker2_map_promote_pending ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 19):: (algebra.Alg.Var 20)::nil))::(algebra.Alg.Var 18)::nil)), (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_locker2_promote_pending ((algebra.Alg.Var 19):: (algebra.Alg.Var 18)::nil))::(algebra.Alg.Term algebra.F.id_locker2_map_promote_pending ((algebra.Alg.Var 20):: (algebra.Alg.Var 18)::nil))::nil)))::R_xml_0_rule_as_list_322. Definition R_xml_0_rule_as_list_324 := ((algebra.Alg.Term algebra.F.id_locker2_map_claim_lock ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Var 21)::(algebra.Alg.Var 22)::nil)), (algebra.Alg.Term algebra.F.id_nil nil))::R_xml_0_rule_as_list_323. Definition R_xml_0_rule_as_list_325 := ((algebra.Alg.Term algebra.F.id_locker2_map_claim_lock ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 19)::(algebra.Alg.Var 20)::nil)):: (algebra.Alg.Var 21)::(algebra.Alg.Var 22)::nil)), (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_locker2_claim_lock ((algebra.Alg.Var 19):: (algebra.Alg.Var 21)::(algebra.Alg.Var 22)::nil))::(algebra.Alg.Term algebra.F.id_locker2_map_claim_lock ((algebra.Alg.Var 20):: (algebra.Alg.Var 21)::(algebra.Alg.Var 22)::nil))::nil))):: R_xml_0_rule_as_list_324. Definition R_xml_0_rule_as_list_326 := ((algebra.Alg.Term algebra.F.id_locker2_map_add_pending ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Var 21):: (algebra.Alg.Var 22)::nil)),(algebra.Alg.Term algebra.F.id_nil nil)):: R_xml_0_rule_as_list_325. Definition R_xml_0_rule_as_list_327 := ((algebra.Alg.Term algebra.F.id_locker2_promote_pending ((algebra.Alg.Var 19)::(algebra.Alg.Var 22)::nil)), (algebra.Alg.Term algebra.F.id_case0 ((algebra.Alg.Var 22):: (algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil))::nil))):: R_xml_0_rule_as_list_326. Definition R_xml_0_rule_as_list_328 := ((algebra.Alg.Term algebra.F.id_case0 ((algebra.Alg.Var 22):: (algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 22)::(algebra.Alg.Var 23)::nil))::nil)), (algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Var 19):: (algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 22)::nil))::nil)):: (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_pending nil):: (algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Var 23)::nil))::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil))::nil)))::R_xml_0_rule_as_list_327. Definition R_xml_0_rule_as_list_329 := ((algebra.Alg.Term algebra.F.id_case0 ((algebra.Alg.Var 22):: (algebra.Alg.Var 19)::(algebra.Alg.Var 24)::nil)),(algebra.Alg.Var 19)):: R_xml_0_rule_as_list_328. Definition R_xml_0_rule_as_list_330 := ((algebra.Alg.Term algebra.F.id_locker2_remove_pending ((algebra.Alg.Var 19)::(algebra.Alg.Var 22)::nil)), (algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Var 19):: (algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Term algebra.F.id_subtract ((algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19):: (algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil))::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 22)::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil))::nil))::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)))::R_xml_0_rule_as_list_329. Definition R_xml_0_rule_as_list_331 := ((algebra.Alg.Term algebra.F.id_locker2_add_pending ((algebra.Alg.Var 19)::(algebra.Alg.Var 21)::(algebra.Alg.Var 22)::nil)), (algebra.Alg.Term algebra.F.id_case1 ((algebra.Alg.Var 22):: (algebra.Alg.Var 21)::(algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_member ((algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_resource nil)::nil)):: (algebra.Alg.Var 21)::nil))::nil)))::R_xml_0_rule_as_list_330. Definition R_xml_0_rule_as_list_332 := ((algebra.Alg.Term algebra.F.id_case1 ((algebra.Alg.Var 22):: (algebra.Alg.Var 21)::(algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_true nil)::nil)), (algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Var 19):: (algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19):: (algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil))::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 22)::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil))::nil))::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)))::R_xml_0_rule_as_list_331. Definition R_xml_0_rule_as_list_333 := ((algebra.Alg.Term algebra.F.id_case1 ((algebra.Alg.Var 22):: (algebra.Alg.Var 21)::(algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_false nil)::nil)),(algebra.Alg.Var 19)):: R_xml_0_rule_as_list_332. Definition R_xml_0_rule_as_list_334 := ((algebra.Alg.Term algebra.F.id_locker2_release_lock ((algebra.Alg.Var 19)::(algebra.Alg.Var 22)::nil)), (algebra.Alg.Term algebra.F.id_case2 ((algebra.Alg.Var 22):: (algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_gen_modtageq ((algebra.Alg.Var 22)::(algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_excl nil)::nil))::nil))::nil))):: R_xml_0_rule_as_list_333. Definition R_xml_0_rule_as_list_335 := ((algebra.Alg.Term algebra.F.id_case2 ((algebra.Alg.Var 22):: (algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_true nil)::nil)), (algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Var 19):: (algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_excllock nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)))::R_xml_0_rule_as_list_334. Definition R_xml_0_rule_as_list_336 := ((algebra.Alg.Term algebra.F.id_case4 ((algebra.Alg.Var 22):: (algebra.Alg.Var 19)::(algebra.Alg.Var 25)::nil)), (algebra.Alg.Term algebra.F.id_false nil))::R_xml_0_rule_as_list_335. Definition R_xml_0_rule_as_list_337 := ((algebra.Alg.Term algebra.F.id_locker2_obtainables ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Var 22)::nil)), (algebra.Alg.Term algebra.F.id_true nil))::R_xml_0_rule_as_list_336. Definition R_xml_0_rule_as_list_338 := ((algebra.Alg.Term algebra.F.id_locker2_obtainables ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 19)::(algebra.Alg.Var 20)::nil)):: (algebra.Alg.Var 22)::nil)), (algebra.Alg.Term algebra.F.id_case5 ((algebra.Alg.Var 22):: (algebra.Alg.Var 20)::(algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_member ((algebra.Alg.Var 22)::(algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil))::nil))::nil)))::R_xml_0_rule_as_list_337. Definition R_xml_0_rule_as_list_339 := ((algebra.Alg.Term algebra.F.id_case5 ((algebra.Alg.Var 22):: (algebra.Alg.Var 20)::(algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_true nil)::nil)), (algebra.Alg.Term algebra.F.id_andt ((algebra.Alg.Term algebra.F.id_locker2_obtainable ((algebra.Alg.Var 19):: (algebra.Alg.Var 22)::nil))::(algebra.Alg.Term algebra.F.id_locker2_obtainables ((algebra.Alg.Var 20):: (algebra.Alg.Var 22)::nil))::nil)))::R_xml_0_rule_as_list_338. Definition R_xml_0_rule_as_list_340 := ((algebra.Alg.Term algebra.F.id_case5 ((algebra.Alg.Var 22):: (algebra.Alg.Var 20)::(algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_false nil)::nil)), (algebra.Alg.Term algebra.F.id_locker2_obtainables ((algebra.Alg.Var 20)::(algebra.Alg.Var 22)::nil))):: R_xml_0_rule_as_list_339. Definition R_xml_0_rule_as_list_341 := ((algebra.Alg.Term algebra.F.id_locker2_check_available ((algebra.Alg.Var 26)::(algebra.Alg.Term algebra.F.id_nil nil)::nil)), (algebra.Alg.Term algebra.F.id_false nil))::R_xml_0_rule_as_list_340. Definition R_xml_0_rule_as_list_342 := ((algebra.Alg.Term algebra.F.id_locker2_check_available ((algebra.Alg.Var 26)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 19)::(algebra.Alg.Var 20)::nil))::nil)), (algebra.Alg.Term algebra.F.id_case6 ((algebra.Alg.Var 20):: (algebra.Alg.Var 19)::(algebra.Alg.Var 26)::(algebra.Alg.Term algebra.F.id_equal ((algebra.Alg.Var 26)::(algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil))::nil))::nil)))::R_xml_0_rule_as_list_341. Definition R_xml_0_rule_as_list_343 := ((algebra.Alg.Term algebra.F.id_case6 ((algebra.Alg.Var 20):: (algebra.Alg.Var 19)::(algebra.Alg.Var 26)::(algebra.Alg.Term algebra.F.id_true nil)::nil)), (algebra.Alg.Term algebra.F.id_andt ((algebra.Alg.Term algebra.F.id_equal ((algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19)::(algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_excl nil)::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::(algebra.Alg.Term algebra.F.id_equal ((algebra.Alg.Term algebra.F.id_record_extract ((algebra.Alg.Var 19):: (algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)))::R_xml_0_rule_as_list_342. Definition R_xml_0_rule_as_list_344 := ((algebra.Alg.Term algebra.F.id_case6 ((algebra.Alg.Var 20):: (algebra.Alg.Var 19)::(algebra.Alg.Var 26)::(algebra.Alg.Term algebra.F.id_false nil)::nil)), (algebra.Alg.Term algebra.F.id_locker2_check_available ((algebra.Alg.Var 26)::(algebra.Alg.Var 20)::nil))):: R_xml_0_rule_as_list_343. Definition R_xml_0_rule_as_list_345 := ((algebra.Alg.Term algebra.F.id_locker2_check_availables ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Var 20)::nil)), (algebra.Alg.Term algebra.F.id_true nil))::R_xml_0_rule_as_list_344. Definition R_xml_0_rule_as_list_346 := ((algebra.Alg.Term algebra.F.id_locker2_check_availables ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 26):: (algebra.Alg.Var 21)::nil))::(algebra.Alg.Var 20)::nil)), (algebra.Alg.Term algebra.F.id_andt ((algebra.Alg.Term algebra.F.id_locker2_check_available ((algebra.Alg.Var 26):: (algebra.Alg.Var 20)::nil))::(algebra.Alg.Term algebra.F.id_locker2_check_availables ((algebra.Alg.Var 21):: (algebra.Alg.Var 20)::nil))::nil)))::R_xml_0_rule_as_list_345. Definition R_xml_0_rule_as_list_347 := ((algebra.Alg.Term algebra.F.id_locker2_adduniq ((algebra.Alg.Term algebra.F.id_nil nil)::(algebra.Alg.Var 27)::nil)),(algebra.Alg.Var 27)):: R_xml_0_rule_as_list_346. Definition R_xml_0_rule_as_list_348 := ((algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 28)::(algebra.Alg.Var 29)::nil)):: (algebra.Alg.Var 27)::nil)), (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 28):: (algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Var 29):: (algebra.Alg.Var 27)::nil))::nil)))::R_xml_0_rule_as_list_347. Definition R_xml_0_rule_as_list_349 := ((algebra.Alg.Term algebra.F.id_subtract ((algebra.Alg.Var 27):: (algebra.Alg.Term algebra.F.id_nil nil)::nil)),(algebra.Alg.Var 27)):: R_xml_0_rule_as_list_348. Definition R_xml_0_rule_as_list_350 := ((algebra.Alg.Term algebra.F.id_subtract ((algebra.Alg.Var 27):: (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 28):: (algebra.Alg.Var 29)::nil))::nil)), (algebra.Alg.Term algebra.F.id_subtract ((algebra.Alg.Term algebra.F.id_delete ((algebra.Alg.Var 28)::(algebra.Alg.Var 27)::nil)):: (algebra.Alg.Var 29)::nil)))::R_xml_0_rule_as_list_349. Definition R_xml_0_rule_as_list_351 := ((algebra.Alg.Term algebra.F.id_delete ((algebra.Alg.Var 30):: (algebra.Alg.Term algebra.F.id_nil nil)::nil)), (algebra.Alg.Term algebra.F.id_nil nil))::R_xml_0_rule_as_list_350. Definition R_xml_0_rule_as_list_352 := ((algebra.Alg.Term algebra.F.id_delete ((algebra.Alg.Var 30):: (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 28):: (algebra.Alg.Var 29)::nil))::nil)), (algebra.Alg.Term algebra.F.id_case8 ((algebra.Alg.Var 29):: (algebra.Alg.Var 28)::(algebra.Alg.Var 30)::(algebra.Alg.Term algebra.F.id_equal ((algebra.Alg.Var 30):: (algebra.Alg.Var 28)::nil))::nil)))::R_xml_0_rule_as_list_351. Definition R_xml_0_rule_as_list_353 := ((algebra.Alg.Term algebra.F.id_case8 ((algebra.Alg.Var 29):: (algebra.Alg.Var 28)::(algebra.Alg.Var 30)::(algebra.Alg.Term algebra.F.id_true nil)::nil)),(algebra.Alg.Var 29)):: R_xml_0_rule_as_list_352. Definition R_xml_0_rule_as_list_354 := ((algebra.Alg.Term algebra.F.id_case8 ((algebra.Alg.Var 29):: (algebra.Alg.Var 28)::(algebra.Alg.Var 30)::(algebra.Alg.Term algebra.F.id_false nil)::nil)), (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 28):: (algebra.Alg.Term algebra.F.id_delete ((algebra.Alg.Var 30):: (algebra.Alg.Var 29)::nil))::nil)))::R_xml_0_rule_as_list_353. Definition R_xml_0_rule_as_list_355 := ((algebra.Alg.Term algebra.F.id_gen_tag ((algebra.Alg.Var 31)::nil)), (algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Var 31):: (algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Term algebra.F.id_tag nil)::nil))::nil)))::R_xml_0_rule_as_list_354. Definition R_xml_0_rule_as_list_356 := ((algebra.Alg.Term algebra.F.id_gen_modtageq ((algebra.Alg.Var 32):: (algebra.Alg.Var 33)::nil)), (algebra.Alg.Term algebra.F.id_equal ((algebra.Alg.Var 32):: (algebra.Alg.Var 33)::nil)))::R_xml_0_rule_as_list_355. Definition R_xml_0_rule_as_list_357 := ((algebra.Alg.Term algebra.F.id_member ((algebra.Alg.Var 30):: (algebra.Alg.Term algebra.F.id_nil nil)::nil)), (algebra.Alg.Term algebra.F.id_false nil))::R_xml_0_rule_as_list_356. Definition R_xml_0_rule_as_list_358 := ((algebra.Alg.Term algebra.F.id_member ((algebra.Alg.Var 30):: (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Var 28):: (algebra.Alg.Var 29)::nil))::nil)), (algebra.Alg.Term algebra.F.id_case9 ((algebra.Alg.Var 29):: (algebra.Alg.Var 28)::(algebra.Alg.Var 30)::(algebra.Alg.Term algebra.F.id_equal ((algebra.Alg.Var 30):: (algebra.Alg.Var 28)::nil))::nil)))::R_xml_0_rule_as_list_357. Definition R_xml_0_rule_as_list_359 := ((algebra.Alg.Term algebra.F.id_case9 ((algebra.Alg.Var 29):: (algebra.Alg.Var 28)::(algebra.Alg.Var 30)::(algebra.Alg.Term algebra.F.id_true nil)::nil)),(algebra.Alg.Term algebra.F.id_true nil)):: R_xml_0_rule_as_list_358. Definition R_xml_0_rule_as_list_360 := ((algebra.Alg.Term algebra.F.id_case9 ((algebra.Alg.Var 29):: (algebra.Alg.Var 28)::(algebra.Alg.Var 30)::(algebra.Alg.Term algebra.F.id_false nil)::nil)), (algebra.Alg.Term algebra.F.id_member ((algebra.Alg.Var 30):: (algebra.Alg.Var 29)::nil)))::R_xml_0_rule_as_list_359. Definition R_xml_0_rule_as_list_361 := ((algebra.Alg.Term algebra.F.id_eqs ((algebra.Alg.Term algebra.F.id_empty nil)::(algebra.Alg.Term algebra.F.id_empty nil)::nil)), (algebra.Alg.Term algebra.F.id_T nil))::R_xml_0_rule_as_list_360. Definition R_xml_0_rule_as_list_362 := ((algebra.Alg.Term algebra.F.id_eqs ((algebra.Alg.Term algebra.F.id_empty nil)::(algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 34):: (algebra.Alg.Var 35)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_361. Definition R_xml_0_rule_as_list_363 := ((algebra.Alg.Term algebra.F.id_eqs ((algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 36)::(algebra.Alg.Var 37)::nil))::(algebra.Alg.Term algebra.F.id_empty nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_362. Definition R_xml_0_rule_as_list_364 := ((algebra.Alg.Term algebra.F.id_eqs ((algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 36)::(algebra.Alg.Var 37)::nil))::(algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 34):: (algebra.Alg.Var 35)::nil))::nil)), (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 36)::(algebra.Alg.Var 34)::nil))::(algebra.Alg.Term algebra.F.id_eqs ((algebra.Alg.Var 37):: (algebra.Alg.Var 35)::nil))::nil)))::R_xml_0_rule_as_list_363. Definition R_xml_0_rule_as_list_365 := ((algebra.Alg.Term algebra.F.id_pushs ((algebra.Alg.Var 36):: (algebra.Alg.Var 37)::nil)), (algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 36):: (algebra.Alg.Var 37)::nil)))::R_xml_0_rule_as_list_364. Definition R_xml_0_rule_as_list_366 := ((algebra.Alg.Term algebra.F.id_pops ((algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 36):: (algebra.Alg.Var 37)::nil))::nil)),(algebra.Alg.Var 37)):: R_xml_0_rule_as_list_365. Definition R_xml_0_rule_as_list_367 := ((algebra.Alg.Term algebra.F.id_tops ((algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 36):: (algebra.Alg.Var 37)::nil))::nil)),(algebra.Alg.Var 36)):: R_xml_0_rule_as_list_366. Definition R_xml_0_rule_as_list_368 := ((algebra.Alg.Term algebra.F.id_istops ((algebra.Alg.Var 36):: (algebra.Alg.Term algebra.F.id_empty nil)::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_367. Definition R_xml_0_rule_as_list_369 := ((algebra.Alg.Term algebra.F.id_istops ((algebra.Alg.Var 36):: (algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 34):: (algebra.Alg.Var 37)::nil))::nil)), (algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 36):: (algebra.Alg.Var 34)::nil)))::R_xml_0_rule_as_list_368. Definition R_xml_0_rule_as_list_370 := ((algebra.Alg.Term algebra.F.id_eqc ((algebra.Alg.Term algebra.F.id_nocalls nil)::(algebra.Alg.Term algebra.F.id_nocalls nil)::nil)),(algebra.Alg.Term algebra.F.id_T nil)):: R_xml_0_rule_as_list_369. Definition R_xml_0_rule_as_list_371 := ((algebra.Alg.Term algebra.F.id_eqc ((algebra.Alg.Term algebra.F.id_nocalls nil)::(algebra.Alg.Term algebra.F.id_calls ((algebra.Alg.Var 34)::(algebra.Alg.Var 35):: (algebra.Alg.Var 38)::nil))::nil)), (algebra.Alg.Term algebra.F.id_F nil))::R_xml_0_rule_as_list_370. Definition R_xml_0_rule_as_list_372 := ((algebra.Alg.Term algebra.F.id_eqc ((algebra.Alg.Term algebra.F.id_calls ((algebra.Alg.Var 36)::(algebra.Alg.Var 37):: (algebra.Alg.Var 39)::nil))::(algebra.Alg.Term algebra.F.id_nocalls nil)::nil)),(algebra.Alg.Term algebra.F.id_F nil)):: R_xml_0_rule_as_list_371. Definition R_xml_0_rule_as_list_373 := ((algebra.Alg.Term algebra.F.id_eqc ((algebra.Alg.Term algebra.F.id_calls ((algebra.Alg.Var 36)::(algebra.Alg.Var 37):: (algebra.Alg.Var 39)::nil))::(algebra.Alg.Term algebra.F.id_calls ((algebra.Alg.Var 34)::(algebra.Alg.Var 35):: (algebra.Alg.Var 38)::nil))::nil)), (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 36)::(algebra.Alg.Var 34)::nil))::(algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqs ((algebra.Alg.Var 37)::(algebra.Alg.Var 35)::nil))::(algebra.Alg.Term algebra.F.id_eqc ((algebra.Alg.Var 39):: (algebra.Alg.Var 38)::nil))::nil))::nil)))::R_xml_0_rule_as_list_372. Definition R_xml_0_rule_as_list_374 := ((algebra.Alg.Term algebra.F.id_push ((algebra.Alg.Var 36):: (algebra.Alg.Var 34)::(algebra.Alg.Term algebra.F.id_nocalls nil)::nil)), (algebra.Alg.Term algebra.F.id_calls ((algebra.Alg.Var 36):: (algebra.Alg.Term algebra.F.id_stack ((algebra.Alg.Var 34):: (algebra.Alg.Term algebra.F.id_empty nil)::nil))::(algebra.Alg.Term algebra.F.id_nocalls nil)::nil)))::R_xml_0_rule_as_list_373. Definition R_xml_0_rule_as_list_375 := ((algebra.Alg.Term algebra.F.id_push ((algebra.Alg.Var 36):: (algebra.Alg.Var 34)::(algebra.Alg.Term algebra.F.id_calls ((algebra.Alg.Var 40)::(algebra.Alg.Var 37):: (algebra.Alg.Var 39)::nil))::nil)), (algebra.Alg.Term algebra.F.id_push1 ((algebra.Alg.Var 36):: (algebra.Alg.Var 34)::(algebra.Alg.Var 40)::(algebra.Alg.Var 37):: (algebra.Alg.Var 39)::(algebra.Alg.Term algebra.F.id_eqt ((algebra.Alg.Var 36)::(algebra.Alg.Var 40)::nil))::nil))):: R_xml_0_rule_as_list_374. Definition R_xml_0_rule_as_list_376 := ((algebra.Alg.Term algebra.F.id_push1 ((algebra.Alg.Var 36):: (algebra.Alg.Var 34)::(algebra.Alg.Var 40)::(algebra.Alg.Var 37):: (algebra.Alg.Var 39)::(algebra.Alg.Term algebra.F.id_T nil)::nil)), (algebra.Alg.Term algebra.F.id_calls ((algebra.Alg.Var 40):: (algebra.Alg.Term algebra.F.id_pushs ((algebra.Alg.Var 34):: (algebra.Alg.Var 37)::nil))::(algebra.Alg.Var 39)::nil))):: R_xml_0_rule_as_list_375. Definition R_xml_0_rule_as_list := R_xml_0_rule_as_list_376. 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 377. injection H;intros ;subst;constructor 376. injection H;intros ;subst;constructor 375. injection H;intros ;subst;constructor 374. injection H;intros ;subst;constructor 373. injection H;intros ;subst;constructor 372. injection H;intros ;subst;constructor 371. injection H;intros ;subst;constructor 370. injection H;intros ;subst;constructor 369. injection H;intros ;subst;constructor 368. injection H;intros ;subst;constructor 367. injection H;intros ;subst;constructor 366. injection H;intros ;subst;constructor 365. injection H;intros ;subst;constructor 364. injection H;intros ;subst;constructor 363. injection H;intros ;subst;constructor 362. injection H;intros ;subst;constructor 361. injection H;intros ;subst;constructor 360. injection H;intros ;subst;constructor 359. injection H;intros ;subst;constructor 358. injection H;intros ;subst;constructor 357. injection H;intros ;subst;constructor 356. injection H;intros ;subst;constructor 355. injection H;intros ;subst;constructor 354. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 353. injection H;intros ;subst;constructor 352. injection H;intros ;subst;constructor 351. injection H;intros ;subst;constructor 350. injection H;intros ;subst;constructor 349. injection H;intros ;subst;constructor 348. injection H;intros ;subst;constructor 347. injection H;intros ;subst;constructor 346. injection H;intros ;subst;constructor 345. injection H;intros ;subst;constructor 344. injection H;intros ;subst;constructor 343. injection H;intros ;subst;constructor 342. injection H;intros ;subst;constructor 341. injection H;intros ;subst;constructor 340. injection H;intros ;subst;constructor 339. injection H;intros ;subst;constructor 338. injection H;intros ;subst;constructor 337. injection H;intros ;subst;constructor 336. injection H;intros ;subst;constructor 335. injection H;intros ;subst;constructor 334. injection H;intros ;subst;constructor 333. injection H;intros ;subst;constructor 332. injection H;intros ;subst;constructor 331. injection H;intros ;subst;constructor 330. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 329. injection H;intros ;subst;constructor 328. injection H;intros ;subst;constructor 327. injection H;intros ;subst;constructor 326. injection H;intros ;subst;constructor 325. injection H;intros ;subst;constructor 324. injection H;intros ;subst;constructor 323. injection H;intros ;subst;constructor 322. injection H;intros ;subst;constructor 321. injection H;intros ;subst;constructor 320. injection H;intros ;subst;constructor 319. injection H;intros ;subst;constructor 318. injection H;intros ;subst;constructor 317. injection H;intros ;subst;constructor 316. injection H;intros ;subst;constructor 315. injection H;intros ;subst;constructor 314. injection H;intros ;subst;constructor 313. injection H;intros ;subst;constructor 312. injection H;intros ;subst;constructor 311. injection H;intros ;subst;constructor 310. injection H;intros ;subst;constructor 309. injection H;intros ;subst;constructor 308. injection H;intros ;subst;constructor 307. injection H;intros ;subst;constructor 306. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 305. injection H;intros ;subst;constructor 304. injection H;intros ;subst;constructor 303. injection H;intros ;subst;constructor 302. injection H;intros ;subst;constructor 301. injection H;intros ;subst;constructor 300. injection H;intros ;subst;constructor 299. injection H;intros ;subst;constructor 298. injection H;intros ;subst;constructor 297. injection H;intros ;subst;constructor 296. injection H;intros ;subst;constructor 295. injection H;intros ;subst;constructor 294. injection H;intros ;subst;constructor 293. injection H;intros ;subst;constructor 292. injection H;intros ;subst;constructor 291. injection H;intros ;subst;constructor 290. injection H;intros ;subst;constructor 289. injection H;intros ;subst;constructor 288. injection H;intros ;subst;constructor 287. injection H;intros ;subst;constructor 286. injection H;intros ;subst;constructor 285. injection H;intros ;subst;constructor 284. injection H;intros ;subst;constructor 283. injection H;intros ;subst;constructor 282. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 281. injection H;intros ;subst;constructor 280. injection H;intros ;subst;constructor 279. injection H;intros ;subst;constructor 278. injection H;intros ;subst;constructor 277. injection H;intros ;subst;constructor 276. injection H;intros ;subst;constructor 275. injection H;intros ;subst;constructor 274. injection H;intros ;subst;constructor 273. injection H;intros ;subst;constructor 272. injection H;intros ;subst;constructor 271. injection H;intros ;subst;constructor 270. injection H;intros ;subst;constructor 269. injection H;intros ;subst;constructor 268. injection H;intros ;subst;constructor 267. injection H;intros ;subst;constructor 266. injection H;intros ;subst;constructor 265. injection H;intros ;subst;constructor 264. injection H;intros ;subst;constructor 263. injection H;intros ;subst;constructor 262. injection H;intros ;subst;constructor 261. injection H;intros ;subst;constructor 260. injection H;intros ;subst;constructor 259. injection H;intros ;subst;constructor 258. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 257. injection H;intros ;subst;constructor 256. injection H;intros ;subst;constructor 255. injection H;intros ;subst;constructor 254. injection H;intros ;subst;constructor 253. injection H;intros ;subst;constructor 252. injection H;intros ;subst;constructor 251. injection H;intros ;subst;constructor 250. injection H;intros ;subst;constructor 249. injection H;intros ;subst;constructor 248. injection H;intros ;subst;constructor 247. injection H;intros ;subst;constructor 246. injection H;intros ;subst;constructor 245. injection H;intros ;subst;constructor 244. injection H;intros ;subst;constructor 243. injection H;intros ;subst;constructor 242. injection H;intros ;subst;constructor 241. injection H;intros ;subst;constructor 240. injection H;intros ;subst;constructor 239. injection H;intros ;subst;constructor 238. injection H;intros ;subst;constructor 237. injection H;intros ;subst;constructor 236. injection H;intros ;subst;constructor 235. injection H;intros ;subst;constructor 234. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 233. injection H;intros ;subst;constructor 232. injection H;intros ;subst;constructor 231. injection H;intros ;subst;constructor 230. injection H;intros ;subst;constructor 229. injection H;intros ;subst;constructor 228. injection H;intros ;subst;constructor 227. injection H;intros ;subst;constructor 226. injection H;intros ;subst;constructor 225. injection H;intros ;subst;constructor 224. injection H;intros ;subst;constructor 223. injection H;intros ;subst;constructor 222. injection H;intros ;subst;constructor 221. injection H;intros ;subst;constructor 220. injection H;intros ;subst;constructor 219. injection H;intros ;subst;constructor 218. injection H;intros ;subst;constructor 217. injection H;intros ;subst;constructor 216. injection H;intros ;subst;constructor 215. injection H;intros ;subst;constructor 214. injection H;intros ;subst;constructor 213. injection H;intros ;subst;constructor 212. injection H;intros ;subst;constructor 211. injection H;intros ;subst;constructor 210. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 209. injection H;intros ;subst;constructor 208. injection H;intros ;subst;constructor 207. injection H;intros ;subst;constructor 206. injection H;intros ;subst;constructor 205. injection H;intros ;subst;constructor 204. injection H;intros ;subst;constructor 203. injection H;intros ;subst;constructor 202. injection H;intros ;subst;constructor 201. injection H;intros ;subst;constructor 200. injection H;intros ;subst;constructor 199. injection H;intros ;subst;constructor 198. injection H;intros ;subst;constructor 197. injection H;intros ;subst;constructor 196. injection H;intros ;subst;constructor 195. injection H;intros ;subst;constructor 194. injection H;intros ;subst;constructor 193. injection H;intros ;subst;constructor 192. injection H;intros ;subst;constructor 191. injection H;intros ;subst;constructor 190. injection H;intros ;subst;constructor 189. injection H;intros ;subst;constructor 188. injection H;intros ;subst;constructor 187. injection H;intros ;subst;constructor 186. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 185. injection H;intros ;subst;constructor 184. injection H;intros ;subst;constructor 183. injection H;intros ;subst;constructor 182. injection H;intros ;subst;constructor 181. injection H;intros ;subst;constructor 180. injection H;intros ;subst;constructor 179. injection H;intros ;subst;constructor 178. injection H;intros ;subst;constructor 177. injection H;intros ;subst;constructor 176. injection H;intros ;subst;constructor 175. injection H;intros ;subst;constructor 174. injection H;intros ;subst;constructor 173. injection H;intros ;subst;constructor 172. injection H;intros ;subst;constructor 171. injection H;intros ;subst;constructor 170. injection H;intros ;subst;constructor 169. injection H;intros ;subst;constructor 168. injection H;intros ;subst;constructor 167. injection H;intros ;subst;constructor 166. injection H;intros ;subst;constructor 165. injection H;intros ;subst;constructor 164. injection H;intros ;subst;constructor 163. injection H;intros ;subst;constructor 162. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 161. injection H;intros ;subst;constructor 160. injection H;intros ;subst;constructor 159. injection H;intros ;subst;constructor 158. injection H;intros ;subst;constructor 157. injection H;intros ;subst;constructor 156. injection H;intros ;subst;constructor 155. injection H;intros ;subst;constructor 154. injection H;intros ;subst;constructor 153. injection H;intros ;subst;constructor 152. injection H;intros ;subst;constructor 151. injection H;intros ;subst;constructor 150. injection H;intros ;subst;constructor 149. injection H;intros ;subst;constructor 148. injection H;intros ;subst;constructor 147. injection H;intros ;subst;constructor 146. injection H;intros ;subst;constructor 145. injection H;intros ;subst;constructor 144. injection H;intros ;subst;constructor 143. injection H;intros ;subst;constructor 142. injection H;intros ;subst;constructor 141. injection H;intros ;subst;constructor 140. injection H;intros ;subst;constructor 139. injection H;intros ;subst;constructor 138. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 137. injection H;intros ;subst;constructor 136. injection H;intros ;subst;constructor 135. injection H;intros ;subst;constructor 134. injection H;intros ;subst;constructor 133. injection H;intros ;subst;constructor 132. injection H;intros ;subst;constructor 131. injection H;intros ;subst;constructor 130. injection H;intros ;subst;constructor 129. injection H;intros ;subst;constructor 128. injection H;intros ;subst;constructor 127. injection H;intros ;subst;constructor 126. injection H;intros ;subst;constructor 125. injection H;intros ;subst;constructor 124. injection H;intros ;subst;constructor 123. injection H;intros ;subst;constructor 122. injection H;intros ;subst;constructor 121. injection H;intros ;subst;constructor 120. injection H;intros ;subst;constructor 119. injection H;intros ;subst;constructor 118. injection H;intros ;subst;constructor 117. injection H;intros ;subst;constructor 116. injection H;intros ;subst;constructor 115. injection H;intros ;subst;constructor 114. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 113. injection H;intros ;subst;constructor 112. injection H;intros ;subst;constructor 111. injection H;intros ;subst;constructor 110. injection H;intros ;subst;constructor 109. injection H;intros ;subst;constructor 108. injection H;intros ;subst;constructor 107. injection H;intros ;subst;constructor 106. injection H;intros ;subst;constructor 105. injection H;intros ;subst;constructor 104. injection H;intros ;subst;constructor 103. injection H;intros ;subst;constructor 102. injection H;intros ;subst;constructor 101. injection H;intros ;subst;constructor 100. injection H;intros ;subst;constructor 99. injection H;intros ;subst;constructor 98. injection H;intros ;subst;constructor 97. injection H;intros ;subst;constructor 96. injection H;intros ;subst;constructor 95. injection H;intros ;subst;constructor 94. injection H;intros ;subst;constructor 93. injection H;intros ;subst;constructor 92. injection H;intros ;subst;constructor 91. injection H;intros ;subst;constructor 90. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 89. injection H;intros ;subst;constructor 88. injection H;intros ;subst;constructor 87. injection H;intros ;subst;constructor 86. injection H;intros ;subst;constructor 85. injection H;intros ;subst;constructor 84. injection H;intros ;subst;constructor 83. injection H;intros ;subst;constructor 82. injection H;intros ;subst;constructor 81. injection H;intros ;subst;constructor 80. injection H;intros ;subst;constructor 79. injection H;intros ;subst;constructor 78. injection H;intros ;subst;constructor 77. injection H;intros ;subst;constructor 76. injection H;intros ;subst;constructor 75. injection H;intros ;subst;constructor 74. injection H;intros ;subst;constructor 73. injection H;intros ;subst;constructor 72. injection H;intros ;subst;constructor 71. injection H;intros ;subst;constructor 70. injection H;intros ;subst;constructor 69. injection H;intros ;subst;constructor 68. injection H;intros ;subst;constructor 67. injection H;intros ;subst;constructor 66. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 65. injection H;intros ;subst;constructor 64. injection H;intros ;subst;constructor 63. injection H;intros ;subst;constructor 62. injection H;intros ;subst;constructor 61. injection H;intros ;subst;constructor 60. injection H;intros ;subst;constructor 59. injection H;intros ;subst;constructor 58. injection H;intros ;subst;constructor 57. injection H;intros ;subst;constructor 56. injection H;intros ;subst;constructor 55. injection H;intros ;subst;constructor 54. injection H;intros ;subst;constructor 53. injection H;intros ;subst;constructor 52. injection H;intros ;subst;constructor 51. injection H;intros ;subst;constructor 50. injection H;intros ;subst;constructor 49. injection H;intros ;subst;constructor 48. injection H;intros ;subst;constructor 47. injection H;intros ;subst;constructor 46. injection H;intros ;subst;constructor 45. injection H;intros ;subst;constructor 44. injection H;intros ;subst;constructor 43. injection H;intros ;subst;constructor 42. rewrite <- or_ext_generated.or25_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 41. injection H;intros ;subst;constructor 40. injection H;intros ;subst;constructor 39. injection H;intros ;subst;constructor 38. injection H;intros ;subst;constructor 37. injection H;intros ;subst;constructor 36. injection H;intros ;subst;constructor 35. injection H;intros ;subst;constructor 34. injection H;intros ;subst;constructor 33. injection H;intros ;subst;constructor 32. injection H;intros ;subst;constructor 31. injection H;intros ;subst;constructor 30. injection H;intros ;subst;constructor 29. 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. rewrite <- or_ext_generated.or18_equiv in H|-. case H;clear H;intros H. injection H;intros ;subst;constructor 17. injection H;intros ;subst;constructor 16. injection H;intros ;subst;constructor 15. injection H;intros ;subst;constructor 14. injection H;intros ;subst;constructor 13. injection H;intros ;subst;constructor 12. injection H;intros ;subst;constructor 11. injection H;intros ;subst;constructor 10. injection H;intros ;subst;constructor 9. injection H;intros ;subst;constructor 8. injection H;intros ;subst;constructor 7. injection H;intros ;subst;constructor 6. injection H;intros ;subst;constructor 5. 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_33 (x42 x43 x44 x45 x46 x47 x48 x49 x50 x51 x52 x53 x54 x55 x56 x57 x58 x59 x60 x61 x62 x63 x64 x65 x66 x67 x68 x69 x70 x71 x72 x73 x74:Prop) : Prop := | conj_33 : x42->x43->x44->x45->x46->x47->x48->x49->x50->x51->x52->x53->x54-> x55->x56->x57->x58->x59->x60->x61->x62->x63->x64->x65->x66->x67-> x68->x69->x70->x71->x72->x73->x74-> and_33 x42 x43 x44 x45 x46 x47 x48 x49 x50 x51 x52 x53 x54 x55 x56 x57 x58 x59 x60 x61 x62 x63 x64 x65 x66 x67 x68 x69 x70 x71 x72 x73 x74 . 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_33 (t = (algebra.Alg.Term algebra.F.id_a nil) -> t' = (algebra.Alg.Term algebra.F.id_a nil)) (t = (algebra.Alg.Term algebra.F.id_release nil) -> t' = (algebra.Alg.Term algebra.F.id_release nil)) (forall x43 x45, t = (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43::x45::nil)) -> exists x42, exists x44, t' = (algebra.Alg.Term algebra.F.id_locker2_obtainable (x42:: x44::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45)) (forall x43 x45, t = (algebra.Alg.Term algebra.F.id_stack (x43::x45::nil)) -> exists x42, exists x44, t' = (algebra.Alg.Term algebra.F.id_stack (x42::x44::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45)) (t = (algebra.Alg.Term algebra.F.id_locker nil) -> t' = (algebra.Alg.Term algebra.F.id_locker nil)) (forall x43, t = (algebra.Alg.Term algebra.F.id_int (x43::nil)) -> exists x42, t' = (algebra.Alg.Term algebra.F.id_int (x42::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)) (t = (algebra.Alg.Term algebra.F.id_true nil) -> t' = (algebra.Alg.Term algebra.F.id_true nil)) (t = (algebra.Alg.Term algebra.F.id_F nil) -> t' = (algebra.Alg.Term algebra.F.id_F nil)) (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_undefined nil) -> t' = (algebra.Alg.Term algebra.F.id_undefined nil)) (t = (algebra.Alg.Term algebra.F.id_nocalls nil) -> t' = (algebra.Alg.Term algebra.F.id_nocalls nil)) (t = (algebra.Alg.Term algebra.F.id_resource nil) -> t' = (algebra.Alg.Term algebra.F.id_resource nil)) (t = (algebra.Alg.Term algebra.F.id_ok nil) -> t' = (algebra.Alg.Term algebra.F.id_ok nil)) (forall x43 x45, t = (algebra.Alg.Term algebra.F.id_tuple (x43::x45::nil)) -> exists x42, exists x44, t' = (algebra.Alg.Term algebra.F.id_tuple (x42::x44::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45)) (forall x43, t = (algebra.Alg.Term algebra.F.id_s (x43::nil)) -> exists x42, t' = (algebra.Alg.Term algebra.F.id_s (x42::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)) (t = (algebra.Alg.Term algebra.F.id_T nil) -> t' = (algebra.Alg.Term algebra.F.id_T nil)) (t = (algebra.Alg.Term algebra.F.id_excl nil) -> t' = (algebra.Alg.Term algebra.F.id_excl nil)) (t = (algebra.Alg.Term algebra.F.id_nil nil) -> t' = (algebra.Alg.Term algebra.F.id_nil nil)) (t = (algebra.Alg.Term algebra.F.id_request nil) -> t' = (algebra.Alg.Term algebra.F.id_request nil)) (t = (algebra.Alg.Term algebra.F.id_mcrlrecord nil) -> t' = (algebra.Alg.Term algebra.F.id_mcrlrecord nil)) (forall x43 x45, t = (algebra.Alg.Term algebra.F.id_cons (x43::x45::nil)) -> exists x42, exists x44, t' = (algebra.Alg.Term algebra.F.id_cons (x42::x44::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45)) (t = (algebra.Alg.Term algebra.F.id_empty nil) -> t' = (algebra.Alg.Term algebra.F.id_empty nil)) (t = (algebra.Alg.Term algebra.F.id_lock nil) -> t' = (algebra.Alg.Term algebra.F.id_lock nil)) (t = (algebra.Alg.Term algebra.F.id_excllock nil) -> t' = (algebra.Alg.Term algebra.F.id_excllock nil)) (forall x43, t = (algebra.Alg.Term algebra.F.id_pid (x43::nil)) -> exists x42, t' = (algebra.Alg.Term algebra.F.id_pid (x42::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)) (forall x43 x45 x47, t = (algebra.Alg.Term algebra.F.id_calls (x43::x45::x47::nil)) -> exists x42, exists x44, exists x46, t' = (algebra.Alg.Term algebra.F.id_calls (x42::x44::x46::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x46 x47)) (t = (algebra.Alg.Term algebra.F.id_tag nil) -> t' = (algebra.Alg.Term algebra.F.id_tag nil)) (forall x43 x45, t = (algebra.Alg.Term algebra.F.id_equal (x43::x45::nil)) -> exists x42, exists x44, t' = (algebra.Alg.Term algebra.F.id_equal (x42::x44::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45)) (forall x43 x45 x47, t = (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x43::x45:: x47::nil)) -> exists x42, exists x44, exists x46, t' = (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x42:: x44::x46::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x46 x47)) (t = (algebra.Alg.Term algebra.F.id_pending nil) -> t' = (algebra.Alg.Term algebra.F.id_pending nil)) (forall x43 x45, t = (algebra.Alg.Term algebra.F.id_andt (x43::x45::nil)) -> exists x42, exists x44, t' = (algebra.Alg.Term algebra.F.id_andt (x42::x44::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45)) (forall x43, t = (algebra.Alg.Term algebra.F.id_tuplenil (x43::nil)) -> exists x42, t' = (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)) (t = (algebra.Alg.Term algebra.F.id_0 nil) -> t' = (algebra.Alg.Term algebra.F.id_0 nil)). Proof. intros t t' H. induction H as [|y IH z z_to_y] using closure_extension.refl_trans_clos_ind2. constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros x43 x45 H;exists x43;exists x45;intuition;constructor 1. intros x43 x45 H;exists x43;exists x45;intuition;constructor 1. intros H;intuition;constructor 1. intros x43 H;exists x43;intuition;constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros x43 x45 H;exists x43;exists x45;intuition;constructor 1. intros x43 H;exists x43;intuition;constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros x43 x45 H;exists x43;exists x45;intuition;constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros H;intuition;constructor 1. intros x43 H;exists x43;intuition;constructor 1. intros x43 x45 x47 H;exists x43;exists x45;exists x47;intuition; constructor 1. intros H;intuition;constructor 1. intros x43 x45 H;exists x43;exists x45;intuition;constructor 1. intros x43 x45 x47 H;exists x43;exists x45;exists x47;intuition; constructor 1. intros H;intuition;constructor 1. intros x43 x45 H;exists x43;exists x45;intuition;constructor 1. intros x43 H;exists x43;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_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0]. constructor. clear H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H; intros ;subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros x43 x45 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 x43 |- _ => destruct (H_id_locker2_obtainable y x45 (refl_equal _)) as [x42 [x44]]; intros ;intuition;exists x42;exists x44;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x45 |- _ => destruct (H_id_locker2_obtainable x43 y (refl_equal _)) as [x42 [x44]]; intros ;intuition;exists x42;exists x44;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_release H_id_locker2_obtainable H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros x43 x45 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 x43 |- _ => destruct (H_id_stack y x45 (refl_equal _)) as [x42 [x44]];intros ; intuition;exists x42;exists x44;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x45 |- _ => destruct (H_id_stack x43 y (refl_equal _)) as [x42 [x44]];intros ; intuition;exists x42;exists x44;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros x43 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 x43 |- _ => destruct (H_id_int y (refl_equal _)) as [x42];intros ;intuition; exists x42;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ;subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ;subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros x43 x45 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 x43 |- _ => destruct (H_id_tuple y x45 (refl_equal _)) as [x42 [x44]];intros ; intuition;exists x42;exists x44;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x45 |- _ => destruct (H_id_tuple x43 y (refl_equal _)) as [x42 [x44]];intros ; intuition;exists x42;exists x44;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros x43 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 x43 |- _ => destruct (H_id_s y (refl_equal _)) as [x42];intros ;intuition; exists x42;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros x43 x45 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 x43 |- _ => destruct (H_id_cons y x45 (refl_equal _)) as [x42 [x44]];intros ; intuition;exists x42;exists x44;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x45 |- _ => destruct (H_id_cons x43 y (refl_equal _)) as [x42 [x44]];intros ; intuition;exists x42;exists x44;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros x43 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 x43 |- _ => destruct (H_id_pid y (refl_equal _)) as [x42];intros ;intuition; exists x42;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros x43 x45 x47 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 x43 |- _ => destruct (H_id_calls y x45 x47 (refl_equal _)) as [x42 [x44 [x46]]]; intros ;intuition;exists x42;exists x44;exists x46;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x45 |- _ => destruct (H_id_calls x43 y x47 (refl_equal _)) as [x42 [x44 [x46]]]; intros ;intuition;exists x42;exists x44;exists x46;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x47 |- _ => destruct (H_id_calls x43 x45 y (refl_equal _)) as [x42 [x44 [x46]]]; intros ;intuition;exists x42;exists x44;exists x46;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ; subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil H_id_0;intros x43 x45 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 x43 |- _ => destruct (H_id_equal y x45 (refl_equal _)) as [x42 [x44]];intros ; intuition;exists x42;exists x44;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x45 |- _ => destruct (H_id_equal x43 y (refl_equal _)) as [x42 [x44]];intros ; intuition;exists x42;exists x44;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_pending H_id_andt H_id_tuplenil H_id_0;intros x43 x45 x47 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 x43 |- _ => destruct (H_id_locker2_claim_lock y x45 x47 (refl_equal _)) as [x42 [x44 [x46]]];intros ;intuition;exists x42;exists x44;exists x46; intuition;eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x45 |- _ => destruct (H_id_locker2_claim_lock x43 y x47 (refl_equal _)) as [x42 [x44 [x46]]];intros ;intuition;exists x42;exists x44;exists x46; intuition;eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x47 |- _ => destruct (H_id_locker2_claim_lock x43 x45 y (refl_equal _)) as [x42 [x44 [x46]]];intros ;intuition;exists x42;exists x44;exists x46; intuition;eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_andt H_id_tuplenil H_id_0;intros H;injection H;clear H;intros ;subst; repeat ( match goal with | H:closure.one_step_list (algebra.EQT.one_step _) _ _ |- _ => inversion H;clear H;subst end ). clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_tuplenil H_id_0;intros x43 x45 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 x43 |- _ => destruct (H_id_andt y x45 (refl_equal _)) as [x42 [x44]];intros ; intuition;exists x42;exists x44;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . match goal with | H:algebra.EQT.one_step _ ?y x45 |- _ => destruct (H_id_andt x43 y (refl_equal _)) as [x42 [x44]];intros ; intuition;exists x42;exists x44;intuition; eapply closure_extension.refl_trans_clos_R;eassumption end . clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_0;intros x43 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 x43 |- _ => destruct (H_id_tuplenil y (refl_equal _)) as [x42];intros ;intuition; exists x42;intuition;eapply closure_extension.refl_trans_clos_R; eassumption end . clear H_id_a H_id_release H_id_locker2_obtainable H_id_stack H_id_locker H_id_int H_id_true H_id_F H_id_false H_id_undefined H_id_nocalls H_id_resource H_id_ok H_id_tuple H_id_s H_id_T H_id_excl H_id_nil H_id_request H_id_mcrlrecord H_id_cons H_id_empty H_id_lock H_id_excllock H_id_pid H_id_calls H_id_tag H_id_equal H_id_locker2_claim_lock H_id_pending H_id_andt H_id_tuplenil;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_a_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> t = (algebra.Alg.Term algebra.F.id_a nil) -> t' = (algebra.Alg.Term algebra.F.id_a nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_release_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_release nil) -> t' = (algebra.Alg.Term algebra.F.id_release nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_locker2_obtainable_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x43 x45, t = (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43::x45::nil)) -> exists x42, exists x44, t' = (algebra.Alg.Term algebra.F.id_locker2_obtainable (x42:: x44::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_stack_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x43 x45, t = (algebra.Alg.Term algebra.F.id_stack (x43::x45::nil)) -> exists x42, exists x44, t' = (algebra.Alg.Term algebra.F.id_stack (x42::x44::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_locker_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_locker nil) -> t' = (algebra.Alg.Term algebra.F.id_locker nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_int_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x43, t = (algebra.Alg.Term algebra.F.id_int (x43::nil)) -> exists x42, t' = (algebra.Alg.Term algebra.F.id_int (x42::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_true_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> t = (algebra.Alg.Term algebra.F.id_true nil) -> t' = (algebra.Alg.Term algebra.F.id_true nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_F_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_F nil) -> t' = (algebra.Alg.Term algebra.F.id_F nil). 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_undefined_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_undefined nil) -> t' = (algebra.Alg.Term algebra.F.id_undefined nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_nocalls_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_nocalls nil) -> t' = (algebra.Alg.Term algebra.F.id_nocalls nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_resource_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_resource nil) -> t' = (algebra.Alg.Term algebra.F.id_resource nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_ok_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> t = (algebra.Alg.Term algebra.F.id_ok nil) -> t' = (algebra.Alg.Term algebra.F.id_ok nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_tuple_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x43 x45, t = (algebra.Alg.Term algebra.F.id_tuple (x43::x45::nil)) -> exists x42, exists x44, t' = (algebra.Alg.Term algebra.F.id_tuple (x42::x44::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_s_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x43, t = (algebra.Alg.Term algebra.F.id_s (x43::nil)) -> exists x42, t' = (algebra.Alg.Term algebra.F.id_s (x42::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_T_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_T nil) -> t' = (algebra.Alg.Term algebra.F.id_T nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_excl_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_excl nil) -> t' = (algebra.Alg.Term algebra.F.id_excl nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_nil_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> t = (algebra.Alg.Term algebra.F.id_nil nil) -> t' = (algebra.Alg.Term algebra.F.id_nil nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_request_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_request nil) -> t' = (algebra.Alg.Term algebra.F.id_request nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_mcrlrecord_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_mcrlrecord nil) -> t' = (algebra.Alg.Term algebra.F.id_mcrlrecord nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_cons_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x43 x45, t = (algebra.Alg.Term algebra.F.id_cons (x43::x45::nil)) -> exists x42, exists x44, t' = (algebra.Alg.Term algebra.F.id_cons (x42::x44::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_empty_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_empty nil) -> t' = (algebra.Alg.Term algebra.F.id_empty nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_lock_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_lock nil) -> t' = (algebra.Alg.Term algebra.F.id_lock nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_excllock_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_excllock nil) -> t' = (algebra.Alg.Term algebra.F.id_excllock nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_pid_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x43, t = (algebra.Alg.Term algebra.F.id_pid (x43::nil)) -> exists x42, t' = (algebra.Alg.Term algebra.F.id_pid (x42::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_calls_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x43 x45 x47, t = (algebra.Alg.Term algebra.F.id_calls (x43::x45::x47::nil)) -> exists x42, exists x44, exists x46, t' = (algebra.Alg.Term algebra.F.id_calls (x42::x44::x46::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x46 x47). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_tag_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_tag nil) -> t' = (algebra.Alg.Term algebra.F.id_tag nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_equal_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x43 x45, t = (algebra.Alg.Term algebra.F.id_equal (x43::x45::nil)) -> exists x42, exists x44, t' = (algebra.Alg.Term algebra.F.id_equal (x42::x44::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_locker2_claim_lock_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x43 x45 x47, t = (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x43::x45:: x47::nil)) -> exists x42, exists x44, exists x46, t' = (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x42:: x44::x46::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x46 x47). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_pending_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_pending nil) -> t' = (algebra.Alg.Term algebra.F.id_pending nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_andt_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x43 x45, t = (algebra.Alg.Term algebra.F.id_andt (x43::x45::nil)) -> exists x42, exists x44, t' = (algebra.Alg.Term algebra.F.id_andt (x42::x44::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43)/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x44 x45). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_tuplenil_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> forall x43, t = (algebra.Alg.Term algebra.F.id_tuplenil (x43::nil)) -> exists x42, t' = (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil))/\ (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) x42 x43). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Lemma id_0_is_R_xml_0_constructor : forall t t', (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) t' t) -> t = (algebra.Alg.Term algebra.F.id_0 nil) -> t' = (algebra.Alg.Term algebra.F.id_0 nil). Proof. intros t t' H. destruct (are_constuctors_of_R_xml_0 H). assumption. Qed. Ltac impossible_star_reduction_R_xml_0 := match goal with | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_a nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_a_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; impossible_star_reduction_R_xml_0 )) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_release nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_release_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_locker2_obtainable (?x43::?x42::nil)) |- _ => let x43 := fresh "x" in (let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_locker2_obtainable_is_R_xml_0_constructor H (refl_equal _)) as [x43 [x42 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_stack (?x43::?x42::nil)) |- _ => let x43 := fresh "x" in (let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_stack_is_R_xml_0_constructor H (refl_equal _)) as [x43 [x42 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_locker nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_locker_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_int (?x42::nil)) |- _ => let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_int_is_R_xml_0_constructor H (refl_equal _)) as [x42 [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 )) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_F nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_F_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_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_undefined nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_undefined_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_nocalls nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_nocalls_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_resource nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_resource_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; impossible_star_reduction_R_xml_0 )) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_ok nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_ok_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_tuple (?x43::?x42::nil)) |- _ => let x43 := fresh "x" in (let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_tuple_is_R_xml_0_constructor H (refl_equal _)) as [x43 [x42 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_s (?x42::nil)) |- _ => let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_s_is_R_xml_0_constructor H (refl_equal _)) as [x42 [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_T nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_T_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_excl nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_excl_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; impossible_star_reduction_R_xml_0 )) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_nil nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_nil_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; impossible_star_reduction_R_xml_0 )) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_request nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_request_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_mcrlrecord nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_mcrlrecord_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; impossible_star_reduction_R_xml_0 )) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_cons (?x43::?x42::nil)) |- _ => let x43 := fresh "x" in (let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_cons_is_R_xml_0_constructor H (refl_equal _)) as [x43 [x42 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_empty nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_empty_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_lock nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_lock_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_excllock nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_excllock_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_pid (?x42::nil)) |- _ => let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_pid_is_R_xml_0_constructor H (refl_equal _)) as [x42 [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_calls (?x44::?x43::?x42::nil)) |- _ => let x44 := fresh "x" in (let x43 := fresh "x" in (let x42 := 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_calls_is_R_xml_0_constructor H (refl_equal _)) as [x44 [x43 [x42 [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_tag nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_tag_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_equal (?x43::?x42::nil)) |- _ => let x43 := fresh "x" in (let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_equal_is_R_xml_0_constructor H (refl_equal _)) as [x43 [x42 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_locker2_claim_lock (?x44::?x43:: ?x42::nil)) |- _ => let x44 := fresh "x" in (let x43 := fresh "x" in (let x42 := 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_locker2_claim_lock_is_R_xml_0_constructor H (refl_equal _)) as [x44 [x43 [x42 [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_pending nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_pending_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_andt (?x43::?x42::nil)) |- _ => let x43 := fresh "x" in (let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_andt_is_R_xml_0_constructor H (refl_equal _)) as [x43 [x42 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_tuplenil (?x42::nil)) |- _ => let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_tuplenil_is_R_xml_0_constructor H (refl_equal _)) as [x42 [Heq Hred1]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; impossible_star_reduction_R_xml_0 )))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_0 nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_0_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; impossible_star_reduction_R_xml_0 )) end . Ltac simplify_star_reduction_R_xml_0 := match goal with | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_a nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_a_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; try (simplify_star_reduction_R_xml_0 ))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_release nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_release_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_locker2_obtainable (?x43::?x42::nil)) |- _ => let x43 := fresh "x" in (let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_locker2_obtainable_is_R_xml_0_constructor H (refl_equal _)) as [x43 [x42 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_stack (?x43::?x42::nil)) |- _ => let x43 := fresh "x" in (let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_stack_is_R_xml_0_constructor H (refl_equal _)) as [x43 [x42 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_locker nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_locker_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_int (?x42::nil)) |- _ => let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_int_is_R_xml_0_constructor H (refl_equal _)) as [x42 [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 ))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_F nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_F_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_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_undefined nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_undefined_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_nocalls nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_nocalls_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_resource nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_resource_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; try (simplify_star_reduction_R_xml_0 ))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_ok nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_ok_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_tuple (?x43::?x42::nil)) |- _ => let x43 := fresh "x" in (let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_tuple_is_R_xml_0_constructor H (refl_equal _)) as [x43 [x42 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_s (?x42::nil)) |- _ => let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_s_is_R_xml_0_constructor H (refl_equal _)) as [x42 [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_T nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_T_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_excl nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_excl_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; try (simplify_star_reduction_R_xml_0 ))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_nil nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_nil_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; try (simplify_star_reduction_R_xml_0 ))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_request nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_request_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_mcrlrecord nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_mcrlrecord_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; try (simplify_star_reduction_R_xml_0 ))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_cons (?x43::?x42::nil)) |- _ => let x43 := fresh "x" in (let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_cons_is_R_xml_0_constructor H (refl_equal _)) as [x43 [x42 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_empty nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_empty_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_lock nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_lock_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_excllock nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_excllock_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_pid (?x42::nil)) |- _ => let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_pid_is_R_xml_0_constructor H (refl_equal _)) as [x42 [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_calls (?x44::?x43::?x42::nil)) |- _ => let x44 := fresh "x" in (let x43 := fresh "x" in (let x42 := 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_calls_is_R_xml_0_constructor H (refl_equal _)) as [x44 [x43 [x42 [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_tag nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_tag_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_equal (?x43::?x42::nil)) |- _ => let x43 := fresh "x" in (let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_equal_is_R_xml_0_constructor H (refl_equal _)) as [x43 [x42 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_locker2_claim_lock (?x44::?x43:: ?x42::nil)) |- _ => let x44 := fresh "x" in (let x43 := fresh "x" in (let x42 := 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_locker2_claim_lock_is_R_xml_0_constructor H (refl_equal _)) as [x44 [x43 [x42 [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_pending nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_pending_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_andt (?x43::?x42::nil)) |- _ => let x43 := fresh "x" in (let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (let Hred2 := fresh "Hred" in (destruct (id_andt_is_R_xml_0_constructor H (refl_equal _)) as [x43 [x42 [Heq [Hred2 Hred1]]]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_tuplenil (?x42::nil)) |- _ => let x42 := fresh "x" in (let Heq := fresh "Heq" in (let Hred1 := fresh "Hred" in (destruct (id_tuplenil_is_R_xml_0_constructor H (refl_equal _)) as [x42 [Heq Hred1]]; (discriminate Heq)|| (injection Heq;intros ;subst;clear Heq;clear H; try (simplify_star_reduction_R_xml_0 ))))) | H:closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_rules) _ (algebra.Alg.Term algebra.F.id_0 nil) |- _ => let Heq := fresh "Heq" in (set (Heq:=id_0_is_R_xml_0_constructor H (refl_equal _)) in *; (discriminate Heq)|| (clearbody Heq;try (subst);try (clear Heq);clear H; try (simplify_star_reduction_R_xml_0 ))) 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_or : A ->A ->A. Hypothesis P_id_gen_tag : A ->A. Hypothesis P_id_record_new : A ->A. Hypothesis P_id_a : A. Hypothesis P_id_locker2_release_lock : A ->A ->A. Hypothesis P_id_eqt : A ->A ->A. Hypothesis P_id_istops : A ->A ->A. Hypothesis P_id_locker2_map_add_pending : A ->A ->A ->A. Hypothesis P_id_release : A. Hypothesis P_id_locker2_obtainable : A ->A ->A. Hypothesis P_id_imp : A ->A ->A. Hypothesis P_id_stack : A ->A ->A. Hypothesis P_id_locker2_map_promote_pending : A ->A ->A. Hypothesis P_id_locker : A. Hypothesis P_id_case4 : A ->A ->A ->A. Hypothesis P_id_int : A ->A. Hypothesis P_id_push : A ->A ->A ->A. Hypothesis P_id_locker2_add_pending : A ->A ->A ->A. Hypothesis P_id_true : A. Hypothesis P_id_locker2_check_availables : A ->A ->A. Hypothesis P_id_F : A. Hypothesis P_id_eqs : A ->A ->A. Hypothesis P_id_record_update : A ->A ->A ->A ->A. Hypothesis P_id_false : A. Hypothesis P_id_gen_modtageq : A ->A ->A. Hypothesis P_id_undefined : A. Hypothesis P_id_nocalls : A. Hypothesis P_id_locker2_remove_pending : A ->A ->A. Hypothesis P_id_resource : A. Hypothesis P_id_case6 : A ->A ->A ->A ->A. Hypothesis P_id_if : A ->A ->A ->A. Hypothesis P_id_pops : A ->A. Hypothesis P_id_locker2_map_claim_lock : A ->A ->A ->A. Hypothesis P_id_ok : A. Hypothesis P_id_case5 : A ->A ->A ->A ->A. Hypothesis P_id_tuple : A ->A ->A. Hypothesis P_id_member : A ->A ->A. Hypothesis P_id_s : A ->A. Hypothesis P_id_delete : A ->A ->A. Hypothesis P_id_T : A. Hypothesis P_id_case9 : A ->A ->A ->A ->A. Hypothesis P_id_record_extract : A ->A ->A ->A. Hypothesis P_id_excl : A. Hypothesis P_id_case2 : A ->A ->A ->A. Hypothesis P_id_nil : A. Hypothesis P_id_eqc : A ->A ->A. Hypothesis P_id_case0 : A ->A ->A ->A. Hypothesis P_id_request : A. Hypothesis P_id_locker2_check_available : A ->A ->A. Hypothesis P_id_not : A ->A. Hypothesis P_id_pushs : A ->A ->A. Hypothesis P_id_locker2_promote_pending : A ->A ->A. Hypothesis P_id_mcrlrecord : A. Hypothesis P_id_locker2_obtainables : A ->A ->A. Hypothesis P_id_cons : A ->A ->A. Hypothesis P_id_push1 : A ->A ->A ->A ->A ->A ->A. Hypothesis P_id_case1 : A ->A ->A ->A ->A. Hypothesis P_id_element : A ->A ->A. Hypothesis P_id_locker2_adduniq : A ->A ->A. Hypothesis P_id_and : A ->A ->A. Hypothesis P_id_empty : A. Hypothesis P_id_record_updates : A ->A ->A ->A. Hypothesis P_id_lock : A. Hypothesis P_id_excllock : A. Hypothesis P_id_pid : A ->A. Hypothesis P_id_calls : A ->A ->A ->A. Hypothesis P_id_subtract : A ->A ->A. Hypothesis P_id_tag : A. Hypothesis P_id_equal : A ->A ->A. Hypothesis P_id_eq : A ->A ->A. Hypothesis P_id_tops : A ->A. Hypothesis P_id_locker2_claim_lock : A ->A ->A ->A. Hypothesis P_id_pending : A. Hypothesis P_id_andt : A ->A ->A. Hypothesis P_id_tuplenil : A ->A. Hypothesis P_id_append : A ->A ->A. Hypothesis P_id_0 : A. Hypothesis P_id_case8 : A ->A ->A ->A ->A. Hypothesis P_id_or_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Hypothesis P_id_gen_tag_monotonic : forall x42 x43, (A0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Hypothesis P_id_record_new_monotonic : forall x42 x43, (A0 <= x43)/\ (x43 <= x42) ->P_id_record_new x43 <= P_id_record_new x42. Hypothesis P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Hypothesis P_id_eqt_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Hypothesis P_id_istops_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_istops x43 x45 <= P_id_istops x42 x44. Hypothesis P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Hypothesis P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Hypothesis P_id_imp_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Hypothesis P_id_stack_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Hypothesis P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Hypothesis P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Hypothesis P_id_int_monotonic : forall x42 x43, (A0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Hypothesis P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Hypothesis P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Hypothesis P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Hypothesis P_id_eqs_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Hypothesis P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (A0 <= x49)/\ (x49 <= x48) -> (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Hypothesis P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Hypothesis P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Hypothesis P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (A0 <= x49)/\ (x49 <= x48) -> (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Hypothesis P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Hypothesis P_id_pops_monotonic : forall x42 x43, (A0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Hypothesis P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Hypothesis P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (A0 <= x49)/\ (x49 <= x48) -> (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Hypothesis P_id_tuple_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Hypothesis P_id_member_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_member x43 x45 <= P_id_member x42 x44. Hypothesis P_id_s_monotonic : forall x42 x43, (A0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Hypothesis P_id_delete_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_delete x43 x45 <= P_id_delete x42 x44. Hypothesis P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (A0 <= x49)/\ (x49 <= x48) -> (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Hypothesis P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Hypothesis P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Hypothesis P_id_eqc_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Hypothesis P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Hypothesis P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Hypothesis P_id_not_monotonic : forall x42 x43, (A0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Hypothesis P_id_pushs_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Hypothesis P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Hypothesis P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Hypothesis P_id_cons_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Hypothesis P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (A0 <= x53)/\ (x53 <= x52) -> (A0 <= x51)/\ (x51 <= x50) -> (A0 <= x49)/\ (x49 <= x48) -> (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Hypothesis P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (A0 <= x49)/\ (x49 <= x48) -> (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Hypothesis P_id_element_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Hypothesis P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Hypothesis P_id_and_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Hypothesis P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Hypothesis P_id_pid_monotonic : forall x42 x43, (A0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Hypothesis P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Hypothesis P_id_subtract_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Hypothesis P_id_equal_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Hypothesis P_id_eq_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Hypothesis P_id_tops_monotonic : forall x42 x43, (A0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Hypothesis P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Hypothesis P_id_andt_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Hypothesis P_id_tuplenil_monotonic : forall x42 x43, (A0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Hypothesis P_id_append_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_append x43 x45 <= P_id_append x42 x44. Hypothesis P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (A0 <= x49)/\ (x49 <= x48) -> (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Hypothesis P_id_or_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_or x43 x42. Hypothesis P_id_gen_tag_bounded : forall x42, (A0 <= x42) ->A0 <= P_id_gen_tag x42. Hypothesis P_id_record_new_bounded : forall x42, (A0 <= x42) ->A0 <= P_id_record_new x42. Hypothesis P_id_a_bounded : A0 <= P_id_a . Hypothesis P_id_locker2_release_lock_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_locker2_release_lock x43 x42. Hypothesis P_id_eqt_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_eqt x43 x42. Hypothesis P_id_istops_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_istops x43 x42. Hypothesis P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (A0 <= x42) -> (A0 <= x43) -> (A0 <= x44) ->A0 <= P_id_locker2_map_add_pending x44 x43 x42. Hypothesis P_id_release_bounded : A0 <= P_id_release . Hypothesis P_id_locker2_obtainable_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_locker2_obtainable x43 x42. Hypothesis P_id_imp_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_imp x43 x42. Hypothesis P_id_stack_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_stack x43 x42. Hypothesis P_id_locker2_map_promote_pending_bounded : forall x42 x43, (A0 <= x42) -> (A0 <= x43) ->A0 <= P_id_locker2_map_promote_pending x43 x42. Hypothesis P_id_locker_bounded : A0 <= P_id_locker . Hypothesis P_id_case4_bounded : forall x44 x42 x43, (A0 <= x42) ->(A0 <= x43) ->(A0 <= x44) ->A0 <= P_id_case4 x44 x43 x42. Hypothesis P_id_int_bounded : forall x42, (A0 <= x42) ->A0 <= P_id_int x42. Hypothesis P_id_push_bounded : forall x44 x42 x43, (A0 <= x42) ->(A0 <= x43) ->(A0 <= x44) ->A0 <= P_id_push x44 x43 x42. Hypothesis P_id_locker2_add_pending_bounded : forall x44 x42 x43, (A0 <= x42) -> (A0 <= x43) ->(A0 <= x44) ->A0 <= P_id_locker2_add_pending x44 x43 x42. Hypothesis P_id_true_bounded : A0 <= P_id_true . Hypothesis P_id_locker2_check_availables_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_locker2_check_availables x43 x42. Hypothesis P_id_F_bounded : A0 <= P_id_F . Hypothesis P_id_eqs_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_eqs x43 x42. Hypothesis P_id_record_update_bounded : forall x44 x42 x45 x43, (A0 <= x42) -> (A0 <= x43) -> (A0 <= x44) ->(A0 <= x45) ->A0 <= P_id_record_update x45 x44 x43 x42. Hypothesis P_id_false_bounded : A0 <= P_id_false . Hypothesis P_id_gen_modtageq_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_gen_modtageq x43 x42. Hypothesis P_id_undefined_bounded : A0 <= P_id_undefined . Hypothesis P_id_nocalls_bounded : A0 <= P_id_nocalls . Hypothesis P_id_locker2_remove_pending_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_locker2_remove_pending x43 x42. Hypothesis P_id_resource_bounded : A0 <= P_id_resource . Hypothesis P_id_case6_bounded : forall x44 x42 x45 x43, (A0 <= x42) -> (A0 <= x43) -> (A0 <= x44) ->(A0 <= x45) ->A0 <= P_id_case6 x45 x44 x43 x42. Hypothesis P_id_if_bounded : forall x44 x42 x43, (A0 <= x42) ->(A0 <= x43) ->(A0 <= x44) ->A0 <= P_id_if x44 x43 x42. Hypothesis P_id_pops_bounded : forall x42, (A0 <= x42) ->A0 <= P_id_pops x42. Hypothesis P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (A0 <= x42) -> (A0 <= x43) -> (A0 <= x44) ->A0 <= P_id_locker2_map_claim_lock x44 x43 x42. Hypothesis P_id_ok_bounded : A0 <= P_id_ok . Hypothesis P_id_case5_bounded : forall x44 x42 x45 x43, (A0 <= x42) -> (A0 <= x43) -> (A0 <= x44) ->(A0 <= x45) ->A0 <= P_id_case5 x45 x44 x43 x42. Hypothesis P_id_tuple_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_tuple x43 x42. Hypothesis P_id_member_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_member x43 x42. Hypothesis P_id_s_bounded : forall x42, (A0 <= x42) ->A0 <= P_id_s x42. Hypothesis P_id_delete_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_delete x43 x42. Hypothesis P_id_T_bounded : A0 <= P_id_T . Hypothesis P_id_case9_bounded : forall x44 x42 x45 x43, (A0 <= x42) -> (A0 <= x43) -> (A0 <= x44) ->(A0 <= x45) ->A0 <= P_id_case9 x45 x44 x43 x42. Hypothesis P_id_record_extract_bounded : forall x44 x42 x43, (A0 <= x42) -> (A0 <= x43) ->(A0 <= x44) ->A0 <= P_id_record_extract x44 x43 x42. Hypothesis P_id_excl_bounded : A0 <= P_id_excl . Hypothesis P_id_case2_bounded : forall x44 x42 x43, (A0 <= x42) ->(A0 <= x43) ->(A0 <= x44) ->A0 <= P_id_case2 x44 x43 x42. Hypothesis P_id_nil_bounded : A0 <= P_id_nil . Hypothesis P_id_eqc_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_eqc x43 x42. Hypothesis P_id_case0_bounded : forall x44 x42 x43, (A0 <= x42) ->(A0 <= x43) ->(A0 <= x44) ->A0 <= P_id_case0 x44 x43 x42. Hypothesis P_id_request_bounded : A0 <= P_id_request . Hypothesis P_id_locker2_check_available_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_locker2_check_available x43 x42. Hypothesis P_id_not_bounded : forall x42, (A0 <= x42) ->A0 <= P_id_not x42. Hypothesis P_id_pushs_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_pushs x43 x42. Hypothesis P_id_locker2_promote_pending_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_locker2_promote_pending x43 x42. Hypothesis P_id_mcrlrecord_bounded : A0 <= P_id_mcrlrecord . Hypothesis P_id_locker2_obtainables_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_locker2_obtainables x43 x42. Hypothesis P_id_cons_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_cons x43 x42. Hypothesis P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (A0 <= x42) -> (A0 <= x43) -> (A0 <= x44) -> (A0 <= x45) -> (A0 <= x46) ->(A0 <= x47) ->A0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Hypothesis P_id_case1_bounded : forall x44 x42 x45 x43, (A0 <= x42) -> (A0 <= x43) -> (A0 <= x44) ->(A0 <= x45) ->A0 <= P_id_case1 x45 x44 x43 x42. Hypothesis P_id_element_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_element x43 x42. Hypothesis P_id_locker2_adduniq_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_locker2_adduniq x43 x42. Hypothesis P_id_and_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_and x43 x42. Hypothesis P_id_empty_bounded : A0 <= P_id_empty . Hypothesis P_id_record_updates_bounded : forall x44 x42 x43, (A0 <= x42) -> (A0 <= x43) ->(A0 <= x44) ->A0 <= P_id_record_updates x44 x43 x42. Hypothesis P_id_lock_bounded : A0 <= P_id_lock . Hypothesis P_id_excllock_bounded : A0 <= P_id_excllock . Hypothesis P_id_pid_bounded : forall x42, (A0 <= x42) ->A0 <= P_id_pid x42. Hypothesis P_id_calls_bounded : forall x44 x42 x43, (A0 <= x42) ->(A0 <= x43) ->(A0 <= x44) ->A0 <= P_id_calls x44 x43 x42. Hypothesis P_id_subtract_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_subtract x43 x42. Hypothesis P_id_tag_bounded : A0 <= P_id_tag . Hypothesis P_id_equal_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_equal x43 x42. Hypothesis P_id_eq_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_eq x43 x42. Hypothesis P_id_tops_bounded : forall x42, (A0 <= x42) ->A0 <= P_id_tops x42. Hypothesis P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (A0 <= x42) -> (A0 <= x43) ->(A0 <= x44) ->A0 <= P_id_locker2_claim_lock x44 x43 x42. Hypothesis P_id_pending_bounded : A0 <= P_id_pending . Hypothesis P_id_andt_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_andt x43 x42. Hypothesis P_id_tuplenil_bounded : forall x42, (A0 <= x42) ->A0 <= P_id_tuplenil x42. Hypothesis P_id_append_bounded : forall x42 x43, (A0 <= x42) ->(A0 <= x43) ->A0 <= P_id_append x43 x42. Hypothesis P_id_0_bounded : A0 <= P_id_0 . Hypothesis P_id_case8_bounded : forall x44 x42 x45 x43, (A0 <= x42) -> (A0 <= x43) -> (A0 <= x44) ->(A0 <= x45) ->A0 <= P_id_case8 x45 x44 x43 x42. Fixpoint measure t { struct t } := match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43::x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43::x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43::x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43:: x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43::x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43::x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43::x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43::x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43::x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45::x44::x43:: x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43::x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43::x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44::x43:: x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43::x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => A0 end. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43::x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43:: x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43::x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44:: x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44:: x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or => P_id_or | algebra.F.id_gen_tag => P_id_gen_tag | algebra.F.id_record_new => P_id_record_new | algebra.F.id_a => P_id_a | algebra.F.id_locker2_release_lock => P_id_locker2_release_lock | algebra.F.id_eqt => P_id_eqt | algebra.F.id_istops => P_id_istops | algebra.F.id_locker2_map_add_pending => P_id_locker2_map_add_pending | algebra.F.id_release => P_id_release | algebra.F.id_locker2_obtainable => P_id_locker2_obtainable | algebra.F.id_imp => P_id_imp | algebra.F.id_stack => P_id_stack | algebra.F.id_locker2_map_promote_pending => P_id_locker2_map_promote_pending | algebra.F.id_locker => P_id_locker | algebra.F.id_case4 => P_id_case4 | algebra.F.id_int => P_id_int | algebra.F.id_push => P_id_push | algebra.F.id_locker2_add_pending => P_id_locker2_add_pending | algebra.F.id_true => P_id_true | algebra.F.id_locker2_check_availables => P_id_locker2_check_availables | algebra.F.id_F => P_id_F | algebra.F.id_eqs => P_id_eqs | algebra.F.id_record_update => P_id_record_update | algebra.F.id_false => P_id_false | algebra.F.id_gen_modtageq => P_id_gen_modtageq | algebra.F.id_undefined => P_id_undefined | algebra.F.id_nocalls => P_id_nocalls | algebra.F.id_locker2_remove_pending => P_id_locker2_remove_pending | algebra.F.id_resource => P_id_resource | algebra.F.id_case6 => P_id_case6 | algebra.F.id_if => P_id_if | algebra.F.id_pops => P_id_pops | algebra.F.id_locker2_map_claim_lock => P_id_locker2_map_claim_lock | algebra.F.id_ok => P_id_ok | algebra.F.id_case5 => P_id_case5 | algebra.F.id_tuple => P_id_tuple | algebra.F.id_member => P_id_member | algebra.F.id_s => P_id_s | algebra.F.id_delete => P_id_delete | algebra.F.id_T => P_id_T | algebra.F.id_case9 => P_id_case9 | algebra.F.id_record_extract => P_id_record_extract | algebra.F.id_excl => P_id_excl | algebra.F.id_case2 => P_id_case2 | algebra.F.id_nil => P_id_nil | algebra.F.id_eqc => P_id_eqc | algebra.F.id_case0 => P_id_case0 | algebra.F.id_request => P_id_request | algebra.F.id_locker2_check_available => P_id_locker2_check_available | algebra.F.id_not => P_id_not | algebra.F.id_pushs => P_id_pushs | algebra.F.id_locker2_promote_pending => P_id_locker2_promote_pending | algebra.F.id_mcrlrecord => P_id_mcrlrecord | algebra.F.id_locker2_obtainables => P_id_locker2_obtainables | algebra.F.id_cons => P_id_cons | algebra.F.id_push1 => P_id_push1 | algebra.F.id_case1 => P_id_case1 | algebra.F.id_element => P_id_element | algebra.F.id_locker2_adduniq => P_id_locker2_adduniq | algebra.F.id_and => P_id_and | algebra.F.id_empty => P_id_empty | algebra.F.id_record_updates => P_id_record_updates | algebra.F.id_lock => P_id_lock | algebra.F.id_excllock => P_id_excllock | algebra.F.id_pid => P_id_pid | algebra.F.id_calls => P_id_calls | algebra.F.id_subtract => P_id_subtract | algebra.F.id_tag => P_id_tag | algebra.F.id_equal => P_id_equal | algebra.F.id_eq => P_id_eq | algebra.F.id_tops => P_id_tops | algebra.F.id_locker2_claim_lock => P_id_locker2_claim_lock | algebra.F.id_pending => P_id_pending | algebra.F.id_andt => P_id_andt | algebra.F.id_tuplenil => P_id_tuplenil | algebra.F.id_append => P_id_append | algebra.F.id_0 => P_id_0 | algebra.F.id_case8 => P_id_case8 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_or => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_gen_tag => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_record_new => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a => match l with | nil => _ | _::_ => _ end | algebra.F.id_locker2_release_lock => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_eqt => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_istops => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_locker2_map_add_pending => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_release => match l with | nil => _ | _::_ => _ end | algebra.F.id_locker2_obtainable => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_imp => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_stack => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_locker2_map_promote_pending => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_locker => match l with | nil => _ | _::_ => _ end | algebra.F.id_case4 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_int => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_push => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_locker2_add_pending => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_true => match l with | nil => _ | _::_ => _ end | algebra.F.id_locker2_check_availables => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_F => match l with | nil => _ | _::_ => _ end | algebra.F.id_eqs => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_record_update => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::_ => _ end | algebra.F.id_false => match l with | nil => _ | _::_ => _ end | algebra.F.id_gen_modtageq => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_undefined => match l with | nil => _ | _::_ => _ end | algebra.F.id_nocalls => match l with | nil => _ | _::_ => _ end | algebra.F.id_locker2_remove_pending => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_resource => match l with | nil => _ | _::_ => _ end | algebra.F.id_case6 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::_ => _ end | algebra.F.id_if => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_pops => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_locker2_map_claim_lock => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_ok => match l with | nil => _ | _::_ => _ end | algebra.F.id_case5 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::_ => _ end | algebra.F.id_tuple => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_member => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_s => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_delete => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_T => match l with | nil => _ | _::_ => _ end | algebra.F.id_case9 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::_ => _ end | algebra.F.id_record_extract => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_excl => match l with | nil => _ | _::_ => _ end | algebra.F.id_case2 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_nil => match l with | nil => _ | _::_ => _ end | algebra.F.id_eqc => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_case0 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_request => match l with | nil => _ | _::_ => _ end | algebra.F.id_locker2_check_available => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_not => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_pushs => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_locker2_promote_pending => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_mcrlrecord => match l with | nil => _ | _::_ => _ end | algebra.F.id_locker2_obtainables => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_cons => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_push1 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::nil => _ | _::_::_::_::_::_::nil => _ | _::_::_::_::_::_::_::_ => _ end | algebra.F.id_case1 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::_ => _ end | algebra.F.id_element => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_locker2_adduniq => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_and => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_empty => match l with | nil => _ | _::_ => _ end | algebra.F.id_record_updates => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_lock => match l with | nil => _ | _::_ => _ end | algebra.F.id_excllock => match l with | nil => _ | _::_ => _ end | algebra.F.id_pid => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_calls => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_subtract => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_tag => match l with | nil => _ | _::_ => _ end | algebra.F.id_equal => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_eq => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_tops => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_locker2_claim_lock => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_pending => match l with | nil => _ | _::_ => _ end | algebra.F.id_andt => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_tuplenil => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_append => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_0 => match l with | nil => _ | _::_ => _ end | algebra.F.id_case8 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::_ => _ end end;simpl in |-*;unfold eq_rect_r, eq_rect, sym_eq in |-*; try (reflexivity );f_equal ;auto. Qed. Lemma measure_bounded : forall t, A0 <= measure t. Proof. intros t. rewrite same_measure in |-*. apply (InterpGen.measure_bounded Aop). intros f. case f. vm_compute in |-*;intros ;apply P_id_or_bounded;assumption. vm_compute in |-*;intros ;apply P_id_gen_tag_bounded;assumption. vm_compute in |-*;intros ;apply P_id_record_new_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_release_lock_bounded; assumption. vm_compute in |-*;intros ;apply P_id_eqt_bounded;assumption. vm_compute in |-*;intros ;apply P_id_istops_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_map_add_pending_bounded; assumption. vm_compute in |-*;intros ;apply P_id_release_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_obtainable_bounded; assumption. vm_compute in |-*;intros ;apply P_id_imp_bounded;assumption. vm_compute in |-*;intros ;apply P_id_stack_bounded;assumption. vm_compute in |-*;intros ; apply P_id_locker2_map_promote_pending_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case4_bounded;assumption. vm_compute in |-*;intros ;apply P_id_int_bounded;assumption. vm_compute in |-*;intros ;apply P_id_push_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_add_pending_bounded; assumption. vm_compute in |-*;intros ;apply P_id_true_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_check_availables_bounded; assumption. vm_compute in |-*;intros ;apply P_id_F_bounded;assumption. vm_compute in |-*;intros ;apply P_id_eqs_bounded;assumption. vm_compute in |-*;intros ;apply P_id_record_update_bounded;assumption. vm_compute in |-*;intros ;apply P_id_false_bounded;assumption. vm_compute in |-*;intros ;apply P_id_gen_modtageq_bounded;assumption. vm_compute in |-*;intros ;apply P_id_undefined_bounded;assumption. vm_compute in |-*;intros ;apply P_id_nocalls_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_remove_pending_bounded; assumption. vm_compute in |-*;intros ;apply P_id_resource_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case6_bounded;assumption. vm_compute in |-*;intros ;apply P_id_if_bounded;assumption. vm_compute in |-*;intros ;apply P_id_pops_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_map_claim_lock_bounded; assumption. vm_compute in |-*;intros ;apply P_id_ok_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case5_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tuple_bounded;assumption. vm_compute in |-*;intros ;apply P_id_member_bounded;assumption. vm_compute in |-*;intros ;apply P_id_s_bounded;assumption. vm_compute in |-*;intros ;apply P_id_delete_bounded;assumption. vm_compute in |-*;intros ;apply P_id_T_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case9_bounded;assumption. vm_compute in |-*;intros ;apply P_id_record_extract_bounded;assumption. vm_compute in |-*;intros ;apply P_id_excl_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case2_bounded;assumption. vm_compute in |-*;intros ;apply P_id_nil_bounded;assumption. vm_compute in |-*;intros ;apply P_id_eqc_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case0_bounded;assumption. vm_compute in |-*;intros ;apply P_id_request_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_check_available_bounded; assumption. vm_compute in |-*;intros ;apply P_id_not_bounded;assumption. vm_compute in |-*;intros ;apply P_id_pushs_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_promote_pending_bounded; assumption. vm_compute in |-*;intros ;apply P_id_mcrlrecord_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_obtainables_bounded; assumption. vm_compute in |-*;intros ;apply P_id_cons_bounded;assumption. vm_compute in |-*;intros ;apply P_id_push1_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case1_bounded;assumption. vm_compute in |-*;intros ;apply P_id_element_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_adduniq_bounded;assumption. vm_compute in |-*;intros ;apply P_id_and_bounded;assumption. vm_compute in |-*;intros ;apply P_id_empty_bounded;assumption. vm_compute in |-*;intros ;apply P_id_record_updates_bounded;assumption. vm_compute in |-*;intros ;apply P_id_lock_bounded;assumption. vm_compute in |-*;intros ;apply P_id_excllock_bounded;assumption. vm_compute in |-*;intros ;apply P_id_pid_bounded;assumption. vm_compute in |-*;intros ;apply P_id_calls_bounded;assumption. vm_compute in |-*;intros ;apply P_id_subtract_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tag_bounded;assumption. vm_compute in |-*;intros ;apply P_id_equal_bounded;assumption. vm_compute in |-*;intros ;apply P_id_eq_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tops_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_claim_lock_bounded; assumption. vm_compute in |-*;intros ;apply P_id_pending_bounded;assumption. vm_compute in |-*;intros ;apply P_id_andt_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tuplenil_bounded;assumption. vm_compute in |-*;intros ;apply P_id_append_bounded;assumption. vm_compute in |-*;intros ;apply P_id_0_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case8_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_or_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_gen_tag_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_record_new_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_locker2_release_lock_monotonic; assumption. vm_compute in |-*;intros ;apply P_id_eqt_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_istops_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_locker2_map_add_pending_monotonic; assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_locker2_obtainable_monotonic; assumption. vm_compute in |-*;intros ;apply P_id_imp_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_stack_monotonic;assumption. vm_compute in |-*;intros ; apply P_id_locker2_map_promote_pending_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_case4_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_int_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_push_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_locker2_add_pending_monotonic; assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_locker2_check_availables_monotonic; assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_eqs_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_record_update_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_gen_modtageq_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_locker2_remove_pending_monotonic; assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_case6_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_if_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_pops_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_locker2_map_claim_lock_monotonic; assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_case5_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_tuple_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_member_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_s_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_delete_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_case9_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_record_extract_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_case2_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_eqc_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_case0_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_locker2_check_available_monotonic; assumption. vm_compute in |-*;intros ;apply P_id_not_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_pushs_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_locker2_promote_pending_monotonic; assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_locker2_obtainables_monotonic; assumption. vm_compute in |-*;intros ;apply P_id_cons_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_push1_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_case1_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_element_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_locker2_adduniq_monotonic; assumption. vm_compute in |-*;intros ;apply P_id_and_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_record_updates_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_pid_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_calls_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_subtract_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_equal_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_eq_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_tops_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_locker2_claim_lock_monotonic; assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_andt_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_tuplenil_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_append_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_case8_monotonic;assumption. intros f. case f. vm_compute in |-*;intros ;apply P_id_or_bounded;assumption. vm_compute in |-*;intros ;apply P_id_gen_tag_bounded;assumption. vm_compute in |-*;intros ;apply P_id_record_new_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_release_lock_bounded; assumption. vm_compute in |-*;intros ;apply P_id_eqt_bounded;assumption. vm_compute in |-*;intros ;apply P_id_istops_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_map_add_pending_bounded; assumption. vm_compute in |-*;intros ;apply P_id_release_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_obtainable_bounded; assumption. vm_compute in |-*;intros ;apply P_id_imp_bounded;assumption. vm_compute in |-*;intros ;apply P_id_stack_bounded;assumption. vm_compute in |-*;intros ; apply P_id_locker2_map_promote_pending_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case4_bounded;assumption. vm_compute in |-*;intros ;apply P_id_int_bounded;assumption. vm_compute in |-*;intros ;apply P_id_push_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_add_pending_bounded; assumption. vm_compute in |-*;intros ;apply P_id_true_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_check_availables_bounded; assumption. vm_compute in |-*;intros ;apply P_id_F_bounded;assumption. vm_compute in |-*;intros ;apply P_id_eqs_bounded;assumption. vm_compute in |-*;intros ;apply P_id_record_update_bounded;assumption. vm_compute in |-*;intros ;apply P_id_false_bounded;assumption. vm_compute in |-*;intros ;apply P_id_gen_modtageq_bounded;assumption. vm_compute in |-*;intros ;apply P_id_undefined_bounded;assumption. vm_compute in |-*;intros ;apply P_id_nocalls_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_remove_pending_bounded; assumption. vm_compute in |-*;intros ;apply P_id_resource_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case6_bounded;assumption. vm_compute in |-*;intros ;apply P_id_if_bounded;assumption. vm_compute in |-*;intros ;apply P_id_pops_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_map_claim_lock_bounded; assumption. vm_compute in |-*;intros ;apply P_id_ok_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case5_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tuple_bounded;assumption. vm_compute in |-*;intros ;apply P_id_member_bounded;assumption. vm_compute in |-*;intros ;apply P_id_s_bounded;assumption. vm_compute in |-*;intros ;apply P_id_delete_bounded;assumption. vm_compute in |-*;intros ;apply P_id_T_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case9_bounded;assumption. vm_compute in |-*;intros ;apply P_id_record_extract_bounded;assumption. vm_compute in |-*;intros ;apply P_id_excl_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case2_bounded;assumption. vm_compute in |-*;intros ;apply P_id_nil_bounded;assumption. vm_compute in |-*;intros ;apply P_id_eqc_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case0_bounded;assumption. vm_compute in |-*;intros ;apply P_id_request_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_check_available_bounded; assumption. vm_compute in |-*;intros ;apply P_id_not_bounded;assumption. vm_compute in |-*;intros ;apply P_id_pushs_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_promote_pending_bounded; assumption. vm_compute in |-*;intros ;apply P_id_mcrlrecord_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_obtainables_bounded; assumption. vm_compute in |-*;intros ;apply P_id_cons_bounded;assumption. vm_compute in |-*;intros ;apply P_id_push1_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case1_bounded;assumption. vm_compute in |-*;intros ;apply P_id_element_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_adduniq_bounded;assumption. vm_compute in |-*;intros ;apply P_id_and_bounded;assumption. vm_compute in |-*;intros ;apply P_id_empty_bounded;assumption. vm_compute in |-*;intros ;apply P_id_record_updates_bounded;assumption. vm_compute in |-*;intros ;apply P_id_lock_bounded;assumption. vm_compute in |-*;intros ;apply P_id_excllock_bounded;assumption. vm_compute in |-*;intros ;apply P_id_pid_bounded;assumption. vm_compute in |-*;intros ;apply P_id_calls_bounded;assumption. vm_compute in |-*;intros ;apply P_id_subtract_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tag_bounded;assumption. vm_compute in |-*;intros ;apply P_id_equal_bounded;assumption. vm_compute in |-*;intros ;apply P_id_eq_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tops_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_claim_lock_bounded; assumption. vm_compute in |-*;intros ;apply P_id_pending_bounded;assumption. vm_compute in |-*;intros ;apply P_id_andt_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tuplenil_bounded;assumption. vm_compute in |-*;intros ;apply P_id_append_bounded;assumption. vm_compute in |-*;intros ;apply P_id_0_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case8_bounded;assumption. intros . do 2 (rewrite <- same_measure in |-*). apply rules_monotonic;assumption. assumption. Qed. Hypothesis P_id_GEN_MODTAGEQ : A ->A ->A. Hypothesis P_id_ELEMENT : A ->A ->A. Hypothesis P_id_Marked_eq : A ->A ->A. Hypothesis P_id_CASE9 : A ->A ->A ->A ->A. Hypothesis P_id_LOCKER2_REMOVE_PENDING : A ->A ->A. Hypothesis P_id_Marked_record_new : A ->A. Hypothesis P_id_CASE6 : A ->A ->A ->A ->A. Hypothesis P_id_LOCKER2_PROMOTE_PENDING : A ->A ->A. Hypothesis P_id_Marked_if : A ->A ->A ->A. Hypothesis P_id_PUSH : A ->A ->A ->A. Hypothesis P_id_MEMBER : A ->A ->A. Hypothesis P_id_CASE5 : A ->A ->A ->A ->A. Hypothesis P_id_RECORD_UPDATE : A ->A ->A ->A ->A. Hypothesis P_id_Marked_not : A ->A. Hypothesis P_id_ISTOPS : A ->A ->A. Hypothesis P_id_LOCKER2_ADD_PENDING : A ->A ->A ->A. Hypothesis P_id_Marked_or : A ->A ->A. Hypothesis P_id_DELETE : A ->A ->A. Hypothesis P_id_CASE0 : A ->A ->A ->A. Hypothesis P_id_Marked_imp : A ->A ->A. Hypothesis P_id_EQT : A ->A ->A. Hypothesis P_id_PUSHS : A ->A ->A. Hypothesis P_id_LOCKER2_RELEASE_LOCK : A ->A ->A. Hypothesis P_id_LOCKER2_OBTAINABLES : A ->A ->A. Hypothesis P_id_RECORD_UPDATES : A ->A ->A ->A. Hypothesis P_id_Marked_locker2_adduniq : A ->A ->A. Hypothesis P_id_EQS : A ->A ->A. Hypothesis P_id_SUBTRACT : A ->A ->A. Hypothesis P_id_Marked_gen_tag : A ->A. Hypothesis P_id_LOCKER2_CHECK_AVAILABLES : A ->A ->A. Hypothesis P_id_LOCKER2_MAP_CLAIM_LOCK : A ->A ->A ->A. Hypothesis P_id_Marked_case4 : A ->A ->A ->A. Hypothesis P_id_PUSH1 : A ->A ->A ->A ->A ->A ->A. Hypothesis P_id_APPEND : A ->A ->A. Hypothesis P_id_LOCKER2_CHECK_AVAILABLE : A ->A ->A. Hypothesis P_id_LOCKER2_MAP_PROMOTE_PENDING : A ->A ->A. Hypothesis P_id_Marked_pops : A ->A. Hypothesis P_id_EQC : A ->A ->A. Hypothesis P_id_CASE1 : A ->A ->A ->A ->A. Hypothesis P_id_CASE8 : A ->A ->A ->A ->A. Hypothesis P_id_RECORD_EXTRACT : A ->A ->A ->A. Hypothesis P_id_Marked_locker2_map_add_pending : A ->A ->A ->A. Hypothesis P_id_AND : A ->A ->A. Hypothesis P_id_Marked_tops : A ->A. Hypothesis P_id_CASE2 : A ->A ->A ->A. Hypothesis P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Hypothesis P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Hypothesis P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Hypothesis P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (A0 <= x49)/\ (x49 <= x48) -> (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Hypothesis P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Hypothesis P_id_Marked_record_new_monotonic : forall x42 x43, (A0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Hypothesis P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (A0 <= x49)/\ (x49 <= x48) -> (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Hypothesis P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Hypothesis P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Hypothesis P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Hypothesis P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Hypothesis P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (A0 <= x49)/\ (x49 <= x48) -> (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Hypothesis P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (A0 <= x49)/\ (x49 <= x48) -> (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Hypothesis P_id_Marked_not_monotonic : forall x42 x43, (A0 <= x43)/\ (x43 <= x42) ->P_id_Marked_not x43 <= P_id_Marked_not x42. Hypothesis P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Hypothesis P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Hypothesis P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Hypothesis P_id_DELETE_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Hypothesis P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Hypothesis P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Hypothesis P_id_EQT_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Hypothesis P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Hypothesis P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Hypothesis P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Hypothesis P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Hypothesis P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Hypothesis P_id_EQS_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Hypothesis P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Hypothesis P_id_Marked_gen_tag_monotonic : forall x42 x43, (A0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Hypothesis P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Hypothesis P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Hypothesis P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Hypothesis P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (A0 <= x53)/\ (x53 <= x52) -> (A0 <= x51)/\ (x51 <= x50) -> (A0 <= x49)/\ (x49 <= x48) -> (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Hypothesis P_id_APPEND_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Hypothesis P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Hypothesis P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Hypothesis P_id_Marked_pops_monotonic : forall x42 x43, (A0 <= x43)/\ (x43 <= x42) -> P_id_Marked_pops x43 <= P_id_Marked_pops x42. Hypothesis P_id_EQC_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Hypothesis P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (A0 <= x49)/\ (x49 <= x48) -> (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Hypothesis P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (A0 <= x49)/\ (x49 <= x48) -> (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Hypothesis P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Hypothesis P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Hypothesis P_id_AND_monotonic : forall x44 x42 x45 x43, (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Hypothesis P_id_Marked_tops_monotonic : forall x42 x43, (A0 <= x43)/\ (x43 <= x42) -> P_id_Marked_tops x43 <= P_id_Marked_tops x42. Hypothesis P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (A0 <= x47)/\ (x47 <= x46) -> (A0 <= x45)/\ (x45 <= x44) -> (A0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Definition marked_measure t := match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43::x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43:: x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43::x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43::x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45::x44::x43:: x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43::x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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 x45 x44 x43 x42; apply (P_id_CASE8 x45 x44 x43 x42). 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 x43 x42;apply (P_id_APPEND x43 x42). 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 x42;apply (P_id_Marked_tops x42). 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 x43 x42;apply (P_id_Marked_eq x43 x42). 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 x43 x42;apply (P_id_SUBTRACT x43 x42). 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 x44 x43 x42; apply (P_id_RECORD_UPDATES x44 x43 x42). 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 x43 x42;apply (P_id_AND x43 x42). 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 x43 x42; apply (P_id_Marked_locker2_adduniq x43 x42). 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 x43 x42;apply (P_id_ELEMENT x43 x42). 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 x45 x44 x43 x42; apply (P_id_CASE1 x45 x44 x43 x42). 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 x47 x46 x45 x44 x43 x42; apply (P_id_PUSH1 x47 x46 x45 x44 x43 x42). 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 x43 x42; apply (P_id_LOCKER2_OBTAINABLES x43 x42). 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 x43 x42; apply (P_id_LOCKER2_PROMOTE_PENDING x43 x42). 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 x43 x42;apply (P_id_PUSHS x43 x42). 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 x42;apply (P_id_Marked_not x42). 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 x43 x42; apply (P_id_LOCKER2_CHECK_AVAILABLE x43 x42). 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 x44 x43 x42;apply (P_id_CASE0 x44 x43 x42). 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 x43 x42;apply (P_id_EQC x43 x42). 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 x44 x43 x42;apply (P_id_CASE2 x44 x43 x42). 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 x44 x43 x42; apply (P_id_RECORD_EXTRACT x44 x43 x42). 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 x45 x44 x43 x42; apply (P_id_CASE9 x45 x44 x43 x42). 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 x43 x42;apply (P_id_DELETE x43 x42). 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 x43 x42;apply (P_id_MEMBER x43 x42). 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 x45 x44 x43 x42; apply (P_id_CASE5 x45 x44 x43 x42). 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 x44 x43 x42; apply (P_id_LOCKER2_MAP_CLAIM_LOCK x44 x43 x42). 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 x42;apply (P_id_Marked_pops x42). 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 x44 x43 x42;apply (P_id_Marked_if x44 x43 x42). 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 x45 x44 x43 x42; apply (P_id_CASE6 x45 x44 x43 x42). 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 x43 x42; apply (P_id_LOCKER2_REMOVE_PENDING x43 x42). 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 x43 x42;apply (P_id_GEN_MODTAGEQ x43 x42). 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 x45 x44 x43 x42; apply (P_id_RECORD_UPDATE x45 x44 x43 x42). 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 x43 x42;apply (P_id_EQS x43 x42). 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 x43 x42; apply (P_id_LOCKER2_CHECK_AVAILABLES x43 x42). 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 x44 x43 x42; apply (P_id_LOCKER2_ADD_PENDING x44 x43 x42). 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 x44 x43 x42;apply (P_id_PUSH x44 x43 x42). 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 x44 x43 x42; apply (P_id_Marked_case4 x44 x43 x42). 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 x43 x42; apply (P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x42). 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 x43 x42;apply (P_id_Marked_imp x43 x42). 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 x44 x43 x42; apply (P_id_Marked_locker2_map_add_pending x44 x43 x42). 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 x43 x42;apply (P_id_ISTOPS x43 x42). 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 x43 x42;apply (P_id_EQT x43 x42). 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 x43 x42; apply (P_id_LOCKER2_RELEASE_LOCK x43 x42). 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 x42;apply (P_id_Marked_record_new x42). 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 x42;apply (P_id_Marked_gen_tag x42). 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 x43 x42;apply (P_id_Marked_or x43 x42). 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_or => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_gen_tag => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_record_new => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_a => match l with | nil => _ | _::_ => _ end | algebra.F.id_locker2_release_lock => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_eqt => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_istops => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_locker2_map_add_pending => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_release => match l with | nil => _ | _::_ => _ end | algebra.F.id_locker2_obtainable => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_imp => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_stack => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_locker2_map_promote_pending => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_locker => match l with | nil => _ | _::_ => _ end | algebra.F.id_case4 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_int => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_push => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_locker2_add_pending => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_true => match l with | nil => _ | _::_ => _ end | algebra.F.id_locker2_check_availables => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_F => match l with | nil => _ | _::_ => _ end | algebra.F.id_eqs => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_record_update => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::_ => _ end | algebra.F.id_false => match l with | nil => _ | _::_ => _ end | algebra.F.id_gen_modtageq => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_undefined => match l with | nil => _ | _::_ => _ end | algebra.F.id_nocalls => match l with | nil => _ | _::_ => _ end | algebra.F.id_locker2_remove_pending => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_resource => match l with | nil => _ | _::_ => _ end | algebra.F.id_case6 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::_ => _ end | algebra.F.id_if => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_pops => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_locker2_map_claim_lock => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_ok => match l with | nil => _ | _::_ => _ end | algebra.F.id_case5 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::_ => _ end | algebra.F.id_tuple => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_member => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_s => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_delete => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_T => match l with | nil => _ | _::_ => _ end | algebra.F.id_case9 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::_ => _ end | algebra.F.id_record_extract => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_excl => match l with | nil => _ | _::_ => _ end | algebra.F.id_case2 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_nil => match l with | nil => _ | _::_ => _ end | algebra.F.id_eqc => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_case0 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_request => match l with | nil => _ | _::_ => _ end | algebra.F.id_locker2_check_available => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_not => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_pushs => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_locker2_promote_pending => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_mcrlrecord => match l with | nil => _ | _::_ => _ end | algebra.F.id_locker2_obtainables => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_cons => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_push1 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::nil => _ | _::_::_::_::_::_::nil => _ | _::_::_::_::_::_::_::_ => _ end | algebra.F.id_case1 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::_ => _ end | algebra.F.id_element => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_locker2_adduniq => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_and => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_empty => match l with | nil => _ | _::_ => _ end | algebra.F.id_record_updates => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_lock => match l with | nil => _ | _::_ => _ end | algebra.F.id_excllock => match l with | nil => _ | _::_ => _ end | algebra.F.id_pid => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_calls => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_subtract => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_tag => match l with | nil => _ | _::_ => _ end | algebra.F.id_equal => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_eq => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_tops => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_locker2_claim_lock => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::_ => _ end | algebra.F.id_pending => match l with | nil => _ | _::_ => _ end | algebra.F.id_andt => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_tuplenil => match l with | nil => _ | _::nil => _ | _::_::_ => _ end | algebra.F.id_append => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::_ => _ end | algebra.F.id_0 => match l with | nil => _ | _::_ => _ end | algebra.F.id_case8 => match l with | nil => _ | _::nil => _ | _::_::nil => _ | _::_::_::nil => _ | _::_::_::_::nil => _ | _::_::_::_::_::_ => _ end end. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . 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 . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . 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 . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . 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 . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . vm_compute in |-*;reflexivity . lazy - [measure InterpGen.measure Pols] in |-* ;f_equal ; apply same_measure. vm_compute in |-*;reflexivity . Qed. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44:: x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44:: x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44:: x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46:: x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44:: x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44:: x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_gen_tag_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_record_new_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_locker2_release_lock_monotonic; assumption. vm_compute in |-*;intros ;apply P_id_eqt_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_istops_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_locker2_map_add_pending_monotonic; assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_locker2_obtainable_monotonic; assumption. vm_compute in |-*;intros ;apply P_id_imp_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_stack_monotonic;assumption. vm_compute in |-*;intros ; apply P_id_locker2_map_promote_pending_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_case4_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_int_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_push_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_locker2_add_pending_monotonic; assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_locker2_check_availables_monotonic; assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_eqs_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_record_update_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_gen_modtageq_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_locker2_remove_pending_monotonic; assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_case6_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_if_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_pops_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_locker2_map_claim_lock_monotonic; assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_case5_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_tuple_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_member_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_s_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_delete_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_case9_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_record_extract_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_case2_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_eqc_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_case0_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_locker2_check_available_monotonic; assumption. vm_compute in |-*;intros ;apply P_id_not_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_pushs_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_locker2_promote_pending_monotonic; assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_locker2_obtainables_monotonic; assumption. vm_compute in |-*;intros ;apply P_id_cons_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_push1_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_case1_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_element_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_locker2_adduniq_monotonic; assumption. vm_compute in |-*;intros ;apply P_id_and_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_record_updates_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_pid_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_calls_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_subtract_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_equal_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_eq_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_tops_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_locker2_claim_lock_monotonic; assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_andt_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_tuplenil_monotonic;assumption. vm_compute in |-*;intros ;apply P_id_append_monotonic;assumption. vm_compute in |-*;apply (Aop.(le_refl)). vm_compute in |-*;intros ;apply P_id_case8_monotonic;assumption. clear f. intros f. case f. vm_compute in |-*;intros ;apply P_id_or_bounded;assumption. vm_compute in |-*;intros ;apply P_id_gen_tag_bounded;assumption. vm_compute in |-*;intros ;apply P_id_record_new_bounded;assumption. vm_compute in |-*;intros ;apply P_id_a_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_release_lock_bounded; assumption. vm_compute in |-*;intros ;apply P_id_eqt_bounded;assumption. vm_compute in |-*;intros ;apply P_id_istops_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_map_add_pending_bounded; assumption. vm_compute in |-*;intros ;apply P_id_release_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_obtainable_bounded; assumption. vm_compute in |-*;intros ;apply P_id_imp_bounded;assumption. vm_compute in |-*;intros ;apply P_id_stack_bounded;assumption. vm_compute in |-*;intros ; apply P_id_locker2_map_promote_pending_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case4_bounded;assumption. vm_compute in |-*;intros ;apply P_id_int_bounded;assumption. vm_compute in |-*;intros ;apply P_id_push_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_add_pending_bounded; assumption. vm_compute in |-*;intros ;apply P_id_true_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_check_availables_bounded; assumption. vm_compute in |-*;intros ;apply P_id_F_bounded;assumption. vm_compute in |-*;intros ;apply P_id_eqs_bounded;assumption. vm_compute in |-*;intros ;apply P_id_record_update_bounded;assumption. vm_compute in |-*;intros ;apply P_id_false_bounded;assumption. vm_compute in |-*;intros ;apply P_id_gen_modtageq_bounded;assumption. vm_compute in |-*;intros ;apply P_id_undefined_bounded;assumption. vm_compute in |-*;intros ;apply P_id_nocalls_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_remove_pending_bounded; assumption. vm_compute in |-*;intros ;apply P_id_resource_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case6_bounded;assumption. vm_compute in |-*;intros ;apply P_id_if_bounded;assumption. vm_compute in |-*;intros ;apply P_id_pops_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_map_claim_lock_bounded; assumption. vm_compute in |-*;intros ;apply P_id_ok_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case5_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tuple_bounded;assumption. vm_compute in |-*;intros ;apply P_id_member_bounded;assumption. vm_compute in |-*;intros ;apply P_id_s_bounded;assumption. vm_compute in |-*;intros ;apply P_id_delete_bounded;assumption. vm_compute in |-*;intros ;apply P_id_T_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case9_bounded;assumption. vm_compute in |-*;intros ;apply P_id_record_extract_bounded;assumption. vm_compute in |-*;intros ;apply P_id_excl_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case2_bounded;assumption. vm_compute in |-*;intros ;apply P_id_nil_bounded;assumption. vm_compute in |-*;intros ;apply P_id_eqc_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case0_bounded;assumption. vm_compute in |-*;intros ;apply P_id_request_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_check_available_bounded; assumption. vm_compute in |-*;intros ;apply P_id_not_bounded;assumption. vm_compute in |-*;intros ;apply P_id_pushs_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_promote_pending_bounded; assumption. vm_compute in |-*;intros ;apply P_id_mcrlrecord_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_obtainables_bounded; assumption. vm_compute in |-*;intros ;apply P_id_cons_bounded;assumption. vm_compute in |-*;intros ;apply P_id_push1_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case1_bounded;assumption. vm_compute in |-*;intros ;apply P_id_element_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_adduniq_bounded;assumption. vm_compute in |-*;intros ;apply P_id_and_bounded;assumption. vm_compute in |-*;intros ;apply P_id_empty_bounded;assumption. vm_compute in |-*;intros ;apply P_id_record_updates_bounded;assumption. vm_compute in |-*;intros ;apply P_id_lock_bounded;assumption. vm_compute in |-*;intros ;apply P_id_excllock_bounded;assumption. vm_compute in |-*;intros ;apply P_id_pid_bounded;assumption. vm_compute in |-*;intros ;apply P_id_calls_bounded;assumption. vm_compute in |-*;intros ;apply P_id_subtract_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tag_bounded;assumption. vm_compute in |-*;intros ;apply P_id_equal_bounded;assumption. vm_compute in |-*;intros ;apply P_id_eq_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tops_bounded;assumption. vm_compute in |-*;intros ;apply P_id_locker2_claim_lock_bounded; assumption. vm_compute in |-*;intros ;apply P_id_pending_bounded;assumption. vm_compute in |-*;intros ;apply P_id_andt_bounded;assumption. vm_compute in |-*;intros ;apply P_id_tuplenil_bounded;assumption. vm_compute in |-*;intros ;apply P_id_append_bounded;assumption. vm_compute in |-*;intros ;apply P_id_0_bounded;assumption. vm_compute in |-*;intros ;apply P_id_case8_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_CASE8_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_APPEND_monotonic;assumption. match type of H with | orb ?a ?b = true => assert (H':{a = true}+{b = true});[ revert H;case a;[left;reflexivity|simpl;intros h;right;exact h]| clear H;destruct H' as [H|H]] end . match type of H with | _ _ ?t = true => generalize (algebra.F.symb_eq_bool_ok f t); unfold algebra.Alg.eq_symb_bool in H; rewrite H;clear H;intros ;subst end . vm_compute in |-*;intros ;apply P_id_Marked_tops_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_eq_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_SUBTRACT_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_RECORD_UPDATES_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_Marked_locker2_adduniq_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_ELEMENT_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_CASE1_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_PUSH1_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_LOCKER2_OBTAINABLES_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_LOCKER2_PROMOTE_PENDING_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_PUSHS_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_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_LOCKER2_CHECK_AVAILABLE_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_CASE0_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_EQC_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_CASE2_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_RECORD_EXTRACT_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_CASE9_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_DELETE_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_MEMBER_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_CASE5_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_LOCKER2_MAP_CLAIM_LOCK_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_pops_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_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_CASE6_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_LOCKER2_REMOVE_PENDING_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_GEN_MODTAGEQ_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_RECORD_UPDATE_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_EQS_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_LOCKER2_CHECK_AVAILABLES_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_LOCKER2_ADD_PENDING_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_PUSH_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_case4_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_LOCKER2_MAP_PROMOTE_PENDING_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_imp_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_locker2_map_add_pending_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_ISTOPS_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_EQT_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_LOCKER2_RELEASE_LOCK_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_record_new_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_gen_tag_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_or_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_or : Z ->Z ->Z. Hypothesis P_id_gen_tag : Z ->Z. Hypothesis P_id_record_new : Z ->Z. Hypothesis P_id_a : Z. Hypothesis P_id_locker2_release_lock : Z ->Z ->Z. Hypothesis P_id_eqt : Z ->Z ->Z. Hypothesis P_id_istops : Z ->Z ->Z. Hypothesis P_id_locker2_map_add_pending : Z ->Z ->Z ->Z. Hypothesis P_id_release : Z. Hypothesis P_id_locker2_obtainable : Z ->Z ->Z. Hypothesis P_id_imp : Z ->Z ->Z. Hypothesis P_id_stack : Z ->Z ->Z. Hypothesis P_id_locker2_map_promote_pending : Z ->Z ->Z. Hypothesis P_id_locker : Z. Hypothesis P_id_case4 : Z ->Z ->Z ->Z. Hypothesis P_id_int : Z ->Z. Hypothesis P_id_push : Z ->Z ->Z ->Z. Hypothesis P_id_locker2_add_pending : Z ->Z ->Z ->Z. Hypothesis P_id_true : Z. Hypothesis P_id_locker2_check_availables : Z ->Z ->Z. Hypothesis P_id_F : Z. Hypothesis P_id_eqs : Z ->Z ->Z. Hypothesis P_id_record_update : Z ->Z ->Z ->Z ->Z. Hypothesis P_id_false : Z. Hypothesis P_id_gen_modtageq : Z ->Z ->Z. Hypothesis P_id_undefined : Z. Hypothesis P_id_nocalls : Z. Hypothesis P_id_locker2_remove_pending : Z ->Z ->Z. Hypothesis P_id_resource : Z. Hypothesis P_id_case6 : Z ->Z ->Z ->Z ->Z. Hypothesis P_id_if : Z ->Z ->Z ->Z. Hypothesis P_id_pops : Z ->Z. Hypothesis P_id_locker2_map_claim_lock : Z ->Z ->Z ->Z. Hypothesis P_id_ok : Z. Hypothesis P_id_case5 : Z ->Z ->Z ->Z ->Z. Hypothesis P_id_tuple : Z ->Z ->Z. Hypothesis P_id_member : Z ->Z ->Z. Hypothesis P_id_s : Z ->Z. Hypothesis P_id_delete : Z ->Z ->Z. Hypothesis P_id_T : Z. Hypothesis P_id_case9 : Z ->Z ->Z ->Z ->Z. Hypothesis P_id_record_extract : Z ->Z ->Z ->Z. Hypothesis P_id_excl : Z. Hypothesis P_id_case2 : Z ->Z ->Z ->Z. Hypothesis P_id_nil : Z. Hypothesis P_id_eqc : Z ->Z ->Z. Hypothesis P_id_case0 : Z ->Z ->Z ->Z. Hypothesis P_id_request : Z. Hypothesis P_id_locker2_check_available : Z ->Z ->Z. Hypothesis P_id_not : Z ->Z. Hypothesis P_id_pushs : Z ->Z ->Z. Hypothesis P_id_locker2_promote_pending : Z ->Z ->Z. Hypothesis P_id_mcrlrecord : Z. Hypothesis P_id_locker2_obtainables : Z ->Z ->Z. Hypothesis P_id_cons : Z ->Z ->Z. Hypothesis P_id_push1 : Z ->Z ->Z ->Z ->Z ->Z ->Z. Hypothesis P_id_case1 : Z ->Z ->Z ->Z ->Z. Hypothesis P_id_element : Z ->Z ->Z. Hypothesis P_id_locker2_adduniq : Z ->Z ->Z. Hypothesis P_id_and : Z ->Z ->Z. Hypothesis P_id_empty : Z. Hypothesis P_id_record_updates : Z ->Z ->Z ->Z. Hypothesis P_id_lock : Z. Hypothesis P_id_excllock : Z. Hypothesis P_id_pid : Z ->Z. Hypothesis P_id_calls : Z ->Z ->Z ->Z. Hypothesis P_id_subtract : Z ->Z ->Z. Hypothesis P_id_tag : Z. Hypothesis P_id_equal : Z ->Z ->Z. Hypothesis P_id_eq : Z ->Z ->Z. Hypothesis P_id_tops : Z ->Z. Hypothesis P_id_locker2_claim_lock : Z ->Z ->Z ->Z. Hypothesis P_id_pending : Z. Hypothesis P_id_andt : Z ->Z ->Z. Hypothesis P_id_tuplenil : Z ->Z. Hypothesis P_id_append : Z ->Z ->Z. Hypothesis P_id_0 : Z. Hypothesis P_id_case8 : Z ->Z ->Z ->Z ->Z. Hypothesis P_id_or_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Hypothesis P_id_gen_tag_monotonic : forall x42 x43, (min_value <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Hypothesis P_id_record_new_monotonic : forall x42 x43, (min_value <= x43)/\ (x43 <= x42) -> P_id_record_new x43 <= P_id_record_new x42. Hypothesis P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Hypothesis P_id_eqt_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_eqt x43 x45 <= P_id_eqt x42 x44. Hypothesis P_id_istops_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_istops x43 x45 <= P_id_istops x42 x44. Hypothesis P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Hypothesis P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Hypothesis P_id_imp_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_imp x43 x45 <= P_id_imp x42 x44. Hypothesis P_id_stack_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_stack x43 x45 <= P_id_stack x42 x44. Hypothesis P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Hypothesis P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Hypothesis P_id_int_monotonic : forall x42 x43, (min_value <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Hypothesis P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Hypothesis P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Hypothesis P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Hypothesis P_id_eqs_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_eqs x43 x45 <= P_id_eqs x42 x44. Hypothesis P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (min_value <= x49)/\ (x49 <= x48) -> (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Hypothesis P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Hypothesis P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Hypothesis P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (min_value <= x49)/\ (x49 <= x48) -> (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Hypothesis P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Hypothesis P_id_pops_monotonic : forall x42 x43, (min_value <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Hypothesis P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Hypothesis P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (min_value <= x49)/\ (x49 <= x48) -> (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Hypothesis P_id_tuple_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_tuple x43 x45 <= P_id_tuple x42 x44. Hypothesis P_id_member_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_member x43 x45 <= P_id_member x42 x44. Hypothesis P_id_s_monotonic : forall x42 x43, (min_value <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Hypothesis P_id_delete_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_delete x43 x45 <= P_id_delete x42 x44. Hypothesis P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (min_value <= x49)/\ (x49 <= x48) -> (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Hypothesis P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Hypothesis P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Hypothesis P_id_eqc_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_eqc x43 x45 <= P_id_eqc x42 x44. Hypothesis P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Hypothesis P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Hypothesis P_id_not_monotonic : forall x42 x43, (min_value <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Hypothesis P_id_pushs_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_pushs x43 x45 <= P_id_pushs x42 x44. Hypothesis P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Hypothesis P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Hypothesis P_id_cons_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_cons x43 x45 <= P_id_cons x42 x44. Hypothesis P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (min_value <= x53)/\ (x53 <= x52) -> (min_value <= x51)/\ (x51 <= x50) -> (min_value <= x49)/\ (x49 <= x48) -> (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Hypothesis P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (min_value <= x49)/\ (x49 <= x48) -> (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Hypothesis P_id_element_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Hypothesis P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Hypothesis P_id_and_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_and x43 x45 <= P_id_and x42 x44. Hypothesis P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Hypothesis P_id_pid_monotonic : forall x42 x43, (min_value <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Hypothesis P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Hypothesis P_id_subtract_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Hypothesis P_id_equal_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_equal x43 x45 <= P_id_equal x42 x44. Hypothesis P_id_eq_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Hypothesis P_id_tops_monotonic : forall x42 x43, (min_value <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Hypothesis P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Hypothesis P_id_andt_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_andt x43 x45 <= P_id_andt x42 x44. Hypothesis P_id_tuplenil_monotonic : forall x42 x43, (min_value <= x43)/\ (x43 <= x42) -> P_id_tuplenil x43 <= P_id_tuplenil x42. Hypothesis P_id_append_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_append x43 x45 <= P_id_append x42 x44. Hypothesis P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (min_value <= x49)/\ (x49 <= x48) -> (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Hypothesis P_id_or_bounded : forall x42 x43, (min_value <= x42) ->(min_value <= x43) ->min_value <= P_id_or x43 x42. Hypothesis P_id_gen_tag_bounded : forall x42, (min_value <= x42) ->min_value <= P_id_gen_tag x42. Hypothesis P_id_record_new_bounded : forall x42, (min_value <= x42) ->min_value <= P_id_record_new x42. Hypothesis P_id_a_bounded : min_value <= P_id_a . Hypothesis P_id_locker2_release_lock_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_locker2_release_lock x43 x42. Hypothesis P_id_eqt_bounded : forall x42 x43, (min_value <= x42) ->(min_value <= x43) ->min_value <= P_id_eqt x43 x42. Hypothesis P_id_istops_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_istops x43 x42. Hypothesis P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) -> min_value <= P_id_locker2_map_add_pending x44 x43 x42. Hypothesis P_id_release_bounded : min_value <= P_id_release . Hypothesis P_id_locker2_obtainable_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_locker2_obtainable x43 x42. Hypothesis P_id_imp_bounded : forall x42 x43, (min_value <= x42) ->(min_value <= x43) ->min_value <= P_id_imp x43 x42. Hypothesis P_id_stack_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_stack x43 x42. Hypothesis P_id_locker2_map_promote_pending_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) -> min_value <= P_id_locker2_map_promote_pending x43 x42. Hypothesis P_id_locker_bounded : min_value <= P_id_locker . Hypothesis P_id_case4_bounded : forall x44 x42 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) ->min_value <= P_id_case4 x44 x43 x42. Hypothesis P_id_int_bounded : forall x42, (min_value <= x42) ->min_value <= P_id_int x42. Hypothesis P_id_push_bounded : forall x44 x42 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) ->min_value <= P_id_push x44 x43 x42. Hypothesis P_id_locker2_add_pending_bounded : forall x44 x42 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) ->min_value <= P_id_locker2_add_pending x44 x43 x42. Hypothesis P_id_true_bounded : min_value <= P_id_true . Hypothesis P_id_locker2_check_availables_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_locker2_check_availables x43 x42. Hypothesis P_id_F_bounded : min_value <= P_id_F . Hypothesis P_id_eqs_bounded : forall x42 x43, (min_value <= x42) ->(min_value <= x43) ->min_value <= P_id_eqs x43 x42. Hypothesis P_id_record_update_bounded : forall x44 x42 x45 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) -> (min_value <= x45) ->min_value <= P_id_record_update x45 x44 x43 x42. Hypothesis P_id_false_bounded : min_value <= P_id_false . Hypothesis P_id_gen_modtageq_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_gen_modtageq x43 x42. Hypothesis P_id_undefined_bounded : min_value <= P_id_undefined . Hypothesis P_id_nocalls_bounded : min_value <= P_id_nocalls . Hypothesis P_id_locker2_remove_pending_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_locker2_remove_pending x43 x42. Hypothesis P_id_resource_bounded : min_value <= P_id_resource . Hypothesis P_id_case6_bounded : forall x44 x42 x45 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) -> (min_value <= x45) ->min_value <= P_id_case6 x45 x44 x43 x42. Hypothesis P_id_if_bounded : forall x44 x42 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) ->min_value <= P_id_if x44 x43 x42. Hypothesis P_id_pops_bounded : forall x42, (min_value <= x42) ->min_value <= P_id_pops x42. Hypothesis P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) -> min_value <= P_id_locker2_map_claim_lock x44 x43 x42. Hypothesis P_id_ok_bounded : min_value <= P_id_ok . Hypothesis P_id_case5_bounded : forall x44 x42 x45 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) -> (min_value <= x45) ->min_value <= P_id_case5 x45 x44 x43 x42. Hypothesis P_id_tuple_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_tuple x43 x42. Hypothesis P_id_member_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_member x43 x42. Hypothesis P_id_s_bounded : forall x42, (min_value <= x42) ->min_value <= P_id_s x42. Hypothesis P_id_delete_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_delete x43 x42. Hypothesis P_id_T_bounded : min_value <= P_id_T . Hypothesis P_id_case9_bounded : forall x44 x42 x45 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) -> (min_value <= x45) ->min_value <= P_id_case9 x45 x44 x43 x42. Hypothesis P_id_record_extract_bounded : forall x44 x42 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) ->min_value <= P_id_record_extract x44 x43 x42. Hypothesis P_id_excl_bounded : min_value <= P_id_excl . Hypothesis P_id_case2_bounded : forall x44 x42 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) ->min_value <= P_id_case2 x44 x43 x42. Hypothesis P_id_nil_bounded : min_value <= P_id_nil . Hypothesis P_id_eqc_bounded : forall x42 x43, (min_value <= x42) ->(min_value <= x43) ->min_value <= P_id_eqc x43 x42. Hypothesis P_id_case0_bounded : forall x44 x42 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) ->min_value <= P_id_case0 x44 x43 x42. Hypothesis P_id_request_bounded : min_value <= P_id_request . Hypothesis P_id_locker2_check_available_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_locker2_check_available x43 x42. Hypothesis P_id_not_bounded : forall x42, (min_value <= x42) ->min_value <= P_id_not x42. Hypothesis P_id_pushs_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_pushs x43 x42. Hypothesis P_id_locker2_promote_pending_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_locker2_promote_pending x43 x42. Hypothesis P_id_mcrlrecord_bounded : min_value <= P_id_mcrlrecord . Hypothesis P_id_locker2_obtainables_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_locker2_obtainables x43 x42. Hypothesis P_id_cons_bounded : forall x42 x43, (min_value <= x42) ->(min_value <= x43) ->min_value <= P_id_cons x43 x42. Hypothesis P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) -> (min_value <= x45) -> (min_value <= x46) -> (min_value <= x47) -> min_value <= P_id_push1 x47 x46 x45 x44 x43 x42. Hypothesis P_id_case1_bounded : forall x44 x42 x45 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) -> (min_value <= x45) ->min_value <= P_id_case1 x45 x44 x43 x42. Hypothesis P_id_element_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_element x43 x42. Hypothesis P_id_locker2_adduniq_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_locker2_adduniq x43 x42. Hypothesis P_id_and_bounded : forall x42 x43, (min_value <= x42) ->(min_value <= x43) ->min_value <= P_id_and x43 x42. Hypothesis P_id_empty_bounded : min_value <= P_id_empty . Hypothesis P_id_record_updates_bounded : forall x44 x42 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) ->min_value <= P_id_record_updates x44 x43 x42. Hypothesis P_id_lock_bounded : min_value <= P_id_lock . Hypothesis P_id_excllock_bounded : min_value <= P_id_excllock . Hypothesis P_id_pid_bounded : forall x42, (min_value <= x42) ->min_value <= P_id_pid x42. Hypothesis P_id_calls_bounded : forall x44 x42 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) ->min_value <= P_id_calls x44 x43 x42. Hypothesis P_id_subtract_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_subtract x43 x42. Hypothesis P_id_tag_bounded : min_value <= P_id_tag . Hypothesis P_id_equal_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_equal x43 x42. Hypothesis P_id_eq_bounded : forall x42 x43, (min_value <= x42) ->(min_value <= x43) ->min_value <= P_id_eq x43 x42. Hypothesis P_id_tops_bounded : forall x42, (min_value <= x42) ->min_value <= P_id_tops x42. Hypothesis P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) ->min_value <= P_id_locker2_claim_lock x44 x43 x42. Hypothesis P_id_pending_bounded : min_value <= P_id_pending . Hypothesis P_id_andt_bounded : forall x42 x43, (min_value <= x42) ->(min_value <= x43) ->min_value <= P_id_andt x43 x42. Hypothesis P_id_tuplenil_bounded : forall x42, (min_value <= x42) ->min_value <= P_id_tuplenil x42. Hypothesis P_id_append_bounded : forall x42 x43, (min_value <= x42) -> (min_value <= x43) ->min_value <= P_id_append x43 x42. Hypothesis P_id_0_bounded : min_value <= P_id_0 . Hypothesis P_id_case8_bounded : forall x44 x42 x45 x43, (min_value <= x42) -> (min_value <= x43) -> (min_value <= x44) -> (min_value <= x45) ->min_value <= P_id_case8 x45 x44 x43 x42. Definition measure := Interp.measure min_value P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43::x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43:: x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43::x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44:: x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44:: x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Hypothesis P_id_GEN_MODTAGEQ : Z ->Z ->Z. Hypothesis P_id_ELEMENT : Z ->Z ->Z. Hypothesis P_id_Marked_eq : Z ->Z ->Z. Hypothesis P_id_CASE9 : Z ->Z ->Z ->Z ->Z. Hypothesis P_id_LOCKER2_REMOVE_PENDING : Z ->Z ->Z. Hypothesis P_id_Marked_record_new : Z ->Z. Hypothesis P_id_CASE6 : Z ->Z ->Z ->Z ->Z. Hypothesis P_id_LOCKER2_PROMOTE_PENDING : Z ->Z ->Z. Hypothesis P_id_Marked_if : Z ->Z ->Z ->Z. Hypothesis P_id_PUSH : Z ->Z ->Z ->Z. Hypothesis P_id_MEMBER : Z ->Z ->Z. Hypothesis P_id_CASE5 : Z ->Z ->Z ->Z ->Z. Hypothesis P_id_RECORD_UPDATE : Z ->Z ->Z ->Z ->Z. Hypothesis P_id_Marked_not : Z ->Z. Hypothesis P_id_ISTOPS : Z ->Z ->Z. Hypothesis P_id_LOCKER2_ADD_PENDING : Z ->Z ->Z ->Z. Hypothesis P_id_Marked_or : Z ->Z ->Z. Hypothesis P_id_DELETE : Z ->Z ->Z. Hypothesis P_id_CASE0 : Z ->Z ->Z ->Z. Hypothesis P_id_Marked_imp : Z ->Z ->Z. Hypothesis P_id_EQT : Z ->Z ->Z. Hypothesis P_id_PUSHS : Z ->Z ->Z. Hypothesis P_id_LOCKER2_RELEASE_LOCK : Z ->Z ->Z. Hypothesis P_id_LOCKER2_OBTAINABLES : Z ->Z ->Z. Hypothesis P_id_RECORD_UPDATES : Z ->Z ->Z ->Z. Hypothesis P_id_Marked_locker2_adduniq : Z ->Z ->Z. Hypothesis P_id_EQS : Z ->Z ->Z. Hypothesis P_id_SUBTRACT : Z ->Z ->Z. Hypothesis P_id_Marked_gen_tag : Z ->Z. Hypothesis P_id_LOCKER2_CHECK_AVAILABLES : Z ->Z ->Z. Hypothesis P_id_LOCKER2_MAP_CLAIM_LOCK : Z ->Z ->Z ->Z. Hypothesis P_id_Marked_case4 : Z ->Z ->Z ->Z. Hypothesis P_id_PUSH1 : Z ->Z ->Z ->Z ->Z ->Z ->Z. Hypothesis P_id_APPEND : Z ->Z ->Z. Hypothesis P_id_LOCKER2_CHECK_AVAILABLE : Z ->Z ->Z. Hypothesis P_id_LOCKER2_MAP_PROMOTE_PENDING : Z ->Z ->Z. Hypothesis P_id_Marked_pops : Z ->Z. Hypothesis P_id_EQC : Z ->Z ->Z. Hypothesis P_id_CASE1 : Z ->Z ->Z ->Z ->Z. Hypothesis P_id_CASE8 : Z ->Z ->Z ->Z ->Z. Hypothesis P_id_RECORD_EXTRACT : Z ->Z ->Z ->Z. Hypothesis P_id_Marked_locker2_map_add_pending : Z ->Z ->Z ->Z. Hypothesis P_id_AND : Z ->Z ->Z. Hypothesis P_id_Marked_tops : Z ->Z. Hypothesis P_id_CASE2 : Z ->Z ->Z ->Z. Hypothesis P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Hypothesis P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Hypothesis P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Hypothesis P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (min_value <= x49)/\ (x49 <= x48) -> (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Hypothesis P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Hypothesis P_id_Marked_record_new_monotonic : forall x42 x43, (min_value <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Hypothesis P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (min_value <= x49)/\ (x49 <= x48) -> (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Hypothesis P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Hypothesis P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Hypothesis P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Hypothesis P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Hypothesis P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (min_value <= x49)/\ (x49 <= x48) -> (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Hypothesis P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (min_value <= x49)/\ (x49 <= x48) -> (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Hypothesis P_id_Marked_not_monotonic : forall x42 x43, (min_value <= x43)/\ (x43 <= x42) -> P_id_Marked_not x43 <= P_id_Marked_not x42. Hypothesis P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Hypothesis P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Hypothesis P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Hypothesis P_id_DELETE_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Hypothesis P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Hypothesis P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Hypothesis P_id_EQT_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_EQT x43 x45 <= P_id_EQT x42 x44. Hypothesis P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Hypothesis P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Hypothesis P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Hypothesis P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Hypothesis P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Hypothesis P_id_EQS_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_EQS x43 x45 <= P_id_EQS x42 x44. Hypothesis P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Hypothesis P_id_Marked_gen_tag_monotonic : forall x42 x43, (min_value <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Hypothesis P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Hypothesis P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Hypothesis P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Hypothesis P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (min_value <= x53)/\ (x53 <= x52) -> (min_value <= x51)/\ (x51 <= x50) -> (min_value <= x49)/\ (x49 <= x48) -> (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Hypothesis P_id_APPEND_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Hypothesis P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Hypothesis P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Hypothesis P_id_Marked_pops_monotonic : forall x42 x43, (min_value <= x43)/\ (x43 <= x42) -> P_id_Marked_pops x43 <= P_id_Marked_pops x42. Hypothesis P_id_EQC_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_EQC x43 x45 <= P_id_EQC x42 x44. Hypothesis P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (min_value <= x49)/\ (x49 <= x48) -> (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Hypothesis P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (min_value <= x49)/\ (x49 <= x48) -> (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Hypothesis P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Hypothesis P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Hypothesis P_id_AND_monotonic : forall x44 x42 x45 x43, (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_AND x43 x45 <= P_id_AND x42 x44. Hypothesis P_id_Marked_tops_monotonic : forall x42 x43, (min_value <= x43)/\ (x43 <= x42) -> P_id_Marked_tops x43 <= P_id_Marked_tops x42. Hypothesis P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (min_value <= x47)/\ (x47 <= x46) -> (min_value <= x45)/\ (x45 <= x44) -> (min_value <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Definition marked_measure := Interp.marked_measure min_value P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44:: x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44:: x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44:: x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46:: x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44:: x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44:: x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_monotonic;assumption. Qed. End S. End InterpZ. Module WF_R_xml_0_deep_rew. Inductive DP_R_xml_0 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_0 : forall x4 x42 x43 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_pid (x7::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_pid (x4::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_eqt (x7::x4::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_1 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt (x8::x5::nil))::(algebra.Alg.Term algebra.F.id_eqt (x9::x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_2 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_eqt (x8::x5::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_3 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_eqt (x9::x6::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_4 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x5::x6::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt (x8::x5::nil))::(algebra.Alg.Term algebra.F.id_eqt (x9::x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_5 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x5::x6::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_eqt (x8::x5::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_6 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x5::x6::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_eqt (x9::x6::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_7 : forall x8 x42 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuplenil (x8::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuplenil (x5::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_eqt (x8::x5::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_8 : forall x42 x6 x9 x43 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_s (x7::nil))::nil))::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x9::x6::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_element ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Term algebra.F.id_s (x7::nil))::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) (* *) | DP_R_xml_0_9 : forall x16 x44 x42 x14 x17 x13 x43 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x14 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x15 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple (x16::(algebra.Alg.Term algebra.F.id_tuplenil (x13::nil))::nil))::x17::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Term algebra.F.id_record_update (x14::x15:: x16::x13::nil))::x15::x17::nil)) (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) (* *) | DP_R_xml_0_10 : forall x16 x44 x42 x14 x17 x13 x43 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x14 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x15 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple (x16::(algebra.Alg.Term algebra.F.id_tuplenil (x13::nil))::nil))::x17::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_record_update (x14::x15:: x16::x13::nil)) (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) (* *) | DP_R_xml_0_11 : forall x20 x18 x42 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x18 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x19::x18::nil)) (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) (* *) | DP_R_xml_0_12 : forall x20 x18 x42 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x18 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x20::x18::nil)) (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) (* *) | DP_R_xml_0_13 : forall x20 x44 x42 x22 x21 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x21 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x20::x21::x22::nil)) (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) (* *) | DP_R_xml_0_14 : forall x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_case0 (x22::x19:: (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) (* *) | DP_R_xml_0_15 : forall x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil)) (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) (* *) | DP_R_xml_0_16 : forall x44 x42 x22 x19 x43 x23, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x22::x23::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_record_updates (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_excl nil)::(algebra.Alg.Term algebra.F.id_tuplenil (x22::nil))::nil))::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_tuplenil (x23::nil))::nil)):: (algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_case0 (x44::x43::x42::nil)) (* *) | DP_R_xml_0_17 : forall x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_record_updates (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Term algebra.F.id_subtract ((algebra.Alg.Term algebra.F.id_record_extract (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil))::(algebra.Alg.Term algebra.F.id_cons (x22::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil))::nil))::nil)):: (algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) (* *) | DP_R_xml_0_18 : forall x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_subtract ((algebra.Alg.Term algebra.F.id_record_extract (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil))::(algebra.Alg.Term algebra.F.id_cons (x22::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) (* *) | DP_R_xml_0_19 : forall x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil)) (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) (* *) | DP_R_xml_0_20 : forall x44 x42 x22 x21 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x21 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_case1 (x22::x21::x19:: (algebra.Alg.Term algebra.F.id_member ((algebra.Alg.Term algebra.F.id_record_extract (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil))::x21::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) (* *) | DP_R_xml_0_21 : forall x44 x42 x22 x21 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x21 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_member ((algebra.Alg.Term algebra.F.id_record_extract (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil))::x21::nil)) (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) (* *) | DP_R_xml_0_22 : forall x44 x42 x22 x21 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x21 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_resource nil)::nil)) (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) (* *) | DP_R_xml_0_23 : forall x44 x42 x22 x21 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x21 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_true nil) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_record_updates (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Term algebra.F.id_record_extract (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil))::(algebra.Alg.Term algebra.F.id_cons (x22::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil))::nil))::nil)):: (algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43::x42::nil)) (* *) | DP_R_xml_0_24 : forall x44 x42 x22 x21 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x21 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_true nil) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Term algebra.F.id_record_extract (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil))::(algebra.Alg.Term algebra.F.id_cons (x22::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43::x42::nil)) (* *) | DP_R_xml_0_25 : forall x44 x42 x22 x21 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x21 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_true nil) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil)) (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43::x42::nil)) (* *) | DP_R_xml_0_26 : forall x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_case2 (x22::x19:: (algebra.Alg.Term algebra.F.id_gen_modtageq (x22:: (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_excl nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) (* *) | DP_R_xml_0_27 : forall x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_gen_modtageq (x22:: (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_excl nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) (* *) | DP_R_xml_0_28 : forall x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_excl nil)::nil)) (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) (* *) | DP_R_xml_0_29 : forall x44 x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_true nil) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_record_updates (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_excllock nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_case2 (x44::x43::x42::nil)) (* *) | DP_R_xml_0_30 : forall x20 x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_case5 (x22::x20::x19:: (algebra.Alg.Term algebra.F.id_member (x22:: (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) (* *) | DP_R_xml_0_31 : forall x20 x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_member (x22:: (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) (* *) | DP_R_xml_0_32 : forall x20 x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil)) (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) (* *) | DP_R_xml_0_33 : forall x20 x44 x42 x22 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_true nil) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_locker2_obtainables (x20:: x22::nil)) (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43::x42::nil)) (* *) | DP_R_xml_0_34 : forall x20 x44 x42 x22 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_false nil) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_locker2_obtainables (x20:: x22::nil)) (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43::x42::nil)) (* *) | DP_R_xml_0_35 : forall x20 x42 x26 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x26 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_case6 (x20::x19::x26:: (algebra.Alg.Term algebra.F.id_equal (x26:: (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_resource nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) (* *) | DP_R_xml_0_36 : forall x20 x42 x26 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x26 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_resource nil)::nil)) (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) (* *) | DP_R_xml_0_37 : forall x20 x44 x42 x26 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x26 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_true nil) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_excl nil)::nil)) (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43::x42::nil)) (* *) | DP_R_xml_0_38 : forall x20 x44 x42 x26 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x26 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_true nil) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil)) (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43::x42::nil)) (* *) | DP_R_xml_0_39 : forall x20 x44 x42 x26 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x26 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_false nil) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_locker2_check_available (x26::x20::nil)) (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43::x42::nil)) (* *) | DP_R_xml_0_40 : forall x20 x42 x26 x21 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x26::x21::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_locker2_check_available (x26::x20::nil)) (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) (* *) | DP_R_xml_0_41 : forall x20 x42 x26 x21 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x26::x21::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_locker2_check_availables (x21::x20::nil)) (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) (* *) | DP_R_xml_0_42 : forall x28 x42 x29 x43 x27, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x28::x29::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x27 x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_append (x29::x27::nil)) (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) (* *) | DP_R_xml_0_43 : forall x28 x42 x29 x43 x27, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x27 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x28::x29::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_subtract ((algebra.Alg.Term algebra.F.id_delete (x28::x27::nil))::x29::nil)) (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) (* *) | DP_R_xml_0_44 : forall x28 x42 x29 x43 x27, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x27 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x28::x29::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_delete (x28::x27::nil)) (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) (* *) | DP_R_xml_0_45 : forall x28 x42 x30 x29 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x30 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x28::x29::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_case8 (x29::x28::x30:: (algebra.Alg.Term algebra.F.id_equal (x30:: x28::nil))::nil)) (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) (* *) | DP_R_xml_0_46 : forall x44 x28 x42 x30 x45 x29 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x29 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x28 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x30 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_false nil) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_delete (x30::x29::nil)) (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43::x42::nil)) (* *) | DP_R_xml_0_47 : forall x28 x42 x30 x29 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x30 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x28::x29::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_case9 (x29::x28::x30:: (algebra.Alg.Term algebra.F.id_equal (x30:: x28::nil))::nil)) (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) (* *) | DP_R_xml_0_48 : forall x44 x28 x42 x30 x45 x29 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x29 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x28 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x30 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_false nil) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_member (x30::x29::nil)) (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43::x42::nil)) (* *) | DP_R_xml_0_49 : forall x36 x34 x42 x37 x35 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_stack (x36::x37::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_stack (x34::x35::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt (x36::x34::nil))::(algebra.Alg.Term algebra.F.id_eqs (x37::x35::nil))::nil)) (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) (* *) | DP_R_xml_0_50 : forall x36 x34 x42 x37 x35 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_stack (x36::x37::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_stack (x34::x35::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_eqt (x36::x34::nil)) (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) (* *) | DP_R_xml_0_51 : forall x36 x34 x42 x37 x35 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_stack (x36::x37::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_stack (x34::x35::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_eqs (x37::x35::nil)) (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) (* *) | DP_R_xml_0_52 : forall x36 x34 x42 x37 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x36 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_stack (x34::x37::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_eqt (x36::x34::nil)) (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) (* *) | DP_R_xml_0_53 : forall x36 x34 x42 x38 x37 x35 x43 x39, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x36::x37::x39::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x34::x35::x38::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt (x36::x34::nil))::(algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqs (x37::x35::nil))::(algebra.Alg.Term algebra.F.id_eqc (x39::x38::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) (* *) | DP_R_xml_0_54 : forall x36 x34 x42 x38 x37 x35 x43 x39, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x36::x37::x39::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x34::x35::x38::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_eqt (x36::x34::nil)) (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) (* *) | DP_R_xml_0_55 : forall x36 x34 x42 x38 x37 x35 x43 x39, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x36::x37::x39::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x34::x35::x38::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqs (x37::x35::nil))::(algebra.Alg.Term algebra.F.id_eqc (x39::x38::nil))::nil)) (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) (* *) | DP_R_xml_0_56 : forall x36 x34 x42 x38 x37 x35 x43 x39, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x36::x37::x39::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x34::x35::x38::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_eqs (x37::x35::nil)) (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) (* *) | DP_R_xml_0_57 : forall x36 x34 x42 x38 x37 x35 x43 x39, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x36::x37::x39::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x34::x35::x38::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_eqc (x39::x38::nil)) (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) (* *) | DP_R_xml_0_58 : forall x40 x36 x44 x34 x42 x37 x43 x39, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x36 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x34 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x40::x37::x39::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_push1 (x36::x34::x40::x37:: x39::(algebra.Alg.Term algebra.F.id_eqt (x36:: x40::nil))::nil)) (algebra.Alg.Term algebra.F.id_push (x44::x43::x42::nil)) (* *) | DP_R_xml_0_59 : forall x40 x36 x44 x34 x42 x37 x43 x39, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x36 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x34 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x40::x37::x39::nil)) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_eqt (x36::x40::nil)) (algebra.Alg.Term algebra.F.id_push (x44::x43::x42::nil)) (* *) | DP_R_xml_0_60 : forall x40 x36 x44 x34 x42 x46 x37 x45 x43 x39 x47, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x36 x47) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x34 x46) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x40 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x37 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x39 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_T nil) x42) -> DP_R_xml_0 (algebra.Alg.Term algebra.F.id_pushs (x34::x37::nil)) (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45::x44::x43:: x42::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 x40 x36 x44 x34 x42 x46 x37 x45 x43 x39 x47, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x36 x47) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x34 x46) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x40 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x37 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x39 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_T nil) x42) -> DP_R_xml_0_non_scc_1 (algebra.Alg.Term algebra.F.id_pushs (x34:: x37::nil)) (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45::x44::x43:: x42::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 x40 x36 x44 x34 x42 x37 x43 x39, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x36 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x34 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x40::x37::x39::nil)) x42) -> DP_R_xml_0_non_scc_2 (algebra.Alg.Term algebra.F.id_push1 (x36:: x34::x40::x37::x39::(algebra.Alg.Term algebra.F.id_eqt (x36::x40::nil))::nil)) (algebra.Alg.Term algebra.F.id_push (x44::x43::x42::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; (eapply acc_DP_R_xml_0_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_non_scc_3 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_3_0 : forall x36 x34 x42 x38 x37 x35 x43 x39, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x36::x37::x39::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x34::x35::x38::nil)) x42) -> DP_R_xml_0_non_scc_3 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqs (x37:: x35::nil))::(algebra.Alg.Term algebra.F.id_eqc (x39::x38::nil))::nil)) (algebra.Alg.Term algebra.F.id_eqc (x43::x42::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 x36 x34 x42 x38 x37 x35 x43 x39, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x36::x37::x39::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x34::x35::x38::nil)) x42) -> DP_R_xml_0_non_scc_4 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt (x36:: x34::nil))::(algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqs (x37:: x35::nil))::(algebra.Alg.Term algebra.F.id_eqc (x39::x38::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_eqc (x43::x42::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 x36 x34 x42 x37 x35 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_stack (x36::x37::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_stack (x34::x35::nil)) x42) -> DP_R_xml_0_non_scc_5 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt (x36:: x34::nil))::(algebra.Alg.Term algebra.F.id_eqs (x37::x35::nil))::nil)) (algebra.Alg.Term algebra.F.id_eqs (x43::x42::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 x28 x42 x30 x29 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x30 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x28::x29::nil)) x42) -> DP_R_xml_0_non_scc_6 (algebra.Alg.Term algebra.F.id_case9 (x29::x28:: x30::(algebra.Alg.Term algebra.F.id_equal (x30::x28::nil))::nil)) (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_6 : forall x y, (DP_R_xml_0_non_scc_6 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_7 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_7_0 : forall x44 x28 x42 x30 x45 x29 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x29 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x28 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x30 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_false nil) x42) -> DP_R_xml_0_non_scc_7 (algebra.Alg.Term algebra.F.id_member (x30:: x29::nil)) (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_7 : forall x y, (DP_R_xml_0_non_scc_7 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_non_scc_6; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_non_scc_8 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_8_0 : forall x28 x42 x30 x29 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x30 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x28::x29::nil)) x42) -> DP_R_xml_0_non_scc_8 (algebra.Alg.Term algebra.F.id_case8 (x29::x28:: x30::(algebra.Alg.Term algebra.F.id_equal (x30::x28::nil))::nil)) (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_8 : forall x y, (DP_R_xml_0_non_scc_8 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_9 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_9_0 : forall x44 x28 x42 x30 x45 x29 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x29 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x28 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x30 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_false nil) x42) -> DP_R_xml_0_non_scc_9 (algebra.Alg.Term algebra.F.id_delete (x30:: x29::nil)) (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_9 : forall x y, (DP_R_xml_0_non_scc_9 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_non_scc_8; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_non_scc_10 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_10_0 : forall x28 x42 x29 x43 x27, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x27 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x28::x29::nil)) x42) -> DP_R_xml_0_non_scc_10 (algebra.Alg.Term algebra.F.id_delete (x28:: x27::nil)) (algebra.Alg.Term algebra.F.id_subtract (x43::x42::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_non_scc_8; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_scc_11 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_11_0 : forall x28 x42 x29 x43 x27, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x27 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x28::x29::nil)) x42) -> DP_R_xml_0_scc_11 (algebra.Alg.Term algebra.F.id_subtract ((algebra.Alg.Term algebra.F.id_delete (x28:: x27::nil))::x29::nil)) (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) . Module WF_DP_R_xml_0_scc_11. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_or (x42:Z) (x43:Z) := 1. Definition P_id_gen_tag (x42:Z) := 2* x42. Definition P_id_record_new (x42:Z) := 0. Definition P_id_a := 0. Definition P_id_locker2_release_lock (x42:Z) (x43:Z) := 1 + 2* x42 + 2* x43. Definition P_id_eqt (x42:Z) (x43:Z) := 0. Definition P_id_istops (x42:Z) (x43:Z) := 2* x42. Definition P_id_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_release := 0. Definition P_id_locker2_obtainable (x42:Z) (x43:Z) := 0. Definition P_id_imp (x42:Z) (x43:Z) := 3 + 1* x43. Definition P_id_stack (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_locker2_map_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_locker := 0. Definition P_id_case4 (x42:Z) (x43:Z) (x44:Z) := 2. Definition P_id_int (x42:Z) := 0. Definition P_id_push (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 2* x43 + 3* x44. Definition P_id_locker2_add_pending (x42:Z) (x43:Z) (x44:Z) := 1 + 3* x42 + 2* x43 + 3* x44. Definition P_id_true := 0. Definition P_id_locker2_check_availables (x42:Z) (x43:Z) := 1* x42. Definition P_id_F := 0. Definition P_id_eqs (x42:Z) (x43:Z) := 0. Definition P_id_record_update (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x42 + 2* x45. Definition P_id_false := 0. Definition P_id_gen_modtageq (x42:Z) (x43:Z) := 2 + 2* x42 + 2* x43. Definition P_id_undefined := 0. Definition P_id_nocalls := 0. Definition P_id_locker2_remove_pending (x42:Z) (x43:Z) := 2 + 3* x42. Definition P_id_resource := 0. Definition P_id_case6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2 + 3* x42 + 2* x43. Definition P_id_if (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 1* x44. Definition P_id_pops (x42:Z) := 1* x42. Definition P_id_locker2_map_claim_lock (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_ok := 0. Definition P_id_case5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2* x43 + 2* x44. Definition P_id_tuple (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_member (x42:Z) (x43:Z) := 0. Definition P_id_s (x42:Z) := 0. Definition P_id_delete (x42:Z) (x43:Z) := 1* x43. Definition P_id_T := 0. Definition P_id_case9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_record_extract (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_excl := 0. Definition P_id_case2 (x42:Z) (x43:Z) (x44:Z) := 1 + 1* x42 + 1* x43. Definition P_id_nil := 0. Definition P_id_eqc (x42:Z) (x43:Z) := 0. Definition P_id_case0 (x42:Z) (x43:Z) (x44:Z) := 2 + 1* x43 + 2* x44. Definition P_id_request := 0. Definition P_id_locker2_check_available (x42:Z) (x43:Z) := 2* x43. Definition P_id_not (x42:Z) := 1. Definition P_id_pushs (x42:Z) (x43:Z) := 1* x42 + 2* x43. Definition P_id_locker2_promote_pending (x42:Z) (x43:Z) := 2 + 3* x42. Definition P_id_mcrlrecord := 0. Definition P_id_locker2_obtainables (x42:Z) (x43:Z) := 2* x42. Definition P_id_cons (x42:Z) (x43:Z) := 1 + 1* x42 + 2* x43. Definition P_id_push1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 1* x42 + 1* x43 + 2* x46. Definition P_id_case1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 3* x42 + 1* x43 + 3* x44. Definition P_id_element (x42:Z) (x43:Z) := 1* x43. Definition P_id_locker2_adduniq (x42:Z) (x43:Z) := 1* x43. Definition P_id_and (x42:Z) (x43:Z) := 2* x42 + 2* x43. Definition P_id_empty := 0. Definition P_id_record_updates (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_lock := 0. Definition P_id_excllock := 0. Definition P_id_pid (x42:Z) := 0. Definition P_id_calls (x42:Z) (x43:Z) (x44:Z) := 1* x44. Definition P_id_subtract (x42:Z) (x43:Z) := 1* x42. Definition P_id_tag := 0. Definition P_id_equal (x42:Z) (x43:Z) := 0. Definition P_id_eq (x42:Z) (x43:Z) := 0. Definition P_id_tops (x42:Z) := 1* x42. Definition P_id_locker2_claim_lock (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_pending := 0. Definition P_id_andt (x42:Z) (x43:Z) := 0. Definition P_id_tuplenil (x42:Z) := 2* x42. Definition P_id_append (x42:Z) (x43:Z) := 1* x42. Definition P_id_0 := 0. Definition P_id_case8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 2* x42 + 1* x43. Lemma P_id_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_record_new x43 <= P_id_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_istops x43 x45 <= P_id_istops x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_member x43 x45 <= P_id_member x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_delete x43 x45 <= P_id_delete x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_append x43 x45 <= P_id_append x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_or_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_or x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_bounded : forall x42, (0 <= x42) ->0 <= P_id_gen_tag x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_bounded : forall x42, (0 <= x42) ->0 <= P_id_record_new x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_release_lock x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_istops x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_add_pending x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_release_bounded : 0 <= P_id_release . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainable x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_imp x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_stack x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_map_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker_bounded : 0 <= P_id_locker . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case4 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_bounded : forall x42, (0 <= x42) ->0 <= P_id_int x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_push x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_add_pending x44 x43 x42. 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_locker2_check_availables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_availables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_F_bounded : 0 <= P_id_F . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_record_update x45 x44 x43 x42. 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_gen_modtageq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_gen_modtageq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_undefined_bounded : 0 <= P_id_undefined . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nocalls_bounded : 0 <= P_id_nocalls . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_remove_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_resource_bounded : 0 <= P_id_resource . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case6 x45 x44 x43 x42. 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 x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_if x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_bounded : forall x42, (0 <= x42) ->0 <= P_id_pops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ok_bounded : 0 <= P_id_ok . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case5 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_tuple x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_member x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x42, (0 <= x42) ->0 <= P_id_s x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_delete x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_T_bounded : 0 <= P_id_T . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case9 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_extract x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excl_bounded : 0 <= P_id_excl . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case2 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nil_bounded : 0 <= P_id_nil . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqc x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case0 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_request_bounded : 0 <= P_id_request . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_available x43 x42. 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 x42, (0 <= x42) ->0 <= P_id_not x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_pushs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_mcrlrecord_bounded : 0 <= P_id_mcrlrecord . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_cons x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (0 <= x42) -> (0 <= x43) -> (0 <= x44) -> (0 <= x45) -> (0 <= x46) ->(0 <= x47) ->0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case1 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_element x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_adduniq x43 x42. 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 x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_and x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_empty_bounded : 0 <= P_id_empty . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_updates x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_lock_bounded : 0 <= P_id_lock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excllock_bounded : 0 <= P_id_excllock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_bounded : forall x42, (0 <= x42) ->0 <= P_id_pid x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_calls x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_subtract x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tag_bounded : 0 <= P_id_tag . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_equal x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_bounded : forall x42, (0 <= x42) ->0 <= P_id_tops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pending_bounded : 0 <= P_id_pending . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_andt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_bounded : forall x42, (0 <= x42) ->0 <= P_id_tuplenil x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_append x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case8 x45 x44 x43 x42. 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_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43::x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43:: x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43::x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44:: x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44:: x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_GEN_MODTAGEQ (x42:Z) (x43:Z) := 0. Definition P_id_ELEMENT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_eq (x42:Z) (x43:Z) := 0. Definition P_id_CASE9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_REMOVE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_record_new (x42:Z) := 0. Definition P_id_CASE6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_if (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_MEMBER (x42:Z) (x43:Z) := 0. Definition P_id_CASE5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_UPDATE (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_Marked_not (x42:Z) := 0. Definition P_id_ISTOPS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_ADD_PENDING (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_or (x42:Z) (x43:Z) := 0. Definition P_id_DELETE (x42:Z) (x43:Z) := 0. Definition P_id_CASE0 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_imp (x42:Z) (x43:Z) := 0. Definition P_id_EQT (x42:Z) (x43:Z) := 0. Definition P_id_PUSHS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_RELEASE_LOCK (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_OBTAINABLES (x42:Z) (x43:Z) := 0. Definition P_id_RECORD_UPDATES (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_adduniq (x42:Z) (x43:Z) := 0. Definition P_id_EQS (x42:Z) (x43:Z) := 0. Definition P_id_SUBTRACT (x42:Z) (x43:Z) := 1* x43. Definition P_id_Marked_gen_tag (x42:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLES (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_CLAIM_LOCK (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_case4 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_APPEND (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLE (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_pops (x42:Z) := 0. Definition P_id_EQC (x42:Z) (x43:Z) := 0. Definition P_id_CASE1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_CASE8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_EXTRACT (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_map_add_pending (x42:Z) (x43:Z) (x44: Z) := 0. Definition P_id_AND (x42:Z) (x43:Z) := 0. Definition P_id_Marked_tops (x42:Z) := 0. Definition P_id_CASE2 (x42:Z) (x43:Z) (x44:Z) := 0. Lemma P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_not_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_not x43 <= P_id_Marked_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_DELETE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_pops x43 <= P_id_Marked_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQC_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_tops x43 <= P_id_Marked_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44:: x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44:: x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44:: x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46:: x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44:: x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44:: x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_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_11. Proof. intros x. apply well_founded_ind with (R:=fun x y => (Zwf.Zwf 0) (marked_measure x) (marked_measure y)). apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. clear x. intros x IHx. repeat ( constructor;inversion 1;subst; full_prove_ineq algebra.Alg.Term ltac:(algebra.Alg_ext.find_replacement ) algebra.EQT_ext.one_step_list_refl_trans_clos marked_measure marked_measure_star_monotonic (Zwf.Zwf 0) (interp.o_Z 0) ltac:(fun _ => R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 ) ltac:(fun _ => rewrite_and_unfold ) ltac:(fun _ => generate_pos_hyp ) ltac:(fun _ => cbv beta iota zeta delta - [Zplus Zmult Zle Zlt] in * ; try (constructor)) IHx ). Qed. End WF_DP_R_xml_0_scc_11. Definition wf_DP_R_xml_0_scc_11 := WF_DP_R_xml_0_scc_11.wf. Lemma acc_DP_R_xml_0_scc_11 : forall x y, (DP_R_xml_0_scc_11 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_11). intros x' _ Hrec y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply Hrec;econstructor eassumption)|| ((eapply acc_DP_R_xml_0_non_scc_10; 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_11. Qed. Inductive DP_R_xml_0_scc_12 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_12_0 : forall x28 x42 x29 x43 x27, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x28::x29::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x27 x42) -> DP_R_xml_0_scc_12 (algebra.Alg.Term algebra.F.id_append (x29:: x27::nil)) (algebra.Alg.Term algebra.F.id_append (x43::x42::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_or (x42:Z) (x43:Z) := 0. Definition P_id_gen_tag (x42:Z) := 2* x42. Definition P_id_record_new (x42:Z) := 3 + 1* x42. Definition P_id_a := 0. Definition P_id_locker2_release_lock (x42:Z) (x43:Z) := 3 + 1* x42 + 3* x43. Definition P_id_eqt (x42:Z) (x43:Z) := 0. Definition P_id_istops (x42:Z) (x43:Z) := 1* x42. Definition P_id_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_release := 0. Definition P_id_locker2_obtainable (x42:Z) (x43:Z) := 0. Definition P_id_imp (x42:Z) (x43:Z) := 1* x43. Definition P_id_stack (x42:Z) (x43:Z) := 2* x42 + 1* x43. Definition P_id_locker2_map_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_locker := 0. Definition P_id_case4 (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_int (x42:Z) := 0. Definition P_id_push (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 2* x43. Definition P_id_locker2_add_pending (x42:Z) (x43:Z) (x44:Z) := 1 + 2* x42 + 2* x43 + 3* x44. Definition P_id_true := 0. Definition P_id_locker2_check_availables (x42:Z) (x43:Z) := 2 + 2* x42. Definition P_id_F := 0. Definition P_id_eqs (x42:Z) (x43:Z) := 0. Definition P_id_record_update (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x42 + 1* x45. Definition P_id_false := 0. Definition P_id_gen_modtageq (x42:Z) (x43:Z) := 1 + 3* x43. Definition P_id_undefined := 0. Definition P_id_nocalls := 0. Definition P_id_locker2_remove_pending (x42:Z) (x43:Z) := 1 + 3* x42. Definition P_id_resource := 1. Definition P_id_case6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_if (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 1* x44. Definition P_id_pops (x42:Z) := 1* x42. Definition P_id_locker2_map_claim_lock (x42:Z) (x43:Z) (x44:Z) := 2 + 2* x42. Definition P_id_ok := 0. Definition P_id_case5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x43. Definition P_id_tuple (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_member (x42:Z) (x43:Z) := 3* x43. Definition P_id_s (x42:Z) := 0. Definition P_id_delete (x42:Z) (x43:Z) := 1* x43. Definition P_id_T := 0. Definition P_id_case9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 3* x42 + 3* x43. Definition P_id_record_extract (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_excl := 0. Definition P_id_case2 (x42:Z) (x43:Z) (x44:Z) := 2 + 1* x43. Definition P_id_nil := 0. Definition P_id_eqc (x42:Z) (x43:Z) := 0. Definition P_id_case0 (x42:Z) (x43:Z) (x44:Z) := 2 + 2* x43 + 1* x44. Definition P_id_request := 0. Definition P_id_locker2_check_available (x42:Z) (x43:Z) := 0. Definition P_id_not (x42:Z) := 1. Definition P_id_pushs (x42:Z) (x43:Z) := 2 + 2* x42 + 1* x43. Definition P_id_locker2_promote_pending (x42:Z) (x43:Z) := 2 + 3* x42. Definition P_id_mcrlrecord := 0. Definition P_id_locker2_obtainables (x42:Z) (x43:Z) := 1* x42. Definition P_id_cons (x42:Z) (x43:Z) := 1 + 1* x42 + 1* x43. Definition P_id_push1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 1* x42. Definition P_id_case1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 3* x42 + 1* x43 + 2* x44. Definition P_id_element (x42:Z) (x43:Z) := 1* x43. Definition P_id_locker2_adduniq (x42:Z) (x43:Z) := 1* x43. Definition P_id_and (x42:Z) (x43:Z) := 1* x42 + 2* x43. Definition P_id_empty := 0. Definition P_id_record_updates (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_lock := 1. Definition P_id_excllock := 0. Definition P_id_pid (x42:Z) := 0. Definition P_id_calls (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_subtract (x42:Z) (x43:Z) := 1* x42. Definition P_id_tag := 0. Definition P_id_equal (x42:Z) (x43:Z) := 0. Definition P_id_eq (x42:Z) (x43:Z) := 0. Definition P_id_tops (x42:Z) := 1* x42. Definition P_id_locker2_claim_lock (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_pending := 0. Definition P_id_andt (x42:Z) (x43:Z) := 0. Definition P_id_tuplenil (x42:Z) := 1* x42. Definition P_id_append (x42:Z) (x43:Z) := 1* x42. Definition P_id_0 := 0. Definition P_id_case8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x42 + 1* x43. Lemma P_id_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_record_new x43 <= P_id_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_istops x43 x45 <= P_id_istops x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_member x43 x45 <= P_id_member x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_delete x43 x45 <= P_id_delete x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_append x43 x45 <= P_id_append x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_or_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_or x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_bounded : forall x42, (0 <= x42) ->0 <= P_id_gen_tag x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_bounded : forall x42, (0 <= x42) ->0 <= P_id_record_new x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_release_lock x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_istops x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_add_pending x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_release_bounded : 0 <= P_id_release . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainable x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_imp x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_stack x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_map_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker_bounded : 0 <= P_id_locker . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case4 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_bounded : forall x42, (0 <= x42) ->0 <= P_id_int x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_push x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_add_pending x44 x43 x42. 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_locker2_check_availables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_availables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_F_bounded : 0 <= P_id_F . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_record_update x45 x44 x43 x42. 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_gen_modtageq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_gen_modtageq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_undefined_bounded : 0 <= P_id_undefined . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nocalls_bounded : 0 <= P_id_nocalls . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_remove_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_resource_bounded : 0 <= P_id_resource . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case6 x45 x44 x43 x42. 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 x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_if x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_bounded : forall x42, (0 <= x42) ->0 <= P_id_pops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ok_bounded : 0 <= P_id_ok . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case5 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_tuple x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_member x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x42, (0 <= x42) ->0 <= P_id_s x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_delete x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_T_bounded : 0 <= P_id_T . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case9 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_extract x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excl_bounded : 0 <= P_id_excl . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case2 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nil_bounded : 0 <= P_id_nil . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqc x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case0 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_request_bounded : 0 <= P_id_request . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_available x43 x42. 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 x42, (0 <= x42) ->0 <= P_id_not x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_pushs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_mcrlrecord_bounded : 0 <= P_id_mcrlrecord . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_cons x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (0 <= x42) -> (0 <= x43) -> (0 <= x44) -> (0 <= x45) -> (0 <= x46) ->(0 <= x47) ->0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case1 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_element x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_adduniq x43 x42. 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 x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_and x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_empty_bounded : 0 <= P_id_empty . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_updates x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_lock_bounded : 0 <= P_id_lock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excllock_bounded : 0 <= P_id_excllock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_bounded : forall x42, (0 <= x42) ->0 <= P_id_pid x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_calls x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_subtract x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tag_bounded : 0 <= P_id_tag . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_equal x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_bounded : forall x42, (0 <= x42) ->0 <= P_id_tops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pending_bounded : 0 <= P_id_pending . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_andt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_bounded : forall x42, (0 <= x42) ->0 <= P_id_tuplenil x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_append x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case8 x45 x44 x43 x42. 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_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43::x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43:: x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43::x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44:: x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44:: x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_GEN_MODTAGEQ (x42:Z) (x43:Z) := 0. Definition P_id_ELEMENT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_eq (x42:Z) (x43:Z) := 0. Definition P_id_CASE9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_REMOVE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_record_new (x42:Z) := 0. Definition P_id_CASE6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_if (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_MEMBER (x42:Z) (x43:Z) := 0. Definition P_id_CASE5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_UPDATE (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_Marked_not (x42:Z) := 0. Definition P_id_ISTOPS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_ADD_PENDING (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_or (x42:Z) (x43:Z) := 0. Definition P_id_DELETE (x42:Z) (x43:Z) := 0. Definition P_id_CASE0 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_imp (x42:Z) (x43:Z) := 0. Definition P_id_EQT (x42:Z) (x43:Z) := 0. Definition P_id_PUSHS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_RELEASE_LOCK (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_OBTAINABLES (x42:Z) (x43:Z) := 0. Definition P_id_RECORD_UPDATES (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_adduniq (x42:Z) (x43:Z) := 0. Definition P_id_EQS (x42:Z) (x43:Z) := 0. Definition P_id_SUBTRACT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_gen_tag (x42:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLES (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_CLAIM_LOCK (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_case4 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_APPEND (x42:Z) (x43:Z) := 1* x42. Definition P_id_LOCKER2_CHECK_AVAILABLE (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_pops (x42:Z) := 0. Definition P_id_EQC (x42:Z) (x43:Z) := 0. Definition P_id_CASE1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_CASE8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_EXTRACT (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_map_add_pending (x42:Z) (x43:Z) (x44: Z) := 0. Definition P_id_AND (x42:Z) (x43:Z) := 0. Definition P_id_Marked_tops (x42:Z) := 0. Definition P_id_CASE2 (x42:Z) (x43:Z) (x44:Z) := 0. Lemma P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_not_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_not x43 <= P_id_Marked_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_DELETE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_pops x43 <= P_id_Marked_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQC_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_tops x43 <= P_id_Marked_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44:: x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44:: x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44:: x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46:: x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44:: x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44:: x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_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)|| ((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_non_scc_13 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_13_0 : forall x20 x44 x42 x26 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x26 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_true nil) x42) -> DP_R_xml_0_non_scc_13 (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil)) (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_13 : forall x y, (DP_R_xml_0_non_scc_13 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_14 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_14_0 : forall x20 x44 x42 x26 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x26 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_true nil) x42) -> DP_R_xml_0_non_scc_14 (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_excl nil)::nil)) (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43::x42::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; (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 x20 x42 x26 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x26 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x42) -> DP_R_xml_0_non_scc_15 (algebra.Alg.Term algebra.F.id_record_extract (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_15 : forall x y, (DP_R_xml_0_non_scc_15 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_16 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_16_0 : forall x20 x42 x26 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x26 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x42) -> DP_R_xml_0_non_scc_16 (algebra.Alg.Term algebra.F.id_case6 (x20:: x19::x26::(algebra.Alg.Term algebra.F.id_equal (x26::(algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_resource nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::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_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' ))|| ((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 x20 x42 x26 x21 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x26::x21::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x42) -> DP_R_xml_0_non_scc_17 (algebra.Alg.Term algebra.F.id_locker2_check_available (x26:: x20::nil)) (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::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_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' ))|| ((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_18 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_18_0 : forall x20 x42 x26 x21 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x26::x21::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x42) -> DP_R_xml_0_scc_18 (algebra.Alg.Term algebra.F.id_locker2_check_availables (x21:: x20::nil)) (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) . Module WF_DP_R_xml_0_scc_18. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_or (x42:Z) (x43:Z) := 0. Definition P_id_gen_tag (x42:Z) := 2* x42. Definition P_id_record_new (x42:Z) := 3 + 1* x42. Definition P_id_a := 0. Definition P_id_locker2_release_lock (x42:Z) (x43:Z) := 3 + 1* x42 + 3* x43. Definition P_id_eqt (x42:Z) (x43:Z) := 0. Definition P_id_istops (x42:Z) (x43:Z) := 1* x42. Definition P_id_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_release := 0. Definition P_id_locker2_obtainable (x42:Z) (x43:Z) := 0. Definition P_id_imp (x42:Z) (x43:Z) := 1* x43. Definition P_id_stack (x42:Z) (x43:Z) := 2* x42 + 1* x43. Definition P_id_locker2_map_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_locker := 0. Definition P_id_case4 (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_int (x42:Z) := 0. Definition P_id_push (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 2* x43. Definition P_id_locker2_add_pending (x42:Z) (x43:Z) (x44:Z) := 1 + 2* x42 + 2* x43 + 3* x44. Definition P_id_true := 0. Definition P_id_locker2_check_availables (x42:Z) (x43:Z) := 2 + 2* x42. Definition P_id_F := 0. Definition P_id_eqs (x42:Z) (x43:Z) := 0. Definition P_id_record_update (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x42 + 1* x45. Definition P_id_false := 0. Definition P_id_gen_modtageq (x42:Z) (x43:Z) := 1 + 3* x43. Definition P_id_undefined := 0. Definition P_id_nocalls := 0. Definition P_id_locker2_remove_pending (x42:Z) (x43:Z) := 1 + 3* x42. Definition P_id_resource := 1. Definition P_id_case6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_if (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 1* x44. Definition P_id_pops (x42:Z) := 1* x42. Definition P_id_locker2_map_claim_lock (x42:Z) (x43:Z) (x44:Z) := 2 + 2* x42. Definition P_id_ok := 0. Definition P_id_case5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x43. Definition P_id_tuple (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_member (x42:Z) (x43:Z) := 3* x43. Definition P_id_s (x42:Z) := 0. Definition P_id_delete (x42:Z) (x43:Z) := 1* x43. Definition P_id_T := 0. Definition P_id_case9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 3* x42 + 3* x43. Definition P_id_record_extract (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_excl := 0. Definition P_id_case2 (x42:Z) (x43:Z) (x44:Z) := 2 + 1* x43. Definition P_id_nil := 0. Definition P_id_eqc (x42:Z) (x43:Z) := 0. Definition P_id_case0 (x42:Z) (x43:Z) (x44:Z) := 2 + 2* x43 + 1* x44. Definition P_id_request := 0. Definition P_id_locker2_check_available (x42:Z) (x43:Z) := 0. Definition P_id_not (x42:Z) := 1. Definition P_id_pushs (x42:Z) (x43:Z) := 2 + 2* x42 + 1* x43. Definition P_id_locker2_promote_pending (x42:Z) (x43:Z) := 2 + 3* x42. Definition P_id_mcrlrecord := 0. Definition P_id_locker2_obtainables (x42:Z) (x43:Z) := 1* x42. Definition P_id_cons (x42:Z) (x43:Z) := 1 + 1* x42 + 1* x43. Definition P_id_push1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 1* x42. Definition P_id_case1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 3* x42 + 1* x43 + 2* x44. Definition P_id_element (x42:Z) (x43:Z) := 1* x43. Definition P_id_locker2_adduniq (x42:Z) (x43:Z) := 1* x43. Definition P_id_and (x42:Z) (x43:Z) := 1* x42 + 2* x43. Definition P_id_empty := 0. Definition P_id_record_updates (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_lock := 1. Definition P_id_excllock := 0. Definition P_id_pid (x42:Z) := 0. Definition P_id_calls (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_subtract (x42:Z) (x43:Z) := 1* x42. Definition P_id_tag := 0. Definition P_id_equal (x42:Z) (x43:Z) := 0. Definition P_id_eq (x42:Z) (x43:Z) := 0. Definition P_id_tops (x42:Z) := 1* x42. Definition P_id_locker2_claim_lock (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_pending := 0. Definition P_id_andt (x42:Z) (x43:Z) := 0. Definition P_id_tuplenil (x42:Z) := 1* x42. Definition P_id_append (x42:Z) (x43:Z) := 1* x42. Definition P_id_0 := 0. Definition P_id_case8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x42 + 1* x43. Lemma P_id_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_record_new x43 <= P_id_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_istops x43 x45 <= P_id_istops x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_member x43 x45 <= P_id_member x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_delete x43 x45 <= P_id_delete x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_append x43 x45 <= P_id_append x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_or_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_or x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_bounded : forall x42, (0 <= x42) ->0 <= P_id_gen_tag x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_bounded : forall x42, (0 <= x42) ->0 <= P_id_record_new x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_release_lock x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_istops x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_add_pending x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_release_bounded : 0 <= P_id_release . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainable x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_imp x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_stack x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_map_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker_bounded : 0 <= P_id_locker . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case4 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_bounded : forall x42, (0 <= x42) ->0 <= P_id_int x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_push x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_add_pending x44 x43 x42. 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_locker2_check_availables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_availables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_F_bounded : 0 <= P_id_F . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_record_update x45 x44 x43 x42. 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_gen_modtageq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_gen_modtageq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_undefined_bounded : 0 <= P_id_undefined . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nocalls_bounded : 0 <= P_id_nocalls . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_remove_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_resource_bounded : 0 <= P_id_resource . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case6 x45 x44 x43 x42. 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 x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_if x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_bounded : forall x42, (0 <= x42) ->0 <= P_id_pops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ok_bounded : 0 <= P_id_ok . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case5 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_tuple x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_member x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x42, (0 <= x42) ->0 <= P_id_s x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_delete x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_T_bounded : 0 <= P_id_T . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case9 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_extract x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excl_bounded : 0 <= P_id_excl . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case2 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nil_bounded : 0 <= P_id_nil . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqc x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case0 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_request_bounded : 0 <= P_id_request . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_available x43 x42. 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 x42, (0 <= x42) ->0 <= P_id_not x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_pushs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_mcrlrecord_bounded : 0 <= P_id_mcrlrecord . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_cons x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (0 <= x42) -> (0 <= x43) -> (0 <= x44) -> (0 <= x45) -> (0 <= x46) ->(0 <= x47) ->0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case1 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_element x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_adduniq x43 x42. 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 x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_and x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_empty_bounded : 0 <= P_id_empty . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_updates x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_lock_bounded : 0 <= P_id_lock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excllock_bounded : 0 <= P_id_excllock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_bounded : forall x42, (0 <= x42) ->0 <= P_id_pid x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_calls x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_subtract x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tag_bounded : 0 <= P_id_tag . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_equal x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_bounded : forall x42, (0 <= x42) ->0 <= P_id_tops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pending_bounded : 0 <= P_id_pending . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_andt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_bounded : forall x42, (0 <= x42) ->0 <= P_id_tuplenil x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_append x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case8 x45 x44 x43 x42. 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_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43::x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43:: x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43::x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44:: x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44:: x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_GEN_MODTAGEQ (x42:Z) (x43:Z) := 0. Definition P_id_ELEMENT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_eq (x42:Z) (x43:Z) := 0. Definition P_id_CASE9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_REMOVE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_record_new (x42:Z) := 0. Definition P_id_CASE6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_if (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_MEMBER (x42:Z) (x43:Z) := 0. Definition P_id_CASE5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_UPDATE (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_Marked_not (x42:Z) := 0. Definition P_id_ISTOPS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_ADD_PENDING (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_or (x42:Z) (x43:Z) := 0. Definition P_id_DELETE (x42:Z) (x43:Z) := 0. Definition P_id_CASE0 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_imp (x42:Z) (x43:Z) := 0. Definition P_id_EQT (x42:Z) (x43:Z) := 0. Definition P_id_PUSHS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_RELEASE_LOCK (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_OBTAINABLES (x42:Z) (x43:Z) := 0. Definition P_id_RECORD_UPDATES (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_adduniq (x42:Z) (x43:Z) := 0. Definition P_id_EQS (x42:Z) (x43:Z) := 0. Definition P_id_SUBTRACT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_gen_tag (x42:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLES (x42:Z) (x43:Z) := 1* x42. Definition P_id_LOCKER2_MAP_CLAIM_LOCK (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_case4 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_APPEND (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLE (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_pops (x42:Z) := 0. Definition P_id_EQC (x42:Z) (x43:Z) := 0. Definition P_id_CASE1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_CASE8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_EXTRACT (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_map_add_pending (x42:Z) (x43:Z) (x44: Z) := 0. Definition P_id_AND (x42:Z) (x43:Z) := 0. Definition P_id_Marked_tops (x42:Z) := 0. Definition P_id_CASE2 (x42:Z) (x43:Z) (x44:Z) := 0. Lemma P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_not_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_not x43 <= P_id_Marked_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_DELETE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_pops x43 <= P_id_Marked_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQC_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_tops x43 <= P_id_Marked_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44:: x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44:: x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44:: x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46:: x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44:: x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44:: x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_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_18. 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_18. Definition wf_DP_R_xml_0_scc_18 := WF_DP_R_xml_0_scc_18.wf. Lemma acc_DP_R_xml_0_scc_18 : forall x y, (DP_R_xml_0_scc_18 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_18). 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_17; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). apply wf_DP_R_xml_0_scc_18. Qed. Inductive DP_R_xml_0_non_scc_19 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_19_0 : forall x20 x44 x42 x26 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x26 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_false nil) x42) -> DP_R_xml_0_non_scc_19 (algebra.Alg.Term algebra.F.id_locker2_check_available (x26:: x20::nil)) (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_19 : forall x y, (DP_R_xml_0_non_scc_19 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_non_scc_16; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_15; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). Qed. Inductive DP_R_xml_0_non_scc_20 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_20_0 : forall x20 x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_non_scc_20 (algebra.Alg.Term algebra.F.id_record_extract (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::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 := (* *) | DP_R_xml_0_non_scc_21_0 : forall x20 x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_non_scc_21 (algebra.Alg.Term algebra.F.id_member (x22:: (algebra.Alg.Term algebra.F.id_record_extract (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::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; (eapply acc_DP_R_xml_0_non_scc_6; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_scc_22 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_22_0 : forall x20 x44 x42 x22 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_true nil) x42) -> DP_R_xml_0_scc_22 (algebra.Alg.Term algebra.F.id_locker2_obtainables (x20:: x22::nil)) (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43::x42::nil)) (* *) | DP_R_xml_0_scc_22_1 : forall x20 x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_scc_22 (algebra.Alg.Term algebra.F.id_case5 (x22::x20:: x19::(algebra.Alg.Term algebra.F.id_member (x22:: (algebra.Alg.Term algebra.F.id_record_extract (x19::(algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) (* *) | DP_R_xml_0_scc_22_2 : forall x20 x44 x42 x22 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_false nil) x42) -> DP_R_xml_0_scc_22 (algebra.Alg.Term algebra.F.id_locker2_obtainables (x20:: x22::nil)) (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43::x42::nil)) . Module WF_DP_R_xml_0_scc_22. Inductive DP_R_xml_0_scc_22_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_22_large_0 : forall x20 x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_scc_22_large (algebra.Alg.Term algebra.F.id_case5 (x22:: x20::x19::(algebra.Alg.Term algebra.F.id_member (x22:: (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) (* *) | DP_R_xml_0_scc_22_large_1 : forall x20 x44 x42 x22 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_false nil) x42) -> DP_R_xml_0_scc_22_large (algebra.Alg.Term algebra.F.id_locker2_obtainables (x20:: x22::nil)) (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43::x42::nil)) . Inductive DP_R_xml_0_scc_22_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_22_strict_0 : forall x20 x44 x42 x22 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_true nil) x42) -> DP_R_xml_0_scc_22_strict (algebra.Alg.Term algebra.F.id_locker2_obtainables (x20:: x22::nil)) (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43::x42::nil)) . Module WF_DP_R_xml_0_scc_22_large. Inductive DP_R_xml_0_scc_22_large_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_22_large_large_0 : forall x20 x44 x42 x22 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_false nil) x42) -> DP_R_xml_0_scc_22_large_large (algebra.Alg.Term algebra.F.id_locker2_obtainables (x20::x22::nil)) (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43::x42::nil)) . Inductive DP_R_xml_0_scc_22_large_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_22_large_strict_0 : forall x20 x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_scc_22_large_strict (algebra.Alg.Term algebra.F.id_case5 (x22::x20::x19:: (algebra.Alg.Term algebra.F.id_member (x22:: (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) . Module WF_DP_R_xml_0_scc_22_large_large. Inductive DP_R_xml_0_scc_22_large_large_non_scc_1 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_22_large_large_non_scc_1_0 : forall x20 x44 x42 x22 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x20 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_false nil) x42) -> DP_R_xml_0_scc_22_large_large_non_scc_1 (algebra.Alg.Term algebra.F.id_locker2_obtainables (x20::x22::nil)) (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43::x42::nil)) . Lemma acc_DP_R_xml_0_scc_22_large_large_non_scc_1 : forall x y, (DP_R_xml_0_scc_22_large_large_non_scc_1 x y) -> Acc WF_DP_R_xml_0_scc_22_large.DP_R_xml_0_scc_22_large_large x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Lemma wf : well_founded WF_DP_R_xml_0_scc_22_large.DP_R_xml_0_scc_22_large_large. Proof. constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_22_large_large_non_scc_1; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_22_large_large_non_scc_0; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )||(fail))). Qed. End WF_DP_R_xml_0_scc_22_large_large. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_or (x42:Z) (x43:Z) := 2. Definition P_id_gen_tag (x42:Z) := 1 + 2* x42. Definition P_id_record_new (x42:Z) := 1. Definition P_id_a := 0. Definition P_id_locker2_release_lock (x42:Z) (x43:Z) := 3 + 3* x42 + 2* x43. Definition P_id_eqt (x42:Z) (x43:Z) := 0. Definition P_id_istops (x42:Z) (x43:Z) := 3. Definition P_id_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_release := 0. Definition P_id_locker2_obtainable (x42:Z) (x43:Z) := 0. Definition P_id_imp (x42:Z) (x43:Z) := 1* x43. Definition P_id_stack (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_locker2_map_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_locker := 0. Definition P_id_case4 (x42:Z) (x43:Z) (x44:Z) := 1. Definition P_id_int (x42:Z) := 1* x42. Definition P_id_push (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_locker2_add_pending (x42:Z) (x43:Z) (x44:Z) := 3 + 3* x42 + 2* x43 + 3* x44. Definition P_id_true := 0. Definition P_id_locker2_check_availables (x42:Z) (x43:Z) := 2* x43. Definition P_id_F := 0. Definition P_id_eqs (x42:Z) (x43:Z) := 0. Definition P_id_record_update (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x42 + 1* x45. Definition P_id_false := 0. Definition P_id_gen_modtageq (x42:Z) (x43:Z) := 3 + 3* x42 + 3* x43. Definition P_id_undefined := 0. Definition P_id_nocalls := 0. Definition P_id_locker2_remove_pending (x42:Z) (x43:Z) := 2 + 3* x42 + 1* x43. Definition P_id_resource := 0. Definition P_id_case6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1. Definition P_id_if (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 1* x44. Definition P_id_pops (x42:Z) := 1* x42. Definition P_id_locker2_map_claim_lock (x42:Z) (x43:Z) (x44:Z) := 2* x42. Definition P_id_ok := 0. Definition P_id_case5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 2* x43 + 1* x44. Definition P_id_tuple (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_member (x42:Z) (x43:Z) := 0. Definition P_id_s (x42:Z) := 3* x42. Definition P_id_delete (x42:Z) (x43:Z) := 1* x43. Definition P_id_T := 0. Definition P_id_case9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_record_extract (x42:Z) (x43:Z) (x44:Z) := 2* x42. Definition P_id_excl := 0. Definition P_id_case2 (x42:Z) (x43:Z) (x44:Z) := 3 + 1* x42 + 2* x43. Definition P_id_nil := 0. Definition P_id_eqc (x42:Z) (x43:Z) := 0. Definition P_id_case0 (x42:Z) (x43:Z) (x44:Z) := 3 + 1* x43 + 1* x44. Definition P_id_request := 0. Definition P_id_locker2_check_available (x42:Z) (x43:Z) := 1. Definition P_id_not (x42:Z) := 0. Definition P_id_pushs (x42:Z) (x43:Z) := 2* x42 + 1* x43. Definition P_id_locker2_promote_pending (x42:Z) (x43:Z) := 3 + 3* x42. Definition P_id_mcrlrecord := 0. Definition P_id_locker2_obtainables (x42:Z) (x43:Z) := 2* x42. Definition P_id_cons (x42:Z) (x43:Z) := 2 + 1* x42 + 1* x43. Definition P_id_push1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_case1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2 + 3* x42 + 1* x43 + 3* x44. Definition P_id_element (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_locker2_adduniq (x42:Z) (x43:Z) := 1* x43. Definition P_id_and (x42:Z) (x43:Z) := 2* x42 + 2* x43. Definition P_id_empty := 0. Definition P_id_record_updates (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44 . Definition P_id_lock := 0. Definition P_id_excllock := 0. Definition P_id_pid (x42:Z) := 0. Definition P_id_calls (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_subtract (x42:Z) (x43:Z) := 1* x42. Definition P_id_tag := 1. Definition P_id_equal (x42:Z) (x43:Z) := 0. Definition P_id_eq (x42:Z) (x43:Z) := 2. Definition P_id_tops (x42:Z) := 1* x42. Definition P_id_locker2_claim_lock (x42:Z) (x43:Z) (x44:Z) := 1 + 1* x42. Definition P_id_pending := 0. Definition P_id_andt (x42:Z) (x43:Z) := 0. Definition P_id_tuplenil (x42:Z) := 1* x42. Definition P_id_append (x42:Z) (x43:Z) := 1* x42. Definition P_id_0 := 0. Definition P_id_case8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2 + 1* x42 + 1* x43. Lemma P_id_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_record_new x43 <= P_id_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_istops x43 x45 <= P_id_istops x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_member x43 x45 <= P_id_member x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_delete x43 x45 <= P_id_delete x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_append x43 x45 <= P_id_append x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_or_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_or x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_bounded : forall x42, (0 <= x42) ->0 <= P_id_gen_tag x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_bounded : forall x42, (0 <= x42) ->0 <= P_id_record_new x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_release_lock x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_istops x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->0 <= P_id_locker2_map_add_pending x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_release_bounded : 0 <= P_id_release . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainable x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_imp x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_stack x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_map_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker_bounded : 0 <= P_id_locker . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case4 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_bounded : forall x42, (0 <= x42) ->0 <= P_id_int x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_push x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_add_pending x44 x43 x42. 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_locker2_check_availables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_availables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_F_bounded : 0 <= P_id_F . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_record_update x45 x44 x43 x42. 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_gen_modtageq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_gen_modtageq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_undefined_bounded : 0 <= P_id_undefined . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nocalls_bounded : 0 <= P_id_nocalls . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_remove_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_resource_bounded : 0 <= P_id_resource . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case6 x45 x44 x43 x42. 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 x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_if x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_bounded : forall x42, (0 <= x42) ->0 <= P_id_pops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ok_bounded : 0 <= P_id_ok . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case5 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_tuple x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_member x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x42, (0 <= x42) ->0 <= P_id_s x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_delete x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_T_bounded : 0 <= P_id_T . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case9 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_extract x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excl_bounded : 0 <= P_id_excl . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case2 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nil_bounded : 0 <= P_id_nil . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqc x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case0 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_request_bounded : 0 <= P_id_request . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_available x43 x42. 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 x42, (0 <= x42) ->0 <= P_id_not x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_pushs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_mcrlrecord_bounded : 0 <= P_id_mcrlrecord . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_cons x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (0 <= x42) -> (0 <= x43) -> (0 <= x44) -> (0 <= x45) -> (0 <= x46) ->(0 <= x47) ->0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case1 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_element x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_adduniq x43 x42. 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 x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_and x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_empty_bounded : 0 <= P_id_empty . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_updates x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_lock_bounded : 0 <= P_id_lock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excllock_bounded : 0 <= P_id_excllock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_bounded : forall x42, (0 <= x42) ->0 <= P_id_pid x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_calls x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_subtract x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tag_bounded : 0 <= P_id_tag . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_equal x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_bounded : forall x42, (0 <= x42) ->0 <= P_id_tops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pending_bounded : 0 <= P_id_pending . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_andt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_bounded : forall x42, (0 <= x42) ->0 <= P_id_tuplenil x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_append x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case8 x45 x44 x43 x42. 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_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43::x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43::x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43::x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45:: x44::x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44::x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_GEN_MODTAGEQ (x42:Z) (x43:Z) := 0. Definition P_id_ELEMENT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_eq (x42:Z) (x43:Z) := 0. Definition P_id_CASE9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_REMOVE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_record_new (x42:Z) := 0. Definition P_id_CASE6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_if (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_MEMBER (x42:Z) (x43:Z) := 0. Definition P_id_CASE5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2* x43. Definition P_id_RECORD_UPDATE (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_Marked_not (x42:Z) := 0. Definition P_id_ISTOPS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_ADD_PENDING (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_or (x42:Z) (x43:Z) := 0. Definition P_id_DELETE (x42:Z) (x43:Z) := 0. Definition P_id_CASE0 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_imp (x42:Z) (x43:Z) := 0. Definition P_id_EQT (x42:Z) (x43:Z) := 0. Definition P_id_PUSHS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_RELEASE_LOCK (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_OBTAINABLES (x42:Z) (x43:Z) := 2* x42. Definition P_id_RECORD_UPDATES (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_adduniq (x42:Z) (x43:Z) := 0. Definition P_id_EQS (x42:Z) (x43:Z) := 0. Definition P_id_SUBTRACT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_gen_tag (x42:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLES (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_CLAIM_LOCK (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_case4 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_APPEND (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLE (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_pops (x42:Z) := 0. Definition P_id_EQC (x42:Z) (x43:Z) := 0. Definition P_id_CASE1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_CASE8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_EXTRACT (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_AND (x42:Z) (x43:Z) := 0. Definition P_id_Marked_tops (x42:Z) := 0. Definition P_id_CASE2 (x42:Z) (x43:Z) (x44:Z) := 0. Lemma P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_not_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_not x43 <= P_id_Marked_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_DELETE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_pops x43 <= P_id_Marked_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQC_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_tops x43 <= P_id_Marked_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44:: x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44:: x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44:: x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46:: x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44:: x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44:: x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44:: x43::x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_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_22_large_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_22_large_strict lt. Proof. unfold Relation_Definitions.inclusion, lt in *. intros a b H;destruct H; match goal with | |- (Zwf.Zwf 0) _ (marked_measure (algebra.Alg.Term ?f ?l)) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_lt_compat_right (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma DP_R_xml_0_scc_22_large_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_22_large_large le. Proof. unfold Relation_Definitions.inclusion, le, Zwf.Zwf in *. intros a b H;destruct H; match goal with | |- _ <= marked_measure (algebra.Alg.Term ?f ?l) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_trans (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition wf_DP_R_xml_0_scc_22_large_large := WF_DP_R_xml_0_scc_22_large_large.wf. Lemma wf : well_founded WF_DP_R_xml_0_scc_22.DP_R_xml_0_scc_22_large. Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_22_large_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx';apply DP_R_xml_0_scc_22_large_strict_in_lt; econstructor eassumption)|| ((apply IHx;[econstructor eassumption| intros y' Hlt;apply IHx';apply lt_le_compat with (1:=Hlt) ; apply DP_R_xml_0_scc_22_large_large_in_le;econstructor eassumption])). apply wf_DP_R_xml_0_scc_22_large_large. Qed. End WF_DP_R_xml_0_scc_22_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_or (x42:Z) (x43:Z) := 0. Definition P_id_gen_tag (x42:Z) := 3 + 2* x42. Definition P_id_record_new (x42:Z) := 1. Definition P_id_a := 0. Definition P_id_locker2_release_lock (x42:Z) (x43:Z) := 3* x42 + 1* x43. Definition P_id_eqt (x42:Z) (x43:Z) := 0. Definition P_id_istops (x42:Z) (x43:Z) := 1 + 1* x42 + 1* x43. Definition P_id_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_release := 0. Definition P_id_locker2_obtainable (x42:Z) (x43:Z) := 0. Definition P_id_imp (x42:Z) (x43:Z) := 2 + 1* x43. Definition P_id_stack (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_locker2_map_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_locker := 0. Definition P_id_case4 (x42:Z) (x43:Z) (x44:Z) := 1. Definition P_id_int (x42:Z) := 0. Definition P_id_push (x42:Z) (x43:Z) (x44:Z) := 3 + 1* x42 + 1* x43 + 2* x44. Definition P_id_locker2_add_pending (x42:Z) (x43:Z) (x44:Z) := 3* x42 + 2* x43 + 3* x44. Definition P_id_true := 1. Definition P_id_locker2_check_availables (x42:Z) (x43:Z) := 2 + 3* x42 + 3* x43. Definition P_id_F := 0. Definition P_id_eqs (x42:Z) (x43:Z) := 0. Definition P_id_record_update (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1* x42 + 1* x44 + 2* x45. Definition P_id_false := 0. Definition P_id_gen_modtageq (x42:Z) (x43:Z) := 0. Definition P_id_undefined := 0. Definition P_id_nocalls := 1. Definition P_id_locker2_remove_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_resource := 0. Definition P_id_case6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_if (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 1* x44. Definition P_id_pops (x42:Z) := 1* x42. Definition P_id_locker2_map_claim_lock (x42:Z) (x43:Z) (x44:Z) := 2* x42. Definition P_id_ok := 0. Definition P_id_case5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2. Definition P_id_tuple (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_member (x42:Z) (x43:Z) := 0. Definition P_id_s (x42:Z) := 0. Definition P_id_delete (x42:Z) (x43:Z) := 1* x43. Definition P_id_T := 0. Definition P_id_case9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2* x45. Definition P_id_record_extract (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_excl := 0. Definition P_id_case2 (x42:Z) (x43:Z) (x44:Z) := 2* x43. Definition P_id_nil := 0. Definition P_id_eqc (x42:Z) (x43:Z) := 0. Definition P_id_case0 (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 2* x44. Definition P_id_request := 0. Definition P_id_locker2_check_available (x42:Z) (x43:Z) := 0. Definition P_id_not (x42:Z) := 3. Definition P_id_pushs (x42:Z) (x43:Z) := 1 + 1* x42 + 2* x43. Definition P_id_locker2_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_mcrlrecord := 0. Definition P_id_locker2_obtainables (x42:Z) (x43:Z) := 2. Definition P_id_cons (x42:Z) (x43:Z) := 1* x42 + 2* x43. Definition P_id_push1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 3 + 1* x42. Definition P_id_case1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 3* x42 + 1* x43 + 3* x44 + 1* x45. Definition P_id_element (x42:Z) (x43:Z) := 2* x43. Definition P_id_locker2_adduniq (x42:Z) (x43:Z) := 1* x43. Definition P_id_and (x42:Z) (x43:Z) := 2* x42 + 2* x43. Definition P_id_empty := 0. Definition P_id_record_updates (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 2* x44. Definition P_id_lock := 0. Definition P_id_excllock := 0. Definition P_id_pid (x42:Z) := 0. Definition P_id_calls (x42:Z) (x43:Z) (x44:Z) := 2. Definition P_id_subtract (x42:Z) (x43:Z) := 1* x42. Definition P_id_tag := 0. Definition P_id_equal (x42:Z) (x43:Z) := 0. Definition P_id_eq (x42:Z) (x43:Z) := 0. Definition P_id_tops (x42:Z) := 1* x42. Definition P_id_locker2_claim_lock (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_pending := 0. Definition P_id_andt (x42:Z) (x43:Z) := 0. Definition P_id_tuplenil (x42:Z) := 1* x42. Definition P_id_append (x42:Z) (x43:Z) := 1* x42. Definition P_id_0 := 0. Definition P_id_case8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2* x42 + 1* x43. Lemma P_id_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_record_new x43 <= P_id_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_istops x43 x45 <= P_id_istops x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_member x43 x45 <= P_id_member x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_delete x43 x45 <= P_id_delete x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_append x43 x45 <= P_id_append x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_or_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_or x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_bounded : forall x42, (0 <= x42) ->0 <= P_id_gen_tag x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_bounded : forall x42, (0 <= x42) ->0 <= P_id_record_new x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_release_lock x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_istops x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_add_pending x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_release_bounded : 0 <= P_id_release . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainable x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_imp x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_stack x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_map_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker_bounded : 0 <= P_id_locker . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case4 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_bounded : forall x42, (0 <= x42) ->0 <= P_id_int x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_push x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_add_pending x44 x43 x42. 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_locker2_check_availables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_availables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_F_bounded : 0 <= P_id_F . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_record_update x45 x44 x43 x42. 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_gen_modtageq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_gen_modtageq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_undefined_bounded : 0 <= P_id_undefined . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nocalls_bounded : 0 <= P_id_nocalls . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_remove_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_resource_bounded : 0 <= P_id_resource . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case6 x45 x44 x43 x42. 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 x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_if x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_bounded : forall x42, (0 <= x42) ->0 <= P_id_pops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ok_bounded : 0 <= P_id_ok . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case5 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_tuple x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_member x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x42, (0 <= x42) ->0 <= P_id_s x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_delete x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_T_bounded : 0 <= P_id_T . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case9 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_extract x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excl_bounded : 0 <= P_id_excl . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case2 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nil_bounded : 0 <= P_id_nil . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqc x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case0 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_request_bounded : 0 <= P_id_request . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_available x43 x42. 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 x42, (0 <= x42) ->0 <= P_id_not x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_pushs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_mcrlrecord_bounded : 0 <= P_id_mcrlrecord . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_cons x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (0 <= x42) -> (0 <= x43) -> (0 <= x44) -> (0 <= x45) -> (0 <= x46) ->(0 <= x47) ->0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case1 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_element x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_adduniq x43 x42. 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 x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_and x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_empty_bounded : 0 <= P_id_empty . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_updates x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_lock_bounded : 0 <= P_id_lock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excllock_bounded : 0 <= P_id_excllock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_bounded : forall x42, (0 <= x42) ->0 <= P_id_pid x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_calls x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_subtract x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tag_bounded : 0 <= P_id_tag . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_equal x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_bounded : forall x42, (0 <= x42) ->0 <= P_id_tops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pending_bounded : 0 <= P_id_pending . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_andt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_bounded : forall x42, (0 <= x42) ->0 <= P_id_tuplenil x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_append x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case8 x45 x44 x43 x42. 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_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43::x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43:: x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43::x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44:: x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44:: x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_GEN_MODTAGEQ (x42:Z) (x43:Z) := 0. Definition P_id_ELEMENT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_eq (x42:Z) (x43:Z) := 0. Definition P_id_CASE9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_REMOVE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_record_new (x42:Z) := 0. Definition P_id_CASE6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_if (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_MEMBER (x42:Z) (x43:Z) := 0. Definition P_id_CASE5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2* x45. Definition P_id_RECORD_UPDATE (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_Marked_not (x42:Z) := 0. Definition P_id_ISTOPS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_ADD_PENDING (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_or (x42:Z) (x43:Z) := 0. Definition P_id_DELETE (x42:Z) (x43:Z) := 0. Definition P_id_CASE0 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_imp (x42:Z) (x43:Z) := 0. Definition P_id_EQT (x42:Z) (x43:Z) := 0. Definition P_id_PUSHS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_RELEASE_LOCK (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_OBTAINABLES (x42:Z) (x43:Z) := 0. Definition P_id_RECORD_UPDATES (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_adduniq (x42:Z) (x43:Z) := 0. Definition P_id_EQS (x42:Z) (x43:Z) := 0. Definition P_id_SUBTRACT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_gen_tag (x42:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLES (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_CLAIM_LOCK (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_case4 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_APPEND (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLE (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_pops (x42:Z) := 0. Definition P_id_EQC (x42:Z) (x43:Z) := 0. Definition P_id_CASE1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_CASE8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_EXTRACT (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_map_add_pending (x42:Z) (x43:Z) (x44: Z) := 0. Definition P_id_AND (x42:Z) (x43:Z) := 0. Definition P_id_Marked_tops (x42:Z) := 0. Definition P_id_CASE2 (x42:Z) (x43:Z) (x44:Z) := 0. Lemma P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_not_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_not x43 <= P_id_Marked_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_DELETE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_pops x43 <= P_id_Marked_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQC_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_tops x43 <= P_id_Marked_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44:: x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44:: x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44:: x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46:: x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44:: x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44:: x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_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_22_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_22_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_22_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_22_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_22_large := WF_DP_R_xml_0_scc_22_large.wf. Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_22. Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_22_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx';apply DP_R_xml_0_scc_22_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_22_large_in_le;econstructor eassumption])). apply wf_DP_R_xml_0_scc_22_large. 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 x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_non_scc_23 (algebra.Alg.Term algebra.F.id_record_extract (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil)) (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::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; (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 x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_non_scc_24 (algebra.Alg.Term algebra.F.id_gen_modtageq (x22::(algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_excl nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::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; (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 := (* *) | DP_R_xml_0_non_scc_25_0 : forall x44 x42 x22 x21 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x21 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_true nil) x42) -> DP_R_xml_0_non_scc_25 (algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil)) (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43::x42::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 := (* *) | DP_R_xml_0_non_scc_26_0 : forall x44 x42 x22 x21 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x21 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_true nil) x42) -> DP_R_xml_0_non_scc_26 (algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil))::(algebra.Alg.Term algebra.F.id_cons (x22::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_26 : forall x y, (DP_R_xml_0_non_scc_26 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_12; 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_27 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_27_0 : forall x44 x42 x22 x21 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x21 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_non_scc_27 (algebra.Alg.Term algebra.F.id_record_extract (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_resource nil)::nil)) (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_27 : forall x y, (DP_R_xml_0_non_scc_27 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_28 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_28_0 : forall x44 x42 x22 x21 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x21 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_non_scc_28 (algebra.Alg.Term algebra.F.id_member ((algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_resource nil)::nil))::x21::nil)) (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::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_non_scc_6; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_non_scc_29 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_29_0 : forall x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_non_scc_29 (algebra.Alg.Term algebra.F.id_record_extract (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::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; (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_30 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_30_0 : forall x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_non_scc_30 (algebra.Alg.Term algebra.F.id_subtract ((algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil))::(algebra.Alg.Term algebra.F.id_cons (x22::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_30 : forall x y, (DP_R_xml_0_non_scc_30 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_11; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_10; 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_31 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_31_0 : forall x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_non_scc_31 (algebra.Alg.Term algebra.F.id_record_extract (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_pending nil)::nil)) (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::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; (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_32 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_32_0 : forall x20 x44 x42 x22 x21 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x21 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_scc_32 (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x20::x21:: x22::nil)) (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) . Module WF_DP_R_xml_0_scc_32. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_or (x42:Z) (x43:Z) := 2. Definition P_id_gen_tag (x42:Z) := 2* x42. Definition P_id_record_new (x42:Z) := 0. Definition P_id_a := 0. Definition P_id_locker2_release_lock (x42:Z) (x43:Z) := 2 + 2* x42 + 3* x43. Definition P_id_eqt (x42:Z) (x43:Z) := 0. Definition P_id_istops (x42:Z) (x43:Z) := 2* x42. Definition P_id_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_release := 0. Definition P_id_locker2_obtainable (x42:Z) (x43:Z) := 0. Definition P_id_imp (x42:Z) (x43:Z) := 1* x43. Definition P_id_stack (x42:Z) (x43:Z) := 1* x42 + 2* x43. Definition P_id_locker2_map_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_locker := 0. Definition P_id_case4 (x42:Z) (x43:Z) (x44:Z) := 2. Definition P_id_int (x42:Z) := 0. Definition P_id_push (x42:Z) (x43:Z) (x44:Z) := 3* x42 + 2* x43 + 2* x44. Definition P_id_locker2_add_pending (x42:Z) (x43:Z) (x44:Z) := 3* x42 + 1* x43 + 1* x44. Definition P_id_true := 2. Definition P_id_locker2_check_availables (x42:Z) (x43:Z) := 2 + 1* x43. Definition P_id_F := 0. Definition P_id_eqs (x42:Z) (x43:Z) := 0. Definition P_id_record_update (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x42 + 1* x44 + 2* x45. Definition P_id_false := 0. Definition P_id_gen_modtageq (x42:Z) (x43:Z) := 1* x42 + 2* x43. Definition P_id_undefined := 0. Definition P_id_nocalls := 0. Definition P_id_locker2_remove_pending (x42:Z) (x43:Z) := 2 + 3* x42 + 3* x43. Definition P_id_resource := 2. Definition P_id_case6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2 + 2* x42 + 2* x45. Definition P_id_if (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 1* x44. Definition P_id_pops (x42:Z) := 1* x42. Definition P_id_locker2_map_claim_lock (x42:Z) (x43:Z) (x44:Z) := 2 + 3* x42. Definition P_id_ok := 0. Definition P_id_case5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2 + 2* x43 + 3* x44. Definition P_id_tuple (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_member (x42:Z) (x43:Z) := 0. Definition P_id_s (x42:Z) := 0. Definition P_id_delete (x42:Z) (x43:Z) := 1* x43. Definition P_id_T := 0. Definition P_id_case9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2* x45. Definition P_id_record_extract (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 2* x44. Definition P_id_excl := 0. Definition P_id_case2 (x42:Z) (x43:Z) (x44:Z) := 1 + 1* x43. Definition P_id_nil := 0. Definition P_id_eqc (x42:Z) (x43:Z) := 0. Definition P_id_case0 (x42:Z) (x43:Z) (x44:Z) := 1 + 1* x43 + 2* x44. Definition P_id_request := 0. Definition P_id_locker2_check_available (x42:Z) (x43:Z) := 1 + 1* x43. Definition P_id_not (x42:Z) := 1. Definition P_id_pushs (x42:Z) (x43:Z) := 1 + 1* x42 + 2* x43. Definition P_id_locker2_promote_pending (x42:Z) (x43:Z) := 1 + 3* x42. Definition P_id_mcrlrecord := 0. Definition P_id_locker2_obtainables (x42:Z) (x43:Z) := 2 + 2* x42. Definition P_id_cons (x42:Z) (x43:Z) := 1 + 2* x42 + 2* x43. Definition P_id_push1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 1* x43 + 2* x44 + 1* x46. Definition P_id_case1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 3* x44 + 2* x45. Definition P_id_element (x42:Z) (x43:Z) := 1* x43. Definition P_id_locker2_adduniq (x42:Z) (x43:Z) := 1* x43. Definition P_id_and (x42:Z) (x43:Z) := 2* x42 + 2* x43. Definition P_id_empty := 0. Definition P_id_record_updates (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_lock := 0. Definition P_id_excllock := 0. Definition P_id_pid (x42:Z) := 0. Definition P_id_calls (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_subtract (x42:Z) (x43:Z) := 1* x42. Definition P_id_tag := 0. Definition P_id_equal (x42:Z) (x43:Z) := 0. Definition P_id_eq (x42:Z) (x43:Z) := 1. Definition P_id_tops (x42:Z) := 1* x42. Definition P_id_locker2_claim_lock (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_pending := 0. Definition P_id_andt (x42:Z) (x43:Z) := 0. Definition P_id_tuplenil (x42:Z) := 1* x42. Definition P_id_append (x42:Z) (x43:Z) := 1* x42. Definition P_id_0 := 0. Definition P_id_case8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 2* x42 + 2* x43. Lemma P_id_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_record_new x43 <= P_id_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_istops x43 x45 <= P_id_istops x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_member x43 x45 <= P_id_member x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_delete x43 x45 <= P_id_delete x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_append x43 x45 <= P_id_append x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_or_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_or x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_bounded : forall x42, (0 <= x42) ->0 <= P_id_gen_tag x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_bounded : forall x42, (0 <= x42) ->0 <= P_id_record_new x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_release_lock x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_istops x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_add_pending x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_release_bounded : 0 <= P_id_release . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainable x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_imp x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_stack x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_map_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker_bounded : 0 <= P_id_locker . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case4 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_bounded : forall x42, (0 <= x42) ->0 <= P_id_int x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_push x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_add_pending x44 x43 x42. 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_locker2_check_availables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_availables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_F_bounded : 0 <= P_id_F . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_record_update x45 x44 x43 x42. 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_gen_modtageq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_gen_modtageq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_undefined_bounded : 0 <= P_id_undefined . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nocalls_bounded : 0 <= P_id_nocalls . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_remove_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_resource_bounded : 0 <= P_id_resource . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case6 x45 x44 x43 x42. 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 x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_if x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_bounded : forall x42, (0 <= x42) ->0 <= P_id_pops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ok_bounded : 0 <= P_id_ok . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case5 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_tuple x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_member x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x42, (0 <= x42) ->0 <= P_id_s x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_delete x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_T_bounded : 0 <= P_id_T . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case9 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_extract x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excl_bounded : 0 <= P_id_excl . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case2 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nil_bounded : 0 <= P_id_nil . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqc x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case0 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_request_bounded : 0 <= P_id_request . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_available x43 x42. 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 x42, (0 <= x42) ->0 <= P_id_not x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_pushs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_mcrlrecord_bounded : 0 <= P_id_mcrlrecord . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_cons x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (0 <= x42) -> (0 <= x43) -> (0 <= x44) -> (0 <= x45) -> (0 <= x46) ->(0 <= x47) ->0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case1 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_element x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_adduniq x43 x42. 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 x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_and x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_empty_bounded : 0 <= P_id_empty . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_updates x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_lock_bounded : 0 <= P_id_lock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excllock_bounded : 0 <= P_id_excllock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_bounded : forall x42, (0 <= x42) ->0 <= P_id_pid x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_calls x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_subtract x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tag_bounded : 0 <= P_id_tag . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_equal x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_bounded : forall x42, (0 <= x42) ->0 <= P_id_tops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pending_bounded : 0 <= P_id_pending . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_andt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_bounded : forall x42, (0 <= x42) ->0 <= P_id_tuplenil x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_append x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case8 x45 x44 x43 x42. 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_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43::x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43:: x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43::x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44:: x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44:: x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_GEN_MODTAGEQ (x42:Z) (x43:Z) := 0. Definition P_id_ELEMENT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_eq (x42:Z) (x43:Z) := 0. Definition P_id_CASE9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_REMOVE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_record_new (x42:Z) := 0. Definition P_id_CASE6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_if (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_MEMBER (x42:Z) (x43:Z) := 0. Definition P_id_CASE5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_UPDATE (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_Marked_not (x42:Z) := 0. Definition P_id_ISTOPS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_ADD_PENDING (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_or (x42:Z) (x43:Z) := 0. Definition P_id_DELETE (x42:Z) (x43:Z) := 0. Definition P_id_CASE0 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_imp (x42:Z) (x43:Z) := 0. Definition P_id_EQT (x42:Z) (x43:Z) := 0. Definition P_id_PUSHS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_RELEASE_LOCK (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_OBTAINABLES (x42:Z) (x43:Z) := 0. Definition P_id_RECORD_UPDATES (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_adduniq (x42:Z) (x43:Z) := 0. Definition P_id_EQS (x42:Z) (x43:Z) := 0. Definition P_id_SUBTRACT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_gen_tag (x42:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLES (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_CLAIM_LOCK (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_Marked_case4 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_APPEND (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLE (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_pops (x42:Z) := 0. Definition P_id_EQC (x42:Z) (x43:Z) := 0. Definition P_id_CASE1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_CASE8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_EXTRACT (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_map_add_pending (x42:Z) (x43:Z) (x44: Z) := 0. Definition P_id_AND (x42:Z) (x43:Z) := 0. Definition P_id_Marked_tops (x42:Z) := 0. Definition P_id_CASE2 (x42:Z) (x43:Z) (x44:Z) := 0. Lemma P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_not_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_not x43 <= P_id_Marked_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_DELETE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_pops x43 <= P_id_Marked_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQC_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_tops x43 <= P_id_Marked_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44:: x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44:: x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44:: x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46:: x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44:: x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44:: x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_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_32. 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_32. Definition wf_DP_R_xml_0_scc_32 := WF_DP_R_xml_0_scc_32.wf. Lemma acc_DP_R_xml_0_scc_32 : forall x y, (DP_R_xml_0_scc_32 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_32). 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_32. Qed. Inductive DP_R_xml_0_non_scc_33 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_33_0 : forall x16 x44 x42 x14 x17 x13 x43 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x14 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x15 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple (x16::(algebra.Alg.Term algebra.F.id_tuplenil (x13::nil))::nil))::x17::nil)) x42) -> DP_R_xml_0_non_scc_33 (algebra.Alg.Term algebra.F.id_record_update (x14::x15::x16::x13::nil)) (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_33 : forall x y, (DP_R_xml_0_non_scc_33 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_34 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_34_0 : forall x16 x44 x42 x14 x17 x13 x43 x15, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x14 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x15 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple (x16::(algebra.Alg.Term algebra.F.id_tuplenil (x13::nil))::nil))::x17::nil)) x42) -> DP_R_xml_0_scc_34 (algebra.Alg.Term algebra.F.id_record_updates ((algebra.Alg.Term algebra.F.id_record_update (x14::x15::x16::x13::nil))::x15::x17::nil)) (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) . Module WF_DP_R_xml_0_scc_34. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_or (x42:Z) (x43:Z) := 1. Definition P_id_gen_tag (x42:Z) := 2* x42. Definition P_id_record_new (x42:Z) := 0. Definition P_id_a := 0. Definition P_id_locker2_release_lock (x42:Z) (x43:Z) := 2 + 3* x42 + 2* x43. Definition P_id_eqt (x42:Z) (x43:Z) := 0. Definition P_id_istops (x42:Z) (x43:Z) := 3* x42 + 2* x43. Definition P_id_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_release := 0. Definition P_id_locker2_obtainable (x42:Z) (x43:Z) := 0. Definition P_id_imp (x42:Z) (x43:Z) := 2 + 1* x43. Definition P_id_stack (x42:Z) (x43:Z) := 1 + 1* x42 + 2* x43. Definition P_id_locker2_map_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_locker := 0. Definition P_id_case4 (x42:Z) (x43:Z) (x44:Z) := 1. Definition P_id_int (x42:Z) := 1* x42. Definition P_id_push (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_locker2_add_pending (x42:Z) (x43:Z) (x44:Z) := 2 + 3* x42 + 2* x43 + 3* x44. Definition P_id_true := 0. Definition P_id_locker2_check_availables (x42:Z) (x43:Z) := 2 + 3* x42 + 1* x43. Definition P_id_F := 0. Definition P_id_eqs (x42:Z) (x43:Z) := 0. Definition P_id_record_update (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1* x42 + 2* x44 + 2* x45. Definition P_id_false := 1. Definition P_id_gen_modtageq (x42:Z) (x43:Z) := 3 + 1* x42. Definition P_id_undefined := 0. Definition P_id_nocalls := 0. Definition P_id_locker2_remove_pending (x42:Z) (x43:Z) := 3 + 3* x42 + 3* x43. Definition P_id_resource := 0. Definition P_id_case6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 3 + 3* x42 + 1* x43 + 1* x45. Definition P_id_if (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 1* x44. Definition P_id_pops (x42:Z) := 1* x42. Definition P_id_locker2_map_claim_lock (x42:Z) (x43:Z) (x44:Z) := 2 + 3* x42. Definition P_id_ok := 0. Definition P_id_case5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_tuple (x42:Z) (x43:Z) := 2* x42 + 1* x43. Definition P_id_member (x42:Z) (x43:Z) := 2 + 3* x43. Definition P_id_s (x42:Z) := 3* x42. Definition P_id_delete (x42:Z) (x43:Z) := 1* x43. Definition P_id_T := 0. Definition P_id_case9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2 + 3* x42 + 2* x43 + 3* x45. Definition P_id_record_extract (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_excl := 0. Definition P_id_case2 (x42:Z) (x43:Z) (x44:Z) := 2 + 1* x42 + 2* x43. Definition P_id_nil := 0. Definition P_id_eqc (x42:Z) (x43:Z) := 0. Definition P_id_case0 (x42:Z) (x43:Z) (x44:Z) := 2 + 1* x43 + 2* x44. Definition P_id_request := 0. Definition P_id_locker2_check_available (x42:Z) (x43:Z) := 2 + 3* x43. Definition P_id_not (x42:Z) := 3. Definition P_id_pushs (x42:Z) (x43:Z) := 1 + 2* x42 + 3* x43. Definition P_id_locker2_promote_pending (x42:Z) (x43:Z) := 2 + 3* x42. Definition P_id_mcrlrecord := 0. Definition P_id_locker2_obtainables (x42:Z) (x43:Z) := 0. Definition P_id_cons (x42:Z) (x43:Z) := 2 + 2* x42 + 2* x43. Definition P_id_push1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_case1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2 + 3* x42 + 1* x43 + 3* x44. Definition P_id_element (x42:Z) (x43:Z) := 1* x42 + 2* x43. Definition P_id_locker2_adduniq (x42:Z) (x43:Z) := 1* x43. Definition P_id_and (x42:Z) (x43:Z) := 2* x42 + 2* x43. Definition P_id_empty := 0. Definition P_id_record_updates (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_lock := 0. Definition P_id_excllock := 0. Definition P_id_pid (x42:Z) := 0. Definition P_id_calls (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_subtract (x42:Z) (x43:Z) := 1* x42. Definition P_id_tag := 0. Definition P_id_equal (x42:Z) (x43:Z) := 2. Definition P_id_eq (x42:Z) (x43:Z) := 0. Definition P_id_tops (x42:Z) := 1* x42. Definition P_id_locker2_claim_lock (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_pending := 0. Definition P_id_andt (x42:Z) (x43:Z) := 0. Definition P_id_tuplenil (x42:Z) := 1* x42. Definition P_id_append (x42:Z) (x43:Z) := 1* x42. Definition P_id_0 := 0. Definition P_id_case8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2 + 2* x42 + 2* x43. Lemma P_id_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_record_new x43 <= P_id_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_istops x43 x45 <= P_id_istops x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_member x43 x45 <= P_id_member x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_delete x43 x45 <= P_id_delete x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_append x43 x45 <= P_id_append x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_or_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_or x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_bounded : forall x42, (0 <= x42) ->0 <= P_id_gen_tag x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_bounded : forall x42, (0 <= x42) ->0 <= P_id_record_new x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_release_lock x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_istops x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_add_pending x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_release_bounded : 0 <= P_id_release . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainable x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_imp x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_stack x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_map_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker_bounded : 0 <= P_id_locker . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case4 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_bounded : forall x42, (0 <= x42) ->0 <= P_id_int x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_push x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_add_pending x44 x43 x42. 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_locker2_check_availables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_availables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_F_bounded : 0 <= P_id_F . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_record_update x45 x44 x43 x42. 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_gen_modtageq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_gen_modtageq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_undefined_bounded : 0 <= P_id_undefined . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nocalls_bounded : 0 <= P_id_nocalls . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_remove_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_resource_bounded : 0 <= P_id_resource . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case6 x45 x44 x43 x42. 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 x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_if x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_bounded : forall x42, (0 <= x42) ->0 <= P_id_pops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ok_bounded : 0 <= P_id_ok . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case5 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_tuple x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_member x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x42, (0 <= x42) ->0 <= P_id_s x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_delete x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_T_bounded : 0 <= P_id_T . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case9 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_extract x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excl_bounded : 0 <= P_id_excl . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case2 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nil_bounded : 0 <= P_id_nil . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqc x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case0 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_request_bounded : 0 <= P_id_request . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_available x43 x42. 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 x42, (0 <= x42) ->0 <= P_id_not x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_pushs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_mcrlrecord_bounded : 0 <= P_id_mcrlrecord . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_cons x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (0 <= x42) -> (0 <= x43) -> (0 <= x44) -> (0 <= x45) -> (0 <= x46) ->(0 <= x47) ->0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case1 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_element x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_adduniq x43 x42. 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 x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_and x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_empty_bounded : 0 <= P_id_empty . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_updates x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_lock_bounded : 0 <= P_id_lock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excllock_bounded : 0 <= P_id_excllock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_bounded : forall x42, (0 <= x42) ->0 <= P_id_pid x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_calls x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_subtract x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tag_bounded : 0 <= P_id_tag . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_equal x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_bounded : forall x42, (0 <= x42) ->0 <= P_id_tops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pending_bounded : 0 <= P_id_pending . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_andt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_bounded : forall x42, (0 <= x42) ->0 <= P_id_tuplenil x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_append x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case8 x45 x44 x43 x42. 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_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43::x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43:: x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43::x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44:: x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44:: x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_GEN_MODTAGEQ (x42:Z) (x43:Z) := 0. Definition P_id_ELEMENT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_eq (x42:Z) (x43:Z) := 0. Definition P_id_CASE9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_REMOVE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_record_new (x42:Z) := 0. Definition P_id_CASE6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_if (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_MEMBER (x42:Z) (x43:Z) := 0. Definition P_id_CASE5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_UPDATE (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_Marked_not (x42:Z) := 0. Definition P_id_ISTOPS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_ADD_PENDING (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_or (x42:Z) (x43:Z) := 0. Definition P_id_DELETE (x42:Z) (x43:Z) := 0. Definition P_id_CASE0 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_imp (x42:Z) (x43:Z) := 0. Definition P_id_EQT (x42:Z) (x43:Z) := 0. Definition P_id_PUSHS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_RELEASE_LOCK (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_OBTAINABLES (x42:Z) (x43:Z) := 0. Definition P_id_RECORD_UPDATES (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_Marked_locker2_adduniq (x42:Z) (x43:Z) := 0. Definition P_id_EQS (x42:Z) (x43:Z) := 0. Definition P_id_SUBTRACT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_gen_tag (x42:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLES (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_CLAIM_LOCK (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_case4 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_APPEND (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLE (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_pops (x42:Z) := 0. Definition P_id_EQC (x42:Z) (x43:Z) := 0. Definition P_id_CASE1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_CASE8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_EXTRACT (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_map_add_pending (x42:Z) (x43:Z) (x44: Z) := 0. Definition P_id_AND (x42:Z) (x43:Z) := 0. Definition P_id_Marked_tops (x42:Z) := 0. Definition P_id_CASE2 (x42:Z) (x43:Z) (x44:Z) := 0. Lemma P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_not_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_not x43 <= P_id_Marked_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_DELETE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_pops x43 <= P_id_Marked_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQC_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_tops x43 <= P_id_Marked_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44:: x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44:: x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44:: x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46:: x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44:: x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44:: x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_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_34. 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_34. Definition wf_DP_R_xml_0_scc_34 := WF_DP_R_xml_0_scc_34.wf. Lemma acc_DP_R_xml_0_scc_34 : forall x y, (DP_R_xml_0_scc_34 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_34). 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_33; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). apply wf_DP_R_xml_0_scc_34. Qed. Inductive DP_R_xml_0_non_scc_35 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_35_0 : forall x44 x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_true nil) x42) -> DP_R_xml_0_non_scc_35 (algebra.Alg.Term algebra.F.id_record_updates (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_excllock nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil))::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_case2 (x44::x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_35 : forall x y, (DP_R_xml_0_non_scc_35 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_34; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_33; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). Qed. Inductive DP_R_xml_0_non_scc_36 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_36_0 : forall x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_non_scc_36 (algebra.Alg.Term algebra.F.id_case2 (x22:: x19::(algebra.Alg.Term algebra.F.id_gen_modtageq (x22:: (algebra.Alg.Term algebra.F.id_record_extract (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_excl nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_36 : forall x y, (DP_R_xml_0_non_scc_36 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_non_scc_35; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_non_scc_37 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_37_0 : forall x44 x42 x22 x21 x45 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x45) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x21 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_true nil) x42) -> DP_R_xml_0_non_scc_37 (algebra.Alg.Term algebra.F.id_record_updates (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Term algebra.F.id_append ((algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil))::(algebra.Alg.Term algebra.F.id_cons (x22::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil))::nil))::nil)):: (algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_37 : forall x y, (DP_R_xml_0_non_scc_37 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_34; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_33; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). Qed. Inductive DP_R_xml_0_non_scc_38 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_38_0 : forall x44 x42 x22 x21 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x21 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_non_scc_38 (algebra.Alg.Term algebra.F.id_case1 (x22:: x21::x19::(algebra.Alg.Term algebra.F.id_member ((algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_resource nil)::nil))::x21::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_38 : forall x y, (DP_R_xml_0_non_scc_38 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_non_scc_37; 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_39 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_39_0 : forall x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_non_scc_39 (algebra.Alg.Term algebra.F.id_record_updates (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_pending nil):: (algebra.Alg.Term algebra.F.id_tuplenil ((algebra.Alg.Term algebra.F.id_subtract ((algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil))::(algebra.Alg.Term algebra.F.id_cons (x22::(algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil))::nil))::nil)):: (algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_39 : forall x y, (DP_R_xml_0_non_scc_39 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_34; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_33; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). Qed. Inductive DP_R_xml_0_non_scc_40 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_40_0 : forall x44 x42 x22 x19 x43 x23, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x22::x23::nil)) x42) -> DP_R_xml_0_non_scc_40 (algebra.Alg.Term algebra.F.id_record_updates (x19::(algebra.Alg.Term algebra.F.id_lock nil)::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_excl nil):: (algebra.Alg.Term algebra.F.id_tuplenil (x22::nil))::nil))::(algebra.Alg.Term algebra.F.id_cons ((algebra.Alg.Term algebra.F.id_tuple ((algebra.Alg.Term algebra.F.id_pending nil)::(algebra.Alg.Term algebra.F.id_tuplenil (x23::nil))::nil)):: (algebra.Alg.Term algebra.F.id_nil nil)::nil))::nil))::nil)) (algebra.Alg.Term algebra.F.id_case0 (x44::x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_40 : forall x y, (DP_R_xml_0_non_scc_40 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_34; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_33; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). Qed. Inductive DP_R_xml_0_non_scc_41 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_41_0 : forall x42 x22 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x19 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x22 x42) -> DP_R_xml_0_non_scc_41 (algebra.Alg.Term algebra.F.id_case0 (x22:: x19::(algebra.Alg.Term algebra.F.id_record_extract (x19:: (algebra.Alg.Term algebra.F.id_lock nil):: (algebra.Alg.Term algebra.F.id_pending nil)::nil))::nil)) (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_41 : forall x y, (DP_R_xml_0_non_scc_41 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_non_scc_40; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))). Qed. Inductive DP_R_xml_0_non_scc_42 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_42_0 : forall x20 x18 x42 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x18 x42) -> DP_R_xml_0_non_scc_42 (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x19:: x18::nil)) (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_42 : forall x y, (DP_R_xml_0_non_scc_42 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_non_scc_41; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_31; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' )))). Qed. Inductive DP_R_xml_0_scc_43 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_43_0 : forall x20 x18 x42 x19 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x19::x20::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x18 x42) -> DP_R_xml_0_scc_43 (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x20:: x18::nil)) (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) . Module WF_DP_R_xml_0_scc_43. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_or (x42:Z) (x43:Z) := 0. Definition P_id_gen_tag (x42:Z) := 2* x42. Definition P_id_record_new (x42:Z) := 3 + 1* x42. Definition P_id_a := 0. Definition P_id_locker2_release_lock (x42:Z) (x43:Z) := 3 + 1* x42 + 3* x43. Definition P_id_eqt (x42:Z) (x43:Z) := 0. Definition P_id_istops (x42:Z) (x43:Z) := 1* x42. Definition P_id_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_release := 0. Definition P_id_locker2_obtainable (x42:Z) (x43:Z) := 0. Definition P_id_imp (x42:Z) (x43:Z) := 1* x43. Definition P_id_stack (x42:Z) (x43:Z) := 2* x42 + 1* x43. Definition P_id_locker2_map_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_locker := 0. Definition P_id_case4 (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_int (x42:Z) := 0. Definition P_id_push (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 2* x43. Definition P_id_locker2_add_pending (x42:Z) (x43:Z) (x44:Z) := 1 + 2* x42 + 2* x43 + 3* x44. Definition P_id_true := 0. Definition P_id_locker2_check_availables (x42:Z) (x43:Z) := 2 + 2* x42. Definition P_id_F := 0. Definition P_id_eqs (x42:Z) (x43:Z) := 0. Definition P_id_record_update (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x42 + 1* x45. Definition P_id_false := 0. Definition P_id_gen_modtageq (x42:Z) (x43:Z) := 1 + 3* x43. Definition P_id_undefined := 0. Definition P_id_nocalls := 0. Definition P_id_locker2_remove_pending (x42:Z) (x43:Z) := 1 + 3* x42. Definition P_id_resource := 1. Definition P_id_case6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_if (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 1* x44. Definition P_id_pops (x42:Z) := 1* x42. Definition P_id_locker2_map_claim_lock (x42:Z) (x43:Z) (x44:Z) := 2 + 2* x42. Definition P_id_ok := 0. Definition P_id_case5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x43. Definition P_id_tuple (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_member (x42:Z) (x43:Z) := 3* x43. Definition P_id_s (x42:Z) := 0. Definition P_id_delete (x42:Z) (x43:Z) := 1* x43. Definition P_id_T := 0. Definition P_id_case9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 3* x42 + 3* x43. Definition P_id_record_extract (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_excl := 0. Definition P_id_case2 (x42:Z) (x43:Z) (x44:Z) := 2 + 1* x43. Definition P_id_nil := 0. Definition P_id_eqc (x42:Z) (x43:Z) := 0. Definition P_id_case0 (x42:Z) (x43:Z) (x44:Z) := 2 + 2* x43 + 1* x44. Definition P_id_request := 0. Definition P_id_locker2_check_available (x42:Z) (x43:Z) := 0. Definition P_id_not (x42:Z) := 1. Definition P_id_pushs (x42:Z) (x43:Z) := 2 + 2* x42 + 1* x43. Definition P_id_locker2_promote_pending (x42:Z) (x43:Z) := 2 + 3* x42. Definition P_id_mcrlrecord := 0. Definition P_id_locker2_obtainables (x42:Z) (x43:Z) := 1* x42. Definition P_id_cons (x42:Z) (x43:Z) := 1 + 1* x42 + 1* x43. Definition P_id_push1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 1* x42. Definition P_id_case1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 3* x42 + 1* x43 + 2* x44. Definition P_id_element (x42:Z) (x43:Z) := 1* x43. Definition P_id_locker2_adduniq (x42:Z) (x43:Z) := 1* x43. Definition P_id_and (x42:Z) (x43:Z) := 1* x42 + 2* x43. Definition P_id_empty := 0. Definition P_id_record_updates (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_lock := 1. Definition P_id_excllock := 0. Definition P_id_pid (x42:Z) := 0. Definition P_id_calls (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_subtract (x42:Z) (x43:Z) := 1* x42. Definition P_id_tag := 0. Definition P_id_equal (x42:Z) (x43:Z) := 0. Definition P_id_eq (x42:Z) (x43:Z) := 0. Definition P_id_tops (x42:Z) := 1* x42. Definition P_id_locker2_claim_lock (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_pending := 0. Definition P_id_andt (x42:Z) (x43:Z) := 0. Definition P_id_tuplenil (x42:Z) := 1* x42. Definition P_id_append (x42:Z) (x43:Z) := 1* x42. Definition P_id_0 := 0. Definition P_id_case8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x42 + 1* x43. Lemma P_id_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_record_new x43 <= P_id_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_istops x43 x45 <= P_id_istops x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_member x43 x45 <= P_id_member x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_delete x43 x45 <= P_id_delete x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_append x43 x45 <= P_id_append x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_or_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_or x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_bounded : forall x42, (0 <= x42) ->0 <= P_id_gen_tag x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_bounded : forall x42, (0 <= x42) ->0 <= P_id_record_new x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_release_lock x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_istops x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_add_pending x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_release_bounded : 0 <= P_id_release . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainable x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_imp x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_stack x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_map_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker_bounded : 0 <= P_id_locker . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case4 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_bounded : forall x42, (0 <= x42) ->0 <= P_id_int x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_push x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_add_pending x44 x43 x42. 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_locker2_check_availables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_availables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_F_bounded : 0 <= P_id_F . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_record_update x45 x44 x43 x42. 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_gen_modtageq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_gen_modtageq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_undefined_bounded : 0 <= P_id_undefined . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nocalls_bounded : 0 <= P_id_nocalls . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_remove_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_resource_bounded : 0 <= P_id_resource . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case6 x45 x44 x43 x42. 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 x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_if x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_bounded : forall x42, (0 <= x42) ->0 <= P_id_pops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ok_bounded : 0 <= P_id_ok . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case5 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_tuple x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_member x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x42, (0 <= x42) ->0 <= P_id_s x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_delete x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_T_bounded : 0 <= P_id_T . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case9 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_extract x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excl_bounded : 0 <= P_id_excl . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case2 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nil_bounded : 0 <= P_id_nil . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqc x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case0 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_request_bounded : 0 <= P_id_request . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_available x43 x42. 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 x42, (0 <= x42) ->0 <= P_id_not x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_pushs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_mcrlrecord_bounded : 0 <= P_id_mcrlrecord . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_cons x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (0 <= x42) -> (0 <= x43) -> (0 <= x44) -> (0 <= x45) -> (0 <= x46) ->(0 <= x47) ->0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case1 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_element x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_adduniq x43 x42. 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 x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_and x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_empty_bounded : 0 <= P_id_empty . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_updates x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_lock_bounded : 0 <= P_id_lock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excllock_bounded : 0 <= P_id_excllock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_bounded : forall x42, (0 <= x42) ->0 <= P_id_pid x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_calls x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_subtract x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tag_bounded : 0 <= P_id_tag . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_equal x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_bounded : forall x42, (0 <= x42) ->0 <= P_id_tops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pending_bounded : 0 <= P_id_pending . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_andt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_bounded : forall x42, (0 <= x42) ->0 <= P_id_tuplenil x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_append x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case8 x45 x44 x43 x42. 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_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43::x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43:: x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43::x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44:: x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44:: x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_GEN_MODTAGEQ (x42:Z) (x43:Z) := 0. Definition P_id_ELEMENT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_eq (x42:Z) (x43:Z) := 0. Definition P_id_CASE9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_REMOVE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_record_new (x42:Z) := 0. Definition P_id_CASE6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_if (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_MEMBER (x42:Z) (x43:Z) := 0. Definition P_id_CASE5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_UPDATE (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_Marked_not (x42:Z) := 0. Definition P_id_ISTOPS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_ADD_PENDING (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_or (x42:Z) (x43:Z) := 0. Definition P_id_DELETE (x42:Z) (x43:Z) := 0. Definition P_id_CASE0 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_imp (x42:Z) (x43:Z) := 0. Definition P_id_EQT (x42:Z) (x43:Z) := 0. Definition P_id_PUSHS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_RELEASE_LOCK (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_OBTAINABLES (x42:Z) (x43:Z) := 0. Definition P_id_RECORD_UPDATES (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_adduniq (x42:Z) (x43:Z) := 0. Definition P_id_EQS (x42:Z) (x43:Z) := 0. Definition P_id_SUBTRACT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_gen_tag (x42:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLES (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_CLAIM_LOCK (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_case4 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_APPEND (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLE (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_PROMOTE_PENDING (x42:Z) (x43:Z) := 1* x42. Definition P_id_Marked_pops (x42:Z) := 0. Definition P_id_EQC (x42:Z) (x43:Z) := 0. Definition P_id_CASE1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_CASE8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_EXTRACT (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_map_add_pending (x42:Z) (x43:Z) (x44: Z) := 0. Definition P_id_AND (x42:Z) (x43:Z) := 0. Definition P_id_Marked_tops (x42:Z) := 0. Definition P_id_CASE2 (x42:Z) (x43:Z) (x44:Z) := 0. Lemma P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_not_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_not x43 <= P_id_Marked_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_DELETE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_pops x43 <= P_id_Marked_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQC_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_tops x43 <= P_id_Marked_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44:: x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44:: x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44:: x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46:: x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44:: x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44:: x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_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_43. 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_43. Definition wf_DP_R_xml_0_scc_43 := WF_DP_R_xml_0_scc_43.wf. Lemma acc_DP_R_xml_0_scc_43 : forall x y, (DP_R_xml_0_scc_43 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_43). 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_42; 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_43. Qed. Inductive DP_R_xml_0_scc_44 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_44_0 : forall x42 x6 x9 x43 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Term algebra.F.id_s ((algebra.Alg.Term algebra.F.id_s (x7::nil))::nil))::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x9::x6::nil)) x42) -> DP_R_xml_0_scc_44 (algebra.Alg.Term algebra.F.id_element ((algebra.Alg.Term algebra.F.id_int ((algebra.Alg.Term algebra.F.id_s (x7::nil))::nil))::x6::nil)) (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) . Module WF_DP_R_xml_0_scc_44. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_or (x42:Z) (x43:Z) := 0. Definition P_id_gen_tag (x42:Z) := 3 + 3* x42. Definition P_id_record_new (x42:Z) := 0. Definition P_id_a := 2. Definition P_id_locker2_release_lock (x42:Z) (x43:Z) := 2* x42 + 2* x43. Definition P_id_eqt (x42:Z) (x43:Z) := 1* x42. Definition P_id_istops (x42:Z) (x43:Z) := 2* x42. Definition P_id_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_release := 0. Definition P_id_locker2_obtainable (x42:Z) (x43:Z) := 0. Definition P_id_imp (x42:Z) (x43:Z) := 2 + 1* x43. Definition P_id_stack (x42:Z) (x43:Z) := 2* x42 + 1* x43. Definition P_id_locker2_map_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_locker := 2. Definition P_id_case4 (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_int (x42:Z) := 1 + 2* x42. Definition P_id_push (x42:Z) (x43:Z) (x44:Z) := 2* x42 + 3* x43 + 2* x44. Definition P_id_locker2_add_pending (x42:Z) (x43:Z) (x44:Z) := 3* x42 + 1* x43 + 1* x44. Definition P_id_true := 0. Definition P_id_locker2_check_availables (x42:Z) (x43:Z) := 3* x43. Definition P_id_F := 0. Definition P_id_eqs (x42:Z) (x43:Z) := 2* x42. Definition P_id_record_update (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1* x42 + 3* x44 + 2* x45. Definition P_id_false := 0. Definition P_id_gen_modtageq (x42:Z) (x43:Z) := 2. Definition P_id_undefined := 0. Definition P_id_nocalls := 0. Definition P_id_locker2_remove_pending (x42:Z) (x43:Z) := 3* x42 + 1* x43. Definition P_id_resource := 1. Definition P_id_case6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_if (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 1* x44. Definition P_id_pops (x42:Z) := 1* x42. Definition P_id_locker2_map_claim_lock (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_ok := 0. Definition P_id_case5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_tuple (x42:Z) (x43:Z) := 2* x42 + 1* x43. Definition P_id_member (x42:Z) (x43:Z) := 0. Definition P_id_s (x42:Z) := 1 + 2* x42. Definition P_id_delete (x42:Z) (x43:Z) := 1* x43. Definition P_id_T := 0. Definition P_id_case9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_record_extract (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_excl := 0. Definition P_id_case2 (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 2* x43. Definition P_id_nil := 0. Definition P_id_eqc (x42:Z) (x43:Z) := 1 + 2* x42. Definition P_id_case0 (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 2* x44. Definition P_id_request := 0. Definition P_id_locker2_check_available (x42:Z) (x43:Z) := 0. Definition P_id_not (x42:Z) := 1. Definition P_id_pushs (x42:Z) (x43:Z) := 2* x42 + 1* x43. Definition P_id_locker2_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_mcrlrecord := 0. Definition P_id_locker2_obtainables (x42:Z) (x43:Z) := 0. Definition P_id_cons (x42:Z) (x43:Z) := 2* x42 + 1* x43. Definition P_id_push1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 1* x42 + 2* x43 + 3* x44 + 2* x45 + 1* x46. Definition P_id_case1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1* x43 + 3* x44. Definition P_id_element (x42:Z) (x43:Z) := 1* x43. Definition P_id_locker2_adduniq (x42:Z) (x43:Z) := 1* x43. Definition P_id_and (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_empty := 0. Definition P_id_record_updates (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_lock := 0. Definition P_id_excllock := 0. Definition P_id_pid (x42:Z) := 1 + 1* x42. Definition P_id_calls (x42:Z) (x43:Z) (x44:Z) := 2* x42 + 1* x43 + 1* x44. Definition P_id_subtract (x42:Z) (x43:Z) := 1* x42. Definition P_id_tag := 2. Definition P_id_equal (x42:Z) (x43:Z) := 0. Definition P_id_eq (x42:Z) (x43:Z) := 0. Definition P_id_tops (x42:Z) := 1* x42. Definition P_id_locker2_claim_lock (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_pending := 0. Definition P_id_andt (x42:Z) (x43:Z) := 0. Definition P_id_tuplenil (x42:Z) := 1* x42. Definition P_id_append (x42:Z) (x43:Z) := 1* x42. Definition P_id_0 := 0. Definition P_id_case8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1* x42 + 2* x43. Lemma P_id_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_record_new x43 <= P_id_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_istops x43 x45 <= P_id_istops x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_member x43 x45 <= P_id_member x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_delete x43 x45 <= P_id_delete x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_append x43 x45 <= P_id_append x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_or_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_or x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_bounded : forall x42, (0 <= x42) ->0 <= P_id_gen_tag x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_bounded : forall x42, (0 <= x42) ->0 <= P_id_record_new x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_release_lock x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_istops x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_add_pending x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_release_bounded : 0 <= P_id_release . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainable x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_imp x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_stack x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_map_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker_bounded : 0 <= P_id_locker . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case4 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_bounded : forall x42, (0 <= x42) ->0 <= P_id_int x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_push x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_add_pending x44 x43 x42. 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_locker2_check_availables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_availables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_F_bounded : 0 <= P_id_F . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_record_update x45 x44 x43 x42. 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_gen_modtageq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_gen_modtageq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_undefined_bounded : 0 <= P_id_undefined . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nocalls_bounded : 0 <= P_id_nocalls . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_remove_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_resource_bounded : 0 <= P_id_resource . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case6 x45 x44 x43 x42. 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 x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_if x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_bounded : forall x42, (0 <= x42) ->0 <= P_id_pops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ok_bounded : 0 <= P_id_ok . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case5 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_tuple x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_member x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x42, (0 <= x42) ->0 <= P_id_s x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_delete x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_T_bounded : 0 <= P_id_T . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case9 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_extract x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excl_bounded : 0 <= P_id_excl . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case2 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nil_bounded : 0 <= P_id_nil . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqc x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case0 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_request_bounded : 0 <= P_id_request . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_available x43 x42. 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 x42, (0 <= x42) ->0 <= P_id_not x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_pushs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_mcrlrecord_bounded : 0 <= P_id_mcrlrecord . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_cons x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (0 <= x42) -> (0 <= x43) -> (0 <= x44) -> (0 <= x45) -> (0 <= x46) ->(0 <= x47) ->0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case1 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_element x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_adduniq x43 x42. 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 x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_and x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_empty_bounded : 0 <= P_id_empty . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_updates x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_lock_bounded : 0 <= P_id_lock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excllock_bounded : 0 <= P_id_excllock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_bounded : forall x42, (0 <= x42) ->0 <= P_id_pid x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_calls x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_subtract x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tag_bounded : 0 <= P_id_tag . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_equal x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_bounded : forall x42, (0 <= x42) ->0 <= P_id_tops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pending_bounded : 0 <= P_id_pending . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_andt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_bounded : forall x42, (0 <= x42) ->0 <= P_id_tuplenil x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_append x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case8 x45 x44 x43 x42. 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_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43::x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43:: x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43::x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44:: x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44:: x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_GEN_MODTAGEQ (x42:Z) (x43:Z) := 0. Definition P_id_ELEMENT (x42:Z) (x43:Z) := 2* x42. Definition P_id_Marked_eq (x42:Z) (x43:Z) := 0. Definition P_id_CASE9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_REMOVE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_record_new (x42:Z) := 0. Definition P_id_CASE6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_if (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_MEMBER (x42:Z) (x43:Z) := 0. Definition P_id_CASE5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_UPDATE (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_Marked_not (x42:Z) := 0. Definition P_id_ISTOPS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_ADD_PENDING (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_or (x42:Z) (x43:Z) := 0. Definition P_id_DELETE (x42:Z) (x43:Z) := 0. Definition P_id_CASE0 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_imp (x42:Z) (x43:Z) := 0. Definition P_id_EQT (x42:Z) (x43:Z) := 0. Definition P_id_PUSHS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_RELEASE_LOCK (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_OBTAINABLES (x42:Z) (x43:Z) := 0. Definition P_id_RECORD_UPDATES (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_adduniq (x42:Z) (x43:Z) := 0. Definition P_id_EQS (x42:Z) (x43:Z) := 0. Definition P_id_SUBTRACT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_gen_tag (x42:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLES (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_CLAIM_LOCK (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_case4 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_APPEND (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLE (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_pops (x42:Z) := 0. Definition P_id_EQC (x42:Z) (x43:Z) := 0. Definition P_id_CASE1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_CASE8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_EXTRACT (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_map_add_pending (x42:Z) (x43:Z) (x44: Z) := 0. Definition P_id_AND (x42:Z) (x43:Z) := 0. Definition P_id_Marked_tops (x42:Z) := 0. Definition P_id_CASE2 (x42:Z) (x43:Z) (x44:Z) := 0. Lemma P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_not_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_not x43 <= P_id_Marked_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_DELETE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_pops x43 <= P_id_Marked_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQC_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_tops x43 <= P_id_Marked_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44:: x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44:: x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44:: x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46:: x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44:: x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44:: x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_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_44. 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_44. Definition wf_DP_R_xml_0_scc_44 := WF_DP_R_xml_0_scc_44.wf. Lemma acc_DP_R_xml_0_scc_44 : forall x y, (DP_R_xml_0_scc_44 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_44). 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_44. Qed. Inductive DP_R_xml_0_non_scc_45 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_45_0 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x5::x6::nil)) x42) -> DP_R_xml_0_non_scc_45 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt (x8:: x5::nil))::(algebra.Alg.Term algebra.F.id_eqt (x9::x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_45 : forall x y, (DP_R_xml_0_non_scc_45 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec;econstructor (eassumption)||(algebra.Alg_ext.star_refl' )). Qed. Inductive DP_R_xml_0_non_scc_46 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_46_0 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x42) -> DP_R_xml_0_non_scc_46 (algebra.Alg.Term algebra.F.id_and ((algebra.Alg.Term algebra.F.id_eqt (x8:: x5::nil))::(algebra.Alg.Term algebra.F.id_eqt (x9::x6::nil))::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_46 : forall x y, (DP_R_xml_0_non_scc_46 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (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_47 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_47_0 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x42) -> DP_R_xml_0_scc_47 (algebra.Alg.Term algebra.F.id_eqt (x8::x5::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_scc_47_1 : forall x4 x42 x43 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_pid (x7::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_pid (x4::nil)) x42) -> DP_R_xml_0_scc_47 (algebra.Alg.Term algebra.F.id_eqt (x7::x4::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_scc_47_2 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x42) -> DP_R_xml_0_scc_47 (algebra.Alg.Term algebra.F.id_eqt (x9::x6::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_scc_47_3 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x5::x6::nil)) x42) -> DP_R_xml_0_scc_47 (algebra.Alg.Term algebra.F.id_eqt (x8::x5::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_scc_47_4 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x5::x6::nil)) x42) -> DP_R_xml_0_scc_47 (algebra.Alg.Term algebra.F.id_eqt (x9::x6::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_scc_47_5 : forall x8 x42 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuplenil (x8::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuplenil (x5::nil)) x42) -> DP_R_xml_0_scc_47 (algebra.Alg.Term algebra.F.id_eqt (x8::x5::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) . Module WF_DP_R_xml_0_scc_47. Inductive DP_R_xml_0_scc_47_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_47_large_0 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x42) -> DP_R_xml_0_scc_47_large (algebra.Alg.Term algebra.F.id_eqt (x8:: x5::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_scc_47_large_1 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x42) -> DP_R_xml_0_scc_47_large (algebra.Alg.Term algebra.F.id_eqt (x9:: x6::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_scc_47_large_2 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x5::x6::nil)) x42) -> DP_R_xml_0_scc_47_large (algebra.Alg.Term algebra.F.id_eqt (x8:: x5::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_scc_47_large_3 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x5::x6::nil)) x42) -> DP_R_xml_0_scc_47_large (algebra.Alg.Term algebra.F.id_eqt (x9:: x6::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_scc_47_large_4 : forall x8 x42 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuplenil (x8::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuplenil (x5::nil)) x42) -> DP_R_xml_0_scc_47_large (algebra.Alg.Term algebra.F.id_eqt (x8:: x5::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) . Inductive DP_R_xml_0_scc_47_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_47_strict_0 : forall x4 x42 x43 x7, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_pid (x7::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_pid (x4::nil)) x42) -> DP_R_xml_0_scc_47_strict (algebra.Alg.Term algebra.F.id_eqt (x7:: x4::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) . Module WF_DP_R_xml_0_scc_47_large. Inductive DP_R_xml_0_scc_47_large_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_47_large_large_0 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x5::x6::nil)) x42) -> DP_R_xml_0_scc_47_large_large (algebra.Alg.Term algebra.F.id_eqt (x8::x5::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_scc_47_large_large_1 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x5::x6::nil)) x42) -> DP_R_xml_0_scc_47_large_large (algebra.Alg.Term algebra.F.id_eqt (x9::x6::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_scc_47_large_large_2 : forall x8 x42 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuplenil (x8::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuplenil (x5::nil)) x42) -> DP_R_xml_0_scc_47_large_large (algebra.Alg.Term algebra.F.id_eqt (x8::x5::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) . Inductive DP_R_xml_0_scc_47_large_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_47_large_strict_0 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x42) -> DP_R_xml_0_scc_47_large_strict (algebra.Alg.Term algebra.F.id_eqt (x8::x5::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_scc_47_large_strict_1 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_cons (x5::x6::nil)) x42) -> DP_R_xml_0_scc_47_large_strict (algebra.Alg.Term algebra.F.id_eqt (x9::x6::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) . Module WF_DP_R_xml_0_scc_47_large_large. Inductive DP_R_xml_0_scc_47_large_large_large : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_47_large_large_large_0 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x5::x6::nil)) x42) -> DP_R_xml_0_scc_47_large_large_large (algebra.Alg.Term algebra.F.id_eqt (x8:: x5::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) (* *) | DP_R_xml_0_scc_47_large_large_large_1 : forall x8 x42 x6 x9 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x8::x9::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuple (x5::x6::nil)) x42) -> DP_R_xml_0_scc_47_large_large_large (algebra.Alg.Term algebra.F.id_eqt (x9:: x6::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) . Inductive DP_R_xml_0_scc_47_large_large_strict : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_47_large_large_strict_0 : forall x8 x42 x5 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuplenil (x8::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_tuplenil (x5::nil)) x42) -> DP_R_xml_0_scc_47_large_large_strict (algebra.Alg.Term algebra.F.id_eqt (x8:: x5::nil)) (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) . Module WF_DP_R_xml_0_scc_47_large_large_large. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_or (x42:Z) (x43:Z) := 0. Definition P_id_gen_tag (x42:Z) := 1 + 3* x42. Definition P_id_record_new (x42:Z) := 3 + 3* x42. Definition P_id_a := 0. Definition P_id_locker2_release_lock (x42:Z) (x43:Z) := 3 + 3* x42 + 2* x43. Definition P_id_eqt (x42:Z) (x43:Z) := 0. Definition P_id_istops (x42:Z) (x43:Z) := 0. Definition P_id_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_release := 0. Definition P_id_locker2_obtainable (x42:Z) (x43:Z) := 0. Definition P_id_imp (x42:Z) (x43:Z) := 1* x43. Definition P_id_stack (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_locker2_map_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_locker := 0. Definition P_id_case4 (x42:Z) (x43:Z) (x44:Z) := 2. Definition P_id_int (x42:Z) := 0. Definition P_id_push (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 2* x43 + 2* x44. Definition P_id_locker2_add_pending (x42:Z) (x43:Z) (x44:Z) := 3 + 3* x42 + 2* x43 + 3* x44. Definition P_id_true := 0. Definition P_id_locker2_check_availables (x42:Z) (x43:Z) := 2* x42. Definition P_id_F := 0. Definition P_id_eqs (x42:Z) (x43:Z) := 0. Definition P_id_record_update (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 3 + 1* x42 + 1* x44 + 2* x45. Definition P_id_false := 0. Definition P_id_gen_modtageq (x42:Z) (x43:Z) := 2 + 3* x42 + 2* x43. Definition P_id_undefined := 0. Definition P_id_nocalls := 0. Definition P_id_locker2_remove_pending (x42:Z) (x43:Z) := 3 + 3* x42. Definition P_id_resource := 0. Definition P_id_case6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2. Definition P_id_if (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 1* x44. Definition P_id_pops (x42:Z) := 1* x42. Definition P_id_locker2_map_claim_lock (x42:Z) (x43:Z) (x44:Z) := 2* x42. Definition P_id_ok := 0. Definition P_id_case5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1. Definition P_id_tuple (x42:Z) (x43:Z) := 1 + 2* x42 + 1* x43. Definition P_id_member (x42:Z) (x43:Z) := 2* x43. Definition P_id_s (x42:Z) := 0. Definition P_id_delete (x42:Z) (x43:Z) := 1* x43. Definition P_id_T := 0. Definition P_id_case9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2* x42 + 2* x43 + 2* x45. Definition P_id_record_extract (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_excl := 0. Definition P_id_case2 (x42:Z) (x43:Z) (x44:Z) := 3 + 1* x42 + 2* x43. Definition P_id_nil := 0. Definition P_id_eqc (x42:Z) (x43:Z) := 0. Definition P_id_case0 (x42:Z) (x43:Z) (x44:Z) := 2 + 1* x43 + 2* x44. Definition P_id_request := 0. Definition P_id_locker2_check_available (x42:Z) (x43:Z) := 2. Definition P_id_not (x42:Z) := 3. Definition P_id_pushs (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_locker2_promote_pending (x42:Z) (x43:Z) := 2 + 3* x42. Definition P_id_mcrlrecord := 0. Definition P_id_locker2_obtainables (x42:Z) (x43:Z) := 1. Definition P_id_cons (x42:Z) (x43:Z) := 2 + 1* x42 + 1* x43. Definition P_id_push1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47:Z) := 1* x42 + 2* x43 + 2* x45 + 2* x46. Definition P_id_case1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 3 + 3* x42 + 1* x43 + 3* x44. Definition P_id_element (x42:Z) (x43:Z) := 2* x43. Definition P_id_locker2_adduniq (x42:Z) (x43:Z) := 1* x43. Definition P_id_and (x42:Z) (x43:Z) := 2* x42 + 2* x43. Definition P_id_empty := 0. Definition P_id_record_updates (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_lock := 2. Definition P_id_excllock := 0. Definition P_id_pid (x42:Z) := 0. Definition P_id_calls (x42:Z) (x43:Z) (x44:Z) := 2* x43 + 1* x44. Definition P_id_subtract (x42:Z) (x43:Z) := 1* x42. Definition P_id_tag := 0. Definition P_id_equal (x42:Z) (x43:Z) := 1. Definition P_id_eq (x42:Z) (x43:Z) := 0. Definition P_id_tops (x42:Z) := 1* x42. Definition P_id_locker2_claim_lock (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_pending := 0. Definition P_id_andt (x42:Z) (x43:Z) := 1. Definition P_id_tuplenil (x42:Z) := 2* x42. Definition P_id_append (x42:Z) (x43:Z) := 1* x42. Definition P_id_0 := 0. Definition P_id_case8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2 + 1* x42 + 1* x43. Lemma P_id_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_record_new x43 <= P_id_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_istops x43 x45 <= P_id_istops x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_member x43 x45 <= P_id_member x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_delete x43 x45 <= P_id_delete x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_append x43 x45 <= P_id_append x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_or_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_or x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_bounded : forall x42, (0 <= x42) ->0 <= P_id_gen_tag x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_bounded : forall x42, (0 <= x42) ->0 <= P_id_record_new x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_release_lock x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_istops x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->0 <= P_id_locker2_map_add_pending x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_release_bounded : 0 <= P_id_release . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainable x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_imp x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_stack x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_bounded : forall x42 x43, (0 <= x42) -> (0 <= x43) ->0 <= P_id_locker2_map_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker_bounded : 0 <= P_id_locker . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case4 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_bounded : forall x42, (0 <= x42) ->0 <= P_id_int x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_push x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_add_pending x44 x43 x42. 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_locker2_check_availables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_availables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_F_bounded : 0 <= P_id_F . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_record_update x45 x44 x43 x42. 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_gen_modtageq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_gen_modtageq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_undefined_bounded : 0 <= P_id_undefined . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nocalls_bounded : 0 <= P_id_nocalls . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_remove_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_resource_bounded : 0 <= P_id_resource . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_case6 x45 x44 x43 x42. 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 x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_if x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_bounded : forall x42, (0 <= x42) ->0 <= P_id_pops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->0 <= P_id_locker2_map_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ok_bounded : 0 <= P_id_ok . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_case5 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_tuple x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_member x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x42, (0 <= x42) ->0 <= P_id_s x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_delete x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_T_bounded : 0 <= P_id_T . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_case9 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_extract x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excl_bounded : 0 <= P_id_excl . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case2 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nil_bounded : 0 <= P_id_nil . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqc x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case0 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_request_bounded : 0 <= P_id_request . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_available x43 x42. 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 x42, (0 <= x42) ->0 <= P_id_not x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_pushs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_mcrlrecord_bounded : 0 <= P_id_mcrlrecord . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_cons x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (0 <= x42) -> (0 <= x43) -> (0 <= x44) -> (0 <= x45) -> (0 <= x46) ->(0 <= x47) ->0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_case1 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_element x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_adduniq x43 x42. 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 x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_and x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_empty_bounded : 0 <= P_id_empty . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_updates x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_lock_bounded : 0 <= P_id_lock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excllock_bounded : 0 <= P_id_excllock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_bounded : forall x42, (0 <= x42) ->0 <= P_id_pid x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_calls x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_subtract x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tag_bounded : 0 <= P_id_tag . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_equal x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_bounded : forall x42, (0 <= x42) ->0 <= P_id_tops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pending_bounded : 0 <= P_id_pending . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_andt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_bounded : forall x42, (0 <= x42) ->0 <= P_id_tuplenil x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_append x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_case8 x45 x44 x43 x42. 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_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43::x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45:: x44::x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44::x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_GEN_MODTAGEQ (x42:Z) (x43:Z) := 0. Definition P_id_ELEMENT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_eq (x42:Z) (x43:Z) := 0. Definition P_id_CASE9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_REMOVE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_record_new (x42:Z) := 0. Definition P_id_CASE6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_if (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_MEMBER (x42:Z) (x43:Z) := 0. Definition P_id_CASE5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_UPDATE (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_Marked_not (x42:Z) := 0. Definition P_id_ISTOPS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_ADD_PENDING (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_or (x42:Z) (x43:Z) := 0. Definition P_id_DELETE (x42:Z) (x43:Z) := 0. Definition P_id_CASE0 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_imp (x42:Z) (x43:Z) := 0. Definition P_id_EQT (x42:Z) (x43:Z) := 1* x43. Definition P_id_PUSHS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_RELEASE_LOCK (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_OBTAINABLES (x42:Z) (x43:Z) := 0. Definition P_id_RECORD_UPDATES (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_adduniq (x42:Z) (x43:Z) := 0. Definition P_id_EQS (x42:Z) (x43:Z) := 0. Definition P_id_SUBTRACT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_gen_tag (x42:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLES (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_CLAIM_LOCK (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_case4 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47:Z) := 0. Definition P_id_APPEND (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLE (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_pops (x42:Z) := 0. Definition P_id_EQC (x42:Z) (x43:Z) := 0. Definition P_id_CASE1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_CASE8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_EXTRACT (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_AND (x42:Z) (x43:Z) := 0. Definition P_id_Marked_tops (x42:Z) := 0. Definition P_id_CASE2 (x42:Z) (x43:Z) (x44:Z) := 0. Lemma P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_not_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_not x43 <= P_id_Marked_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_DELETE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_pops x43 <= P_id_Marked_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQC_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_tops x43 <= P_id_Marked_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45:: x44::x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45:: x44::x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44:: x43::x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45:: x44::x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44:: x43::x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44:: x43::x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44:: x43::x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47:: x46::x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43::x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45:: x44::x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45:: x44::x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44:: x43::x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44:: x43::x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic; assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_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_47_large_large.DP_R_xml_0_scc_47_large_large_large . Proof. intros x. apply well_founded_ind with (R:=fun x y => (Zwf.Zwf 0) (marked_measure x) (marked_measure y)). apply Inverse_Image.wf_inverse_image with (B:=Z). apply Zwf.Zwf_well_founded. clear x. intros x IHx. repeat ( constructor;inversion 1;subst; full_prove_ineq algebra.Alg.Term ltac:(algebra.Alg_ext.find_replacement ) algebra.EQT_ext.one_step_list_refl_trans_clos marked_measure marked_measure_star_monotonic (Zwf.Zwf 0) (interp.o_Z 0) ltac:(fun _ => R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 ) ltac:(fun _ => rewrite_and_unfold ) ltac:(fun _ => generate_pos_hyp ) ltac:(fun _ => cbv beta iota zeta delta - [Zplus Zmult Zle Zlt] in * ; try (constructor)) IHx ). Qed. End WF_DP_R_xml_0_scc_47_large_large_large. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_or (x42:Z) (x43:Z) := 1. Definition P_id_gen_tag (x42:Z) := 2 + 2* x42. Definition P_id_record_new (x42:Z) := 1. Definition P_id_a := 0. Definition P_id_locker2_release_lock (x42:Z) (x43:Z) := 2 + 3* x42 + 2* x43. Definition P_id_eqt (x42:Z) (x43:Z) := 0. Definition P_id_istops (x42:Z) (x43:Z) := 0. Definition P_id_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_release := 0. Definition P_id_locker2_obtainable (x42:Z) (x43:Z) := 0. Definition P_id_imp (x42:Z) (x43:Z) := 1* x43. Definition P_id_stack (x42:Z) (x43:Z) := 2* x42 + 1* x43. Definition P_id_locker2_map_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_locker := 0. Definition P_id_case4 (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_int (x42:Z) := 0. Definition P_id_push (x42:Z) (x43:Z) (x44:Z) := 1 + 2* x42 + 2* x43. Definition P_id_locker2_add_pending (x42:Z) (x43:Z) (x44:Z) := 3 + 3* x42 + 2* x43 + 1* x44. Definition P_id_true := 0. Definition P_id_locker2_check_availables (x42:Z) (x43:Z) := 2 + 3* x42. Definition P_id_F := 0. Definition P_id_eqs (x42:Z) (x43:Z) := 0. Definition P_id_record_update (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 3 + 1* x42 + 1* x44 + 2* x45. Definition P_id_false := 2. Definition P_id_gen_modtageq (x42:Z) (x43:Z) := 3* x42 + 3* x43. Definition P_id_undefined := 0. Definition P_id_nocalls := 0. Definition P_id_locker2_remove_pending (x42:Z) (x43:Z) := 3 + 3* x42. Definition P_id_resource := 0. Definition P_id_case6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 2* x45. Definition P_id_if (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 1* x44. Definition P_id_pops (x42:Z) := 1* x42. Definition P_id_locker2_map_claim_lock (x42:Z) (x43:Z) (x44:Z) := 2* x42. Definition P_id_ok := 0. Definition P_id_case5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 3 + 1* x43. Definition P_id_tuple (x42:Z) (x43:Z) := 2* x42 + 1* x43. Definition P_id_member (x42:Z) (x43:Z) := 2. Definition P_id_s (x42:Z) := 0. Definition P_id_delete (x42:Z) (x43:Z) := 1* x43. Definition P_id_T := 0. Definition P_id_case9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2* x45. Definition P_id_record_extract (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_excl := 0. Definition P_id_case2 (x42:Z) (x43:Z) (x44:Z) := 2 + 1* x42 + 2* x43. Definition P_id_nil := 0. Definition P_id_eqc (x42:Z) (x43:Z) := 0. Definition P_id_case0 (x42:Z) (x43:Z) (x44:Z) := 2 + 1* x43 + 2* x44. Definition P_id_request := 0. Definition P_id_locker2_check_available (x42:Z) (x43:Z) := 2. Definition P_id_not (x42:Z) := 2. Definition P_id_pushs (x42:Z) (x43:Z) := 1 + 2* x42 + 1* x43. Definition P_id_locker2_promote_pending (x42:Z) (x43:Z) := 2 + 3* x42. Definition P_id_mcrlrecord := 0. Definition P_id_locker2_obtainables (x42:Z) (x43:Z) := 3 + 1* x42. Definition P_id_cons (x42:Z) (x43:Z) := 2 + 1* x42 + 1* x43. Definition P_id_push1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47:Z) := 0. Definition P_id_case1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 3 + 1* x43 + 3* x44. Definition P_id_element (x42:Z) (x43:Z) := 2* x43. Definition P_id_locker2_adduniq (x42:Z) (x43:Z) := 1* x43. Definition P_id_and (x42:Z) (x43:Z) := 2* x42 + 2* x43. Definition P_id_empty := 0. Definition P_id_record_updates (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_lock := 0. Definition P_id_excllock := 0. Definition P_id_pid (x42:Z) := 0. Definition P_id_calls (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_subtract (x42:Z) (x43:Z) := 1* x42. Definition P_id_tag := 0. Definition P_id_equal (x42:Z) (x43:Z) := 0. Definition P_id_eq (x42:Z) (x43:Z) := 0. Definition P_id_tops (x42:Z) := 1* x42. Definition P_id_locker2_claim_lock (x42:Z) (x43:Z) (x44:Z) := 2 + 1* x42 . Definition P_id_pending := 0. Definition P_id_andt (x42:Z) (x43:Z) := 1. Definition P_id_tuplenil (x42:Z) := 1 + 2* x42. Definition P_id_append (x42:Z) (x43:Z) := 1* x42. Definition P_id_0 := 0. Definition P_id_case8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2 + 1* x42 + 1* x43 + 3* x45. Lemma P_id_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_record_new x43 <= P_id_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_istops x43 x45 <= P_id_istops x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_member x43 x45 <= P_id_member x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_delete x43 x45 <= P_id_delete x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_append x43 x45 <= P_id_append x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_or_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_or x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_bounded : forall x42, (0 <= x42) ->0 <= P_id_gen_tag x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_bounded : forall x42, (0 <= x42) ->0 <= P_id_record_new x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_release_lock x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_istops x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->0 <= P_id_locker2_map_add_pending x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_release_bounded : 0 <= P_id_release . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainable x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_imp x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_stack x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_bounded : forall x42 x43, (0 <= x42) -> (0 <= x43) ->0 <= P_id_locker2_map_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker_bounded : 0 <= P_id_locker . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case4 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_bounded : forall x42, (0 <= x42) ->0 <= P_id_int x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_push x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_add_pending x44 x43 x42. 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_locker2_check_availables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_availables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_F_bounded : 0 <= P_id_F . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_record_update x45 x44 x43 x42. 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_gen_modtageq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_gen_modtageq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_undefined_bounded : 0 <= P_id_undefined . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nocalls_bounded : 0 <= P_id_nocalls . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_remove_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_resource_bounded : 0 <= P_id_resource . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_case6 x45 x44 x43 x42. 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 x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_if x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_bounded : forall x42, (0 <= x42) ->0 <= P_id_pops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->0 <= P_id_locker2_map_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ok_bounded : 0 <= P_id_ok . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_case5 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_tuple x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_member x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x42, (0 <= x42) ->0 <= P_id_s x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_delete x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_T_bounded : 0 <= P_id_T . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_case9 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_extract x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excl_bounded : 0 <= P_id_excl . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case2 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nil_bounded : 0 <= P_id_nil . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqc x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case0 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_request_bounded : 0 <= P_id_request . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_available x43 x42. 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 x42, (0 <= x42) ->0 <= P_id_not x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_pushs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_mcrlrecord_bounded : 0 <= P_id_mcrlrecord . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_cons x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (0 <= x42) -> (0 <= x43) -> (0 <= x44) -> (0 <= x45) -> (0 <= x46) ->(0 <= x47) ->0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_case1 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_element x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_adduniq x43 x42. 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 x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_and x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_empty_bounded : 0 <= P_id_empty . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_updates x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_lock_bounded : 0 <= P_id_lock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excllock_bounded : 0 <= P_id_excllock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_bounded : forall x42, (0 <= x42) ->0 <= P_id_pid x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_calls x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_subtract x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tag_bounded : 0 <= P_id_tag . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_equal x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_bounded : forall x42, (0 <= x42) ->0 <= P_id_tops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pending_bounded : 0 <= P_id_pending . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_andt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_bounded : forall x42, (0 <= x42) ->0 <= P_id_tuplenil x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_append x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_case8 x45 x44 x43 x42. 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_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43::x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45:: x44::x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44::x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_GEN_MODTAGEQ (x42:Z) (x43:Z) := 0. Definition P_id_ELEMENT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_eq (x42:Z) (x43:Z) := 0. Definition P_id_CASE9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_REMOVE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_record_new (x42:Z) := 0. Definition P_id_CASE6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_if (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_MEMBER (x42:Z) (x43:Z) := 0. Definition P_id_CASE5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_UPDATE (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_Marked_not (x42:Z) := 0. Definition P_id_ISTOPS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_ADD_PENDING (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_or (x42:Z) (x43:Z) := 0. Definition P_id_DELETE (x42:Z) (x43:Z) := 0. Definition P_id_CASE0 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_imp (x42:Z) (x43:Z) := 0. Definition P_id_EQT (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_PUSHS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_RELEASE_LOCK (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_OBTAINABLES (x42:Z) (x43:Z) := 0. Definition P_id_RECORD_UPDATES (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_adduniq (x42:Z) (x43:Z) := 0. Definition P_id_EQS (x42:Z) (x43:Z) := 0. Definition P_id_SUBTRACT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_gen_tag (x42:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLES (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_CLAIM_LOCK (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_case4 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47:Z) := 0. Definition P_id_APPEND (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLE (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_pops (x42:Z) := 0. Definition P_id_EQC (x42:Z) (x43:Z) := 0. Definition P_id_CASE1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_CASE8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_EXTRACT (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_AND (x42:Z) (x43:Z) := 0. Definition P_id_Marked_tops (x42:Z) := 0. Definition P_id_CASE2 (x42:Z) (x43:Z) (x44:Z) := 0. Lemma P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_not_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_not x43 <= P_id_Marked_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_DELETE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_pops x43 <= P_id_Marked_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQC_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_tops x43 <= P_id_Marked_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45:: x44::x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45:: x44::x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45:: x44::x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44:: x43::x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44:: x43::x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44:: x43::x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47:: x46::x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43::x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45:: x44::x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45:: x44::x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44:: x43::x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44:: x43::x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_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_47_large_large_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_47_large_large_strict lt. Proof. unfold Relation_Definitions.inclusion, lt in *. intros a b H;destruct H; match goal with | |- (Zwf.Zwf 0) _ (marked_measure (algebra.Alg.Term ?f ?l)) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_lt_compat_right (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma DP_R_xml_0_scc_47_large_large_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_47_large_large_large le. Proof. unfold Relation_Definitions.inclusion, le, Zwf.Zwf in *. intros a b H;destruct H; match goal with | |- _ <= marked_measure (algebra.Alg.Term ?f ?l) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_trans (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition wf_DP_R_xml_0_scc_47_large_large_large := WF_DP_R_xml_0_scc_47_large_large_large.wf. Lemma wf : well_founded WF_DP_R_xml_0_scc_47_large.DP_R_xml_0_scc_47_large_large. Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_47_large_large_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx';apply DP_R_xml_0_scc_47_large_large_strict_in_lt; econstructor eassumption)|| ((apply IHx;[econstructor eassumption| intros y' Hlt;apply IHx';apply lt_le_compat with (1:=Hlt) ; apply DP_R_xml_0_scc_47_large_large_large_in_le; econstructor eassumption])). apply wf_DP_R_xml_0_scc_47_large_large_large. Qed. End WF_DP_R_xml_0_scc_47_large_large. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_or (x42:Z) (x43:Z) := 0. Definition P_id_gen_tag (x42:Z) := 1* x42. Definition P_id_record_new (x42:Z) := 1. Definition P_id_a := 0. Definition P_id_locker2_release_lock (x42:Z) (x43:Z) := 3 + 3* x42 + 1* x43. Definition P_id_eqt (x42:Z) (x43:Z) := 0. Definition P_id_istops (x42:Z) (x43:Z) := 2 + 1* x42. Definition P_id_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_release := 0. Definition P_id_locker2_obtainable (x42:Z) (x43:Z) := 0. Definition P_id_imp (x42:Z) (x43:Z) := 1 + 1* x43. Definition P_id_stack (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_locker2_map_promote_pending (x42:Z) (x43:Z) := 3 + 3* x42 . Definition P_id_locker := 0. Definition P_id_case4 (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_int (x42:Z) := 1* x42. Definition P_id_push (x42:Z) (x43:Z) (x44:Z) := 1 + 3* x42 + 3* x43. Definition P_id_locker2_add_pending (x42:Z) (x43:Z) (x44:Z) := 3 + 3* x42 + 2* x43 + 3* x44. Definition P_id_true := 0. Definition P_id_locker2_check_availables (x42:Z) (x43:Z) := 2 + 2* x43. Definition P_id_F := 0. Definition P_id_eqs (x42:Z) (x43:Z) := 0. Definition P_id_record_update (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x42 + 2* x45. Definition P_id_false := 2. Definition P_id_gen_modtageq (x42:Z) (x43:Z) := 3* x42. Definition P_id_undefined := 0. Definition P_id_nocalls := 0. Definition P_id_locker2_remove_pending (x42:Z) (x43:Z) := 2 + 3* x42 + 3* x43. Definition P_id_resource := 0. Definition P_id_case6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2* x45. Definition P_id_if (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 1* x44. Definition P_id_pops (x42:Z) := 1* x42. Definition P_id_locker2_map_claim_lock (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_ok := 0. Definition P_id_case5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_tuple (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_member (x42:Z) (x43:Z) := 2. Definition P_id_s (x42:Z) := 2* x42. Definition P_id_delete (x42:Z) (x43:Z) := 1* x43. Definition P_id_T := 0. Definition P_id_case9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x45. Definition P_id_record_extract (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_excl := 0. Definition P_id_case2 (x42:Z) (x43:Z) (x44:Z) := 2 + 1* x42 + 2* x43. Definition P_id_nil := 0. Definition P_id_eqc (x42:Z) (x43:Z) := 0. Definition P_id_case0 (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 2* x44. Definition P_id_request := 0. Definition P_id_locker2_check_available (x42:Z) (x43:Z) := 2. Definition P_id_not (x42:Z) := 0. Definition P_id_pushs (x42:Z) (x43:Z) := 2 + 2* x42 + 1* x43. Definition P_id_locker2_promote_pending (x42:Z) (x43:Z) := 2 + 3* x42. Definition P_id_mcrlrecord := 0. Definition P_id_locker2_obtainables (x42:Z) (x43:Z) := 0. Definition P_id_cons (x42:Z) (x43:Z) := 1 + 1* x42 + 1* x43. Definition P_id_push1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 1 + 3* x42. Definition P_id_case1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2 + 3* x42 + 1* x43 + 3* x44. Definition P_id_element (x42:Z) (x43:Z) := 1* x42 + 2* x43. Definition P_id_locker2_adduniq (x42:Z) (x43:Z) := 1* x43. Definition P_id_and (x42:Z) (x43:Z) := 2* x42 + 2* x43. Definition P_id_empty := 0. Definition P_id_record_updates (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44 . Definition P_id_lock := 0. Definition P_id_excllock := 0. Definition P_id_pid (x42:Z) := 0. Definition P_id_calls (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_subtract (x42:Z) (x43:Z) := 1* x42. Definition P_id_tag := 0. Definition P_id_equal (x42:Z) (x43:Z) := 0. Definition P_id_eq (x42:Z) (x43:Z) := 0. Definition P_id_tops (x42:Z) := 1* x42. Definition P_id_locker2_claim_lock (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_pending := 0. Definition P_id_andt (x42:Z) (x43:Z) := 1* x42. Definition P_id_tuplenil (x42:Z) := 2* x42. Definition P_id_append (x42:Z) (x43:Z) := 1* x42. Definition P_id_0 := 0. Definition P_id_case8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x42 + 1* x43. Lemma P_id_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_record_new x43 <= P_id_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_istops x43 x45 <= P_id_istops x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_member x43 x45 <= P_id_member x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_delete x43 x45 <= P_id_delete x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_append x43 x45 <= P_id_append x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_or_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_or x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_bounded : forall x42, (0 <= x42) ->0 <= P_id_gen_tag x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_bounded : forall x42, (0 <= x42) ->0 <= P_id_record_new x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_release_lock x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_istops x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->0 <= P_id_locker2_map_add_pending x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_release_bounded : 0 <= P_id_release . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainable x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_imp x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_stack x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_map_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker_bounded : 0 <= P_id_locker . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case4 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_bounded : forall x42, (0 <= x42) ->0 <= P_id_int x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_push x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_add_pending x44 x43 x42. 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_locker2_check_availables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_availables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_F_bounded : 0 <= P_id_F . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_record_update x45 x44 x43 x42. 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_gen_modtageq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_gen_modtageq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_undefined_bounded : 0 <= P_id_undefined . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nocalls_bounded : 0 <= P_id_nocalls . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_remove_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_resource_bounded : 0 <= P_id_resource . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case6 x45 x44 x43 x42. 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 x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_if x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_bounded : forall x42, (0 <= x42) ->0 <= P_id_pops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ok_bounded : 0 <= P_id_ok . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case5 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_tuple x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_member x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x42, (0 <= x42) ->0 <= P_id_s x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_delete x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_T_bounded : 0 <= P_id_T . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case9 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_extract x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excl_bounded : 0 <= P_id_excl . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case2 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nil_bounded : 0 <= P_id_nil . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqc x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case0 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_request_bounded : 0 <= P_id_request . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_available x43 x42. 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 x42, (0 <= x42) ->0 <= P_id_not x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_pushs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_mcrlrecord_bounded : 0 <= P_id_mcrlrecord . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_cons x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (0 <= x42) -> (0 <= x43) -> (0 <= x44) -> (0 <= x45) -> (0 <= x46) ->(0 <= x47) ->0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case1 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_element x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_adduniq x43 x42. 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 x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_and x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_empty_bounded : 0 <= P_id_empty . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_updates x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_lock_bounded : 0 <= P_id_lock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excllock_bounded : 0 <= P_id_excllock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_bounded : forall x42, (0 <= x42) ->0 <= P_id_pid x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_calls x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_subtract x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tag_bounded : 0 <= P_id_tag . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_equal x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_bounded : forall x42, (0 <= x42) ->0 <= P_id_tops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pending_bounded : 0 <= P_id_pending . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_andt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_bounded : forall x42, (0 <= x42) ->0 <= P_id_tuplenil x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_append x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case8 x45 x44 x43 x42. 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_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43::x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43::x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43::x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45:: x44::x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44::x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_GEN_MODTAGEQ (x42:Z) (x43:Z) := 0. Definition P_id_ELEMENT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_eq (x42:Z) (x43:Z) := 0. Definition P_id_CASE9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_REMOVE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_record_new (x42:Z) := 0. Definition P_id_CASE6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_if (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_MEMBER (x42:Z) (x43:Z) := 0. Definition P_id_CASE5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_UPDATE (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_Marked_not (x42:Z) := 0. Definition P_id_ISTOPS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_ADD_PENDING (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_or (x42:Z) (x43:Z) := 0. Definition P_id_DELETE (x42:Z) (x43:Z) := 0. Definition P_id_CASE0 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_imp (x42:Z) (x43:Z) := 0. Definition P_id_EQT (x42:Z) (x43:Z) := 3* x42 + 3* x43. Definition P_id_PUSHS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_RELEASE_LOCK (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_OBTAINABLES (x42:Z) (x43:Z) := 0. Definition P_id_RECORD_UPDATES (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_adduniq (x42:Z) (x43:Z) := 0. Definition P_id_EQS (x42:Z) (x43:Z) := 0. Definition P_id_SUBTRACT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_gen_tag (x42:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLES (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_CLAIM_LOCK (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_case4 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_APPEND (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLE (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_pops (x42:Z) := 0. Definition P_id_EQC (x42:Z) (x43:Z) := 0. Definition P_id_CASE1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_CASE8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_EXTRACT (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_AND (x42:Z) (x43:Z) := 0. Definition P_id_Marked_tops (x42:Z) := 0. Definition P_id_CASE2 (x42:Z) (x43:Z) (x44:Z) := 0. Lemma P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_not_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_not x43 <= P_id_Marked_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_DELETE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_pops x43 <= P_id_Marked_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQC_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_tops x43 <= P_id_Marked_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44:: x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44:: x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44:: x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46:: x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44:: x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44:: x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44:: x43::x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_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_47_large_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_47_large_strict lt. Proof. unfold Relation_Definitions.inclusion, lt in *. intros a b H;destruct H; match goal with | |- (Zwf.Zwf 0) _ (marked_measure (algebra.Alg.Term ?f ?l)) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_lt_compat_right (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma DP_R_xml_0_scc_47_large_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_47_large_large le. Proof. unfold Relation_Definitions.inclusion, le, Zwf.Zwf in *. intros a b H;destruct H; match goal with | |- _ <= marked_measure (algebra.Alg.Term ?f ?l) => let l'' := algebra.Alg_ext.find_replacement l in ((apply (interp.le_trans (interp.o_Z 0)) with (marked_measure (algebra.Alg.Term f l''));[idtac| apply marked_measure_star_monotonic; repeat (apply algebra.EQT_ext.one_step_list_refl_trans_clos); (assumption)||(constructor 1)])) end ;clear ;rewrite_and_unfold ;repeat (generate_pos_hyp ); cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition wf_DP_R_xml_0_scc_47_large_large := WF_DP_R_xml_0_scc_47_large_large.wf. Lemma wf : well_founded WF_DP_R_xml_0_scc_47.DP_R_xml_0_scc_47_large. Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_47_large_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx';apply DP_R_xml_0_scc_47_large_strict_in_lt; econstructor eassumption)|| ((apply IHx;[econstructor eassumption| intros y' Hlt;apply IHx';apply lt_le_compat with (1:=Hlt) ; apply DP_R_xml_0_scc_47_large_large_in_le;econstructor eassumption])). apply wf_DP_R_xml_0_scc_47_large_large. Qed. End WF_DP_R_xml_0_scc_47_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_or (x42:Z) (x43:Z) := 1. Definition P_id_gen_tag (x42:Z) := 3* x42. Definition P_id_record_new (x42:Z) := 1. Definition P_id_a := 0. Definition P_id_locker2_release_lock (x42:Z) (x43:Z) := 2* x42 + 2* x43. Definition P_id_eqt (x42:Z) (x43:Z) := 0. Definition P_id_istops (x42:Z) (x43:Z) := 1 + 2* x42 + 1* x43. Definition P_id_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_release := 0. Definition P_id_locker2_obtainable (x42:Z) (x43:Z) := 0. Definition P_id_imp (x42:Z) (x43:Z) := 1* x43. Definition P_id_stack (x42:Z) (x43:Z) := 2* x42 + 2* x43. Definition P_id_locker2_map_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_locker := 0. Definition P_id_case4 (x42:Z) (x43:Z) (x44:Z) := 1. Definition P_id_int (x42:Z) := 0. Definition P_id_push (x42:Z) (x43:Z) (x44:Z) := 3 + 2* x42 + 1* x43 + 3* x44. Definition P_id_locker2_add_pending (x42:Z) (x43:Z) (x44:Z) := 3* x42 + 2* x43 + 3* x44. Definition P_id_true := 0. Definition P_id_locker2_check_availables (x42:Z) (x43:Z) := 1 + 3* x43. Definition P_id_F := 0. Definition P_id_eqs (x42:Z) (x43:Z) := 0. Definition P_id_record_update (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1* x42 + 1* x44 + 2* x45. Definition P_id_false := 0. Definition P_id_gen_modtageq (x42:Z) (x43:Z) := 2 + 3* x42. Definition P_id_undefined := 0. Definition P_id_nocalls := 2. Definition P_id_locker2_remove_pending (x42:Z) (x43:Z) := 3* x42 + 1* x43. Definition P_id_resource := 0. Definition P_id_case6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2 + 3* x42. Definition P_id_if (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 1* x44. Definition P_id_pops (x42:Z) := 1* x42. Definition P_id_locker2_map_claim_lock (x42:Z) (x43:Z) (x44:Z) := 2* x42. Definition P_id_ok := 0. Definition P_id_case5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2 + 1* x43. Definition P_id_tuple (x42:Z) (x43:Z) := 2* x42 + 1* x43. Definition P_id_member (x42:Z) (x43:Z) := 0. Definition P_id_s (x42:Z) := 0. Definition P_id_delete (x42:Z) (x43:Z) := 1* x43. Definition P_id_T := 0. Definition P_id_case9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_record_extract (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_excl := 0. Definition P_id_case2 (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x43. Definition P_id_nil := 0. Definition P_id_eqc (x42:Z) (x43:Z) := 0. Definition P_id_case0 (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 2* x44. Definition P_id_request := 0. Definition P_id_locker2_check_available (x42:Z) (x43:Z) := 2 + 2* x43. Definition P_id_not (x42:Z) := 1. Definition P_id_pushs (x42:Z) (x43:Z) := 2* x42 + 2* x43. Definition P_id_locker2_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_mcrlrecord := 0. Definition P_id_locker2_obtainables (x42:Z) (x43:Z) := 2 + 1* x42. Definition P_id_cons (x42:Z) (x43:Z) := 1* x42 + 2* x43. Definition P_id_push1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_case1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 3* x42 + 1* x43 + 3* x44. Definition P_id_element (x42:Z) (x43:Z) := 1* x43. Definition P_id_locker2_adduniq (x42:Z) (x43:Z) := 1* x43. Definition P_id_and (x42:Z) (x43:Z) := 1* x42 + 2* x43. Definition P_id_empty := 0. Definition P_id_record_updates (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_lock := 0. Definition P_id_excllock := 0. Definition P_id_pid (x42:Z) := 3 + 2* x42. Definition P_id_calls (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_subtract (x42:Z) (x43:Z) := 1* x42. Definition P_id_tag := 0. Definition P_id_equal (x42:Z) (x43:Z) := 0. Definition P_id_eq (x42:Z) (x43:Z) := 2. Definition P_id_tops (x42:Z) := 1* x42. Definition P_id_locker2_claim_lock (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_pending := 0. Definition P_id_andt (x42:Z) (x43:Z) := 1. Definition P_id_tuplenil (x42:Z) := 2* x42. Definition P_id_append (x42:Z) (x43:Z) := 1* x42. Definition P_id_0 := 0. Definition P_id_case8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2* x42 + 1* x43. Lemma P_id_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_record_new x43 <= P_id_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_istops x43 x45 <= P_id_istops x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_member x43 x45 <= P_id_member x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_delete x43 x45 <= P_id_delete x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_append x43 x45 <= P_id_append x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_or_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_or x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_bounded : forall x42, (0 <= x42) ->0 <= P_id_gen_tag x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_bounded : forall x42, (0 <= x42) ->0 <= P_id_record_new x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_release_lock x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_istops x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_add_pending x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_release_bounded : 0 <= P_id_release . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainable x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_imp x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_stack x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_map_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker_bounded : 0 <= P_id_locker . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case4 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_bounded : forall x42, (0 <= x42) ->0 <= P_id_int x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_push x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_add_pending x44 x43 x42. 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_locker2_check_availables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_availables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_F_bounded : 0 <= P_id_F . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_record_update x45 x44 x43 x42. 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_gen_modtageq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_gen_modtageq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_undefined_bounded : 0 <= P_id_undefined . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nocalls_bounded : 0 <= P_id_nocalls . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_remove_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_resource_bounded : 0 <= P_id_resource . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case6 x45 x44 x43 x42. 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 x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_if x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_bounded : forall x42, (0 <= x42) ->0 <= P_id_pops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ok_bounded : 0 <= P_id_ok . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case5 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_tuple x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_member x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x42, (0 <= x42) ->0 <= P_id_s x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_delete x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_T_bounded : 0 <= P_id_T . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case9 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_extract x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excl_bounded : 0 <= P_id_excl . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case2 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nil_bounded : 0 <= P_id_nil . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqc x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case0 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_request_bounded : 0 <= P_id_request . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_available x43 x42. 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 x42, (0 <= x42) ->0 <= P_id_not x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_pushs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_mcrlrecord_bounded : 0 <= P_id_mcrlrecord . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_cons x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (0 <= x42) -> (0 <= x43) -> (0 <= x44) -> (0 <= x45) -> (0 <= x46) ->(0 <= x47) ->0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case1 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_element x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_adduniq x43 x42. 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 x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_and x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_empty_bounded : 0 <= P_id_empty . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_updates x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_lock_bounded : 0 <= P_id_lock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excllock_bounded : 0 <= P_id_excllock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_bounded : forall x42, (0 <= x42) ->0 <= P_id_pid x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_calls x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_subtract x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tag_bounded : 0 <= P_id_tag . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_equal x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_bounded : forall x42, (0 <= x42) ->0 <= P_id_tops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pending_bounded : 0 <= P_id_pending . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_andt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_bounded : forall x42, (0 <= x42) ->0 <= P_id_tuplenil x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_append x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case8 x45 x44 x43 x42. 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_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43::x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43:: x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43::x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44:: x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44:: x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_GEN_MODTAGEQ (x42:Z) (x43:Z) := 0. Definition P_id_ELEMENT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_eq (x42:Z) (x43:Z) := 0. Definition P_id_CASE9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_REMOVE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_record_new (x42:Z) := 0. Definition P_id_CASE6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_if (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_MEMBER (x42:Z) (x43:Z) := 0. Definition P_id_CASE5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_UPDATE (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_Marked_not (x42:Z) := 0. Definition P_id_ISTOPS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_ADD_PENDING (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_or (x42:Z) (x43:Z) := 0. Definition P_id_DELETE (x42:Z) (x43:Z) := 0. Definition P_id_CASE0 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_imp (x42:Z) (x43:Z) := 0. Definition P_id_EQT (x42:Z) (x43:Z) := 1* x42 + 2* x43. Definition P_id_PUSHS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_RELEASE_LOCK (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_OBTAINABLES (x42:Z) (x43:Z) := 0. Definition P_id_RECORD_UPDATES (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_adduniq (x42:Z) (x43:Z) := 0. Definition P_id_EQS (x42:Z) (x43:Z) := 0. Definition P_id_SUBTRACT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_gen_tag (x42:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLES (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_CLAIM_LOCK (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_case4 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_APPEND (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLE (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_pops (x42:Z) := 0. Definition P_id_EQC (x42:Z) (x43:Z) := 0. Definition P_id_CASE1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_CASE8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_EXTRACT (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_map_add_pending (x42:Z) (x43:Z) (x44: Z) := 0. Definition P_id_AND (x42:Z) (x43:Z) := 0. Definition P_id_Marked_tops (x42:Z) := 0. Definition P_id_CASE2 (x42:Z) (x43:Z) (x44:Z) := 0. Lemma P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_not_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_not x43 <= P_id_Marked_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_DELETE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_pops x43 <= P_id_Marked_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQC_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_tops x43 <= P_id_Marked_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44:: x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44:: x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44:: x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46:: x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44:: x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44:: x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_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_47_strict_in_lt : Relation_Definitions.inclusion _ DP_R_xml_0_scc_47_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_47_large_in_le : Relation_Definitions.inclusion _ DP_R_xml_0_scc_47_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_47_large := WF_DP_R_xml_0_scc_47_large.wf. Lemma wf : well_founded WF_DP_R_xml_0.DP_R_xml_0_scc_47. Proof. intros x. apply (well_founded_ind wf_lt). clear x. intros x. pattern x. apply (@Acc_ind _ DP_R_xml_0_scc_47_large). clear x. intros x _ IHx IHx'. constructor. intros y H. destruct H; (apply IHx';apply DP_R_xml_0_scc_47_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_47_large_in_le;econstructor eassumption])). apply wf_DP_R_xml_0_scc_47_large. Qed. End WF_DP_R_xml_0_scc_47. Definition wf_DP_R_xml_0_scc_47 := WF_DP_R_xml_0_scc_47.wf. Lemma acc_DP_R_xml_0_scc_47 : forall x y, (DP_R_xml_0_scc_47 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_47). 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_46; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_45; 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_47. Qed. Inductive DP_R_xml_0_non_scc_48 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_48_0 : forall x40 x36 x44 x34 x42 x37 x43 x39, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x36 x44) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x34 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x40::x37::x39::nil)) x42) -> DP_R_xml_0_non_scc_48 (algebra.Alg.Term algebra.F.id_eqt (x36:: x40::nil)) (algebra.Alg.Term algebra.F.id_push (x44::x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_48 : forall x y, (DP_R_xml_0_non_scc_48 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_47; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_46; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_45; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). Qed. Inductive DP_R_xml_0_non_scc_49 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_49_0 : forall x36 x34 x42 x38 x37 x35 x43 x39, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x36::x37::x39::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x34::x35::x38::nil)) x42) -> DP_R_xml_0_non_scc_49 (algebra.Alg.Term algebra.F.id_eqt (x36:: x34::nil)) (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_49 : forall x y, (DP_R_xml_0_non_scc_49 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_47; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_46; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_45; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). Qed. Inductive DP_R_xml_0_non_scc_50 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_50_0 : forall x36 x34 x42 x37 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) x36 x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_stack (x34::x37::nil)) x42) -> DP_R_xml_0_non_scc_50 (algebra.Alg.Term algebra.F.id_eqt (x36:: x34::nil)) (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_50 : forall x y, (DP_R_xml_0_non_scc_50 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_47; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_46; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_45; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). Qed. Inductive DP_R_xml_0_non_scc_51 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_51_0 : forall x36 x34 x42 x37 x35 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_stack (x36::x37::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_stack (x34::x35::nil)) x42) -> DP_R_xml_0_non_scc_51 (algebra.Alg.Term algebra.F.id_eqt (x36:: x34::nil)) (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_51 : forall x y, (DP_R_xml_0_non_scc_51 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_47; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_46; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_45; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((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_52 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_52_0 : forall x36 x34 x42 x37 x35 x43, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_stack (x36::x37::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_stack (x34::x35::nil)) x42) -> DP_R_xml_0_scc_52 (algebra.Alg.Term algebra.F.id_eqs (x37::x35::nil)) (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) . Module WF_DP_R_xml_0_scc_52. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_or (x42:Z) (x43:Z) := 0. Definition P_id_gen_tag (x42:Z) := 2* x42. Definition P_id_record_new (x42:Z) := 0. Definition P_id_a := 0. Definition P_id_locker2_release_lock (x42:Z) (x43:Z) := 2 + 2* x42 + 2* x43. Definition P_id_eqt (x42:Z) (x43:Z) := 0. Definition P_id_istops (x42:Z) (x43:Z) := 3* x42. Definition P_id_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_release := 0. Definition P_id_locker2_obtainable (x42:Z) (x43:Z) := 0. Definition P_id_imp (x42:Z) (x43:Z) := 1* x43. Definition P_id_stack (x42:Z) (x43:Z) := 2 + 2* x42 + 2* x43. Definition P_id_locker2_map_promote_pending (x42:Z) (x43:Z) := 2 + 3* x42. Definition P_id_locker := 0. Definition P_id_case4 (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_int (x42:Z) := 0. Definition P_id_push (x42:Z) (x43:Z) (x44:Z) := 3* x42 + 2* x43. Definition P_id_locker2_add_pending (x42:Z) (x43:Z) (x44:Z) := 3 + 3* x42 + 2* x43 + 3* x44. Definition P_id_true := 0. Definition P_id_locker2_check_availables (x42:Z) (x43:Z) := 2 + 3* x43. Definition P_id_F := 0. Definition P_id_eqs (x42:Z) (x43:Z) := 0. Definition P_id_record_update (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1* x42 + 3* x44 + 2* x45. Definition P_id_false := 0. Definition P_id_gen_modtageq (x42:Z) (x43:Z) := 2* x42 + 3* x43. Definition P_id_undefined := 0. Definition P_id_nocalls := 0. Definition P_id_locker2_remove_pending (x42:Z) (x43:Z) := 3 + 3* x42 + 1* x43. Definition P_id_resource := 0. Definition P_id_case6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 3* x42 + 1* x43. Definition P_id_if (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 1* x44. Definition P_id_pops (x42:Z) := 1* x42. Definition P_id_locker2_map_claim_lock (x42:Z) (x43:Z) (x44:Z) := 2* x42. Definition P_id_ok := 0. Definition P_id_case5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2* x43. Definition P_id_tuple (x42:Z) (x43:Z) := 2* x42 + 1* x43. Definition P_id_member (x42:Z) (x43:Z) := 0. Definition P_id_s (x42:Z) := 0. Definition P_id_delete (x42:Z) (x43:Z) := 1* x43. Definition P_id_T := 0. Definition P_id_case9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_record_extract (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_excl := 0. Definition P_id_case2 (x42:Z) (x43:Z) (x44:Z) := 2 + 1* x42 + 2* x43. Definition P_id_nil := 0. Definition P_id_eqc (x42:Z) (x43:Z) := 0. Definition P_id_case0 (x42:Z) (x43:Z) (x44:Z) := 2 + 1* x43 + 2* x44. Definition P_id_request := 0. Definition P_id_locker2_check_available (x42:Z) (x43:Z) := 3* x43. Definition P_id_not (x42:Z) := 0. Definition P_id_pushs (x42:Z) (x43:Z) := 3 + 3* x42 + 2* x43. Definition P_id_locker2_promote_pending (x42:Z) (x43:Z) := 2 + 3* x42. Definition P_id_mcrlrecord := 0. Definition P_id_locker2_obtainables (x42:Z) (x43:Z) := 2* x42. Definition P_id_cons (x42:Z) (x43:Z) := 1 + 1* x42 + 1* x43. Definition P_id_push1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 3* x42 + 2* x43. Definition P_id_case1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 3 + 3* x42 + 1* x43 + 3* x44. Definition P_id_element (x42:Z) (x43:Z) := 2* x43. Definition P_id_locker2_adduniq (x42:Z) (x43:Z) := 1* x43. Definition P_id_and (x42:Z) (x43:Z) := 1* x42 + 2* x43. Definition P_id_empty := 0. Definition P_id_record_updates (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 2* x44. Definition P_id_lock := 0. Definition P_id_excllock := 0. Definition P_id_pid (x42:Z) := 0. Definition P_id_calls (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_subtract (x42:Z) (x43:Z) := 1* x42. Definition P_id_tag := 0. Definition P_id_equal (x42:Z) (x43:Z) := 2* x43. Definition P_id_eq (x42:Z) (x43:Z) := 0. Definition P_id_tops (x42:Z) := 1* x42. Definition P_id_locker2_claim_lock (x42:Z) (x43:Z) (x44:Z) := 1. Definition P_id_pending := 0. Definition P_id_andt (x42:Z) (x43:Z) := 0. Definition P_id_tuplenil (x42:Z) := 1* x42. Definition P_id_append (x42:Z) (x43:Z) := 1* x42. Definition P_id_0 := 0. Definition P_id_case8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 1* x42 + 1* x43. Lemma P_id_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_record_new x43 <= P_id_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_istops x43 x45 <= P_id_istops x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_member x43 x45 <= P_id_member x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_delete x43 x45 <= P_id_delete x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_append x43 x45 <= P_id_append x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_or_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_or x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_bounded : forall x42, (0 <= x42) ->0 <= P_id_gen_tag x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_bounded : forall x42, (0 <= x42) ->0 <= P_id_record_new x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_release_lock x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_istops x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_add_pending x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_release_bounded : 0 <= P_id_release . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainable x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_imp x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_stack x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_map_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker_bounded : 0 <= P_id_locker . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case4 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_bounded : forall x42, (0 <= x42) ->0 <= P_id_int x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_push x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_add_pending x44 x43 x42. 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_locker2_check_availables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_availables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_F_bounded : 0 <= P_id_F . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_record_update x45 x44 x43 x42. 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_gen_modtageq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_gen_modtageq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_undefined_bounded : 0 <= P_id_undefined . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nocalls_bounded : 0 <= P_id_nocalls . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_remove_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_resource_bounded : 0 <= P_id_resource . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case6 x45 x44 x43 x42. 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 x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_if x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_bounded : forall x42, (0 <= x42) ->0 <= P_id_pops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ok_bounded : 0 <= P_id_ok . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case5 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_tuple x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_member x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x42, (0 <= x42) ->0 <= P_id_s x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_delete x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_T_bounded : 0 <= P_id_T . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case9 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_extract x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excl_bounded : 0 <= P_id_excl . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case2 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nil_bounded : 0 <= P_id_nil . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqc x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case0 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_request_bounded : 0 <= P_id_request . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_available x43 x42. 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 x42, (0 <= x42) ->0 <= P_id_not x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_pushs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_mcrlrecord_bounded : 0 <= P_id_mcrlrecord . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_cons x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (0 <= x42) -> (0 <= x43) -> (0 <= x44) -> (0 <= x45) -> (0 <= x46) ->(0 <= x47) ->0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case1 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_element x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_adduniq x43 x42. 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 x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_and x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_empty_bounded : 0 <= P_id_empty . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_updates x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_lock_bounded : 0 <= P_id_lock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excllock_bounded : 0 <= P_id_excllock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_bounded : forall x42, (0 <= x42) ->0 <= P_id_pid x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_calls x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_subtract x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tag_bounded : 0 <= P_id_tag . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_equal x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_bounded : forall x42, (0 <= x42) ->0 <= P_id_tops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pending_bounded : 0 <= P_id_pending . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_andt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_bounded : forall x42, (0 <= x42) ->0 <= P_id_tuplenil x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_append x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case8 x45 x44 x43 x42. 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_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43::x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43:: x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43::x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44:: x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44:: x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_GEN_MODTAGEQ (x42:Z) (x43:Z) := 0. Definition P_id_ELEMENT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_eq (x42:Z) (x43:Z) := 0. Definition P_id_CASE9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_REMOVE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_record_new (x42:Z) := 0. Definition P_id_CASE6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_if (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_MEMBER (x42:Z) (x43:Z) := 0. Definition P_id_CASE5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_UPDATE (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_Marked_not (x42:Z) := 0. Definition P_id_ISTOPS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_ADD_PENDING (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_or (x42:Z) (x43:Z) := 0. Definition P_id_DELETE (x42:Z) (x43:Z) := 0. Definition P_id_CASE0 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_imp (x42:Z) (x43:Z) := 0. Definition P_id_EQT (x42:Z) (x43:Z) := 0. Definition P_id_PUSHS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_RELEASE_LOCK (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_OBTAINABLES (x42:Z) (x43:Z) := 0. Definition P_id_RECORD_UPDATES (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_adduniq (x42:Z) (x43:Z) := 0. Definition P_id_EQS (x42:Z) (x43:Z) := 1* x43. Definition P_id_SUBTRACT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_gen_tag (x42:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLES (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_CLAIM_LOCK (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_case4 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_APPEND (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLE (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_pops (x42:Z) := 0. Definition P_id_EQC (x42:Z) (x43:Z) := 0. Definition P_id_CASE1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_CASE8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_EXTRACT (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_map_add_pending (x42:Z) (x43:Z) (x44: Z) := 0. Definition P_id_AND (x42:Z) (x43:Z) := 0. Definition P_id_Marked_tops (x42:Z) := 0. Definition P_id_CASE2 (x42:Z) (x43:Z) (x44:Z) := 0. Lemma P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_not_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_not x43 <= P_id_Marked_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_DELETE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_pops x43 <= P_id_Marked_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQC_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_tops x43 <= P_id_Marked_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44:: x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44:: x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44:: x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46:: x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44:: x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44:: x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_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_52. 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_52. Definition wf_DP_R_xml_0_scc_52 := WF_DP_R_xml_0_scc_52.wf. Lemma acc_DP_R_xml_0_scc_52 : forall x y, (DP_R_xml_0_scc_52 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_52). 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_51; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_5; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). apply wf_DP_R_xml_0_scc_52. Qed. Inductive DP_R_xml_0_non_scc_53 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_non_scc_53_0 : forall x36 x34 x42 x38 x37 x35 x43 x39, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x36::x37::x39::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x34::x35::x38::nil)) x42) -> DP_R_xml_0_non_scc_53 (algebra.Alg.Term algebra.F.id_eqs (x37:: x35::nil)) (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) . Lemma acc_DP_R_xml_0_non_scc_53 : forall x y, (DP_R_xml_0_non_scc_53 x y) ->Acc WF_R_xml_0_deep_rew.DP_R_xml_0 x. Proof. intros x y h. inversion h;clear h;subst; constructor;intros _y _h;inversion _h;clear _h;subst; (eapply acc_DP_R_xml_0_scc_52; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_51; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_5; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))). Qed. Inductive DP_R_xml_0_scc_54 : algebra.Alg.term ->algebra.Alg.term ->Prop := (* *) | DP_R_xml_0_scc_54_0 : forall x36 x34 x42 x38 x37 x35 x43 x39, (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x36::x37::x39::nil)) x43) -> (closure.refl_trans_clos (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules) (algebra.Alg.Term algebra.F.id_calls (x34::x35::x38::nil)) x42) -> DP_R_xml_0_scc_54 (algebra.Alg.Term algebra.F.id_eqc (x39::x38::nil)) (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) . Module WF_DP_R_xml_0_scc_54. Open Scope Z_scope. Import ring_extention. Notation Local "a <= b" := (Zle a b). Notation Local "a < b" := (Zlt a b). Definition P_id_or (x42:Z) (x43:Z) := 0. Definition P_id_gen_tag (x42:Z) := 2 + 3* x42. Definition P_id_record_new (x42:Z) := 1. Definition P_id_a := 0. Definition P_id_locker2_release_lock (x42:Z) (x43:Z) := 1 + 2* x42 + 2* x43. Definition P_id_eqt (x42:Z) (x43:Z) := 0. Definition P_id_istops (x42:Z) (x43:Z) := 2 + 3* x42 + 2* x43. Definition P_id_locker2_map_add_pending (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_release := 0. Definition P_id_locker2_obtainable (x42:Z) (x43:Z) := 0. Definition P_id_imp (x42:Z) (x43:Z) := 1* x43. Definition P_id_stack (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_locker2_map_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_locker := 0. Definition P_id_case4 (x42:Z) (x43:Z) (x44:Z) := 3. Definition P_id_int (x42:Z) := 0. Definition P_id_push (x42:Z) (x43:Z) (x44:Z) := 2 + 2* x42 + 2* x43 + 2* x44. Definition P_id_locker2_add_pending (x42:Z) (x43:Z) (x44:Z) := 1 + 3* x42 + 2* x43 + 3* x44. Definition P_id_true := 0. Definition P_id_locker2_check_availables (x42:Z) (x43:Z) := 3. Definition P_id_F := 0. Definition P_id_eqs (x42:Z) (x43:Z) := 0. Definition P_id_record_update (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1* x42 + 3* x44 + 1* x45. Definition P_id_false := 0. Definition P_id_gen_modtageq (x42:Z) (x43:Z) := 1* x42. Definition P_id_undefined := 0. Definition P_id_nocalls := 0. Definition P_id_locker2_remove_pending (x42:Z) (x43:Z) := 1 + 3* x42 + 1* x43. Definition P_id_resource := 0. Definition P_id_case6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 3* x42 + 3* x43. Definition P_id_if (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 1* x44. Definition P_id_pops (x42:Z) := 1* x42. Definition P_id_locker2_map_claim_lock (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_ok := 0. Definition P_id_case5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_tuple (x42:Z) (x43:Z) := 2* x42 + 1* x43. Definition P_id_member (x42:Z) (x43:Z) := 1* x43. Definition P_id_s (x42:Z) := 0. Definition P_id_delete (x42:Z) (x43:Z) := 1* x43. Definition P_id_T := 0. Definition P_id_case9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2* x42 + 2* x43. Definition P_id_record_extract (x42:Z) (x43:Z) (x44:Z) := 1* x42. Definition P_id_excl := 0. Definition P_id_case2 (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x43. Definition P_id_nil := 0. Definition P_id_eqc (x42:Z) (x43:Z) := 0. Definition P_id_case0 (x42:Z) (x43:Z) (x44:Z) := 1* x43 + 2* x44. Definition P_id_request := 0. Definition P_id_locker2_check_available (x42:Z) (x43:Z) := 2* x43. Definition P_id_not (x42:Z) := 2. Definition P_id_pushs (x42:Z) (x43:Z) := 1* x42 + 2* x43. Definition P_id_locker2_promote_pending (x42:Z) (x43:Z) := 3* x42. Definition P_id_mcrlrecord := 0. Definition P_id_locker2_obtainables (x42:Z) (x43:Z) := 0. Definition P_id_cons (x42:Z) (x43:Z) := 2* x42 + 2* x43. Definition P_id_push1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 2 + 2* x42 + 2* x43 + 2* x46. Definition P_id_case1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 1 + 3* x42 + 1* x43 + 3* x44. Definition P_id_element (x42:Z) (x43:Z) := 1* x43. Definition P_id_locker2_adduniq (x42:Z) (x43:Z) := 1* x43. Definition P_id_and (x42:Z) (x43:Z) := 2* x42 + 2* x43. Definition P_id_empty := 0. Definition P_id_record_updates (x42:Z) (x43:Z) (x44:Z) := 1* x42 + 1* x44. Definition P_id_lock := 0. Definition P_id_excllock := 0. Definition P_id_pid (x42:Z) := 0. Definition P_id_calls (x42:Z) (x43:Z) (x44:Z) := 1 + 2* x44. Definition P_id_subtract (x42:Z) (x43:Z) := 1* x42. Definition P_id_tag := 0. Definition P_id_equal (x42:Z) (x43:Z) := 0. Definition P_id_eq (x42:Z) (x43:Z) := 0. Definition P_id_tops (x42:Z) := 1* x42. Definition P_id_locker2_claim_lock (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_pending := 0. Definition P_id_andt (x42:Z) (x43:Z) := 0. Definition P_id_tuplenil (x42:Z) := 1* x42. Definition P_id_append (x42:Z) (x43:Z) := 1* x42. Definition P_id_0 := 0. Definition P_id_case8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 2* x42 + 2* x43. Lemma P_id_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_or x43 x45 <= P_id_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_gen_tag x43 <= P_id_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_record_new x43 <= P_id_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_release_lock x43 x45 <= P_id_locker2_release_lock x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqt x43 x45 <= P_id_eqt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_istops x43 x45 <= P_id_istops x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_add_pending x43 x45 x47 <= P_id_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainable x43 x45 <= P_id_locker2_obtainable x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_imp x43 x45 <= P_id_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_stack x43 x45 <= P_id_stack x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_promote_pending x43 x45 <= P_id_locker2_map_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case4 x43 x45 x47 <= P_id_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_int x43 <= P_id_int x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push x43 x45 x47 <= P_id_push x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_add_pending x43 x45 x47 <= P_id_locker2_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_availables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_availables x43 x45 <= P_id_locker2_check_availables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqs x43 x45 <= P_id_eqs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_update x43 x45 x47 x49 <= P_id_record_update x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_modtageq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_gen_modtageq x43 x45 <= P_id_gen_modtageq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_remove_pending x43 x45 <= P_id_locker2_remove_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case6 x43 x45 x47 x49 <= P_id_case6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_if x43 x45 x47 <= P_id_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pops x43 <= P_id_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_map_claim_lock x43 x45 x47 <= P_id_locker2_map_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case5 x43 x45 x47 x49 <= P_id_case5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_tuple x43 x45 <= P_id_tuple x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_member x43 x45 <= P_id_member x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_s x43 <= P_id_s x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_delete x43 x45 <= P_id_delete x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case9 x43 x45 x47 x49 <= P_id_case9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_extract x43 x45 x47 <= P_id_record_extract x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case2 x43 x45 x47 <= P_id_case2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eqc x43 x45 <= P_id_eqc x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case0 x43 x45 x47 <= P_id_case0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_check_available x43 x45 <= P_id_locker2_check_available x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_not x43 <= P_id_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_pushs x43 x45 <= P_id_pushs x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_promote_pending x43 x45 <= P_id_locker2_promote_pending x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_obtainables x43 x45 <= P_id_locker2_obtainables x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_cons x43 x45 <= P_id_cons x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_push1 x43 x45 x47 x49 x51 x53 <= P_id_push1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case1 x43 x45 x47 x49 <= P_id_case1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_element x43 x45 <= P_id_element x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_adduniq x43 x45 <= P_id_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_and x43 x45 <= P_id_and x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_record_updates x43 x45 x47 <= P_id_record_updates x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_pid x43 <= P_id_pid x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_calls x43 x45 x47 <= P_id_calls x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_subtract x43 x45 <= P_id_subtract x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_equal x43 x45 <= P_id_equal x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_eq x43 x45 <= P_id_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tops x43 <= P_id_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_locker2_claim_lock x43 x45 x47 <= P_id_locker2_claim_lock x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_andt x43 x45 <= P_id_andt x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_tuplenil x43 <= P_id_tuplenil x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_append x43 x45 <= P_id_append x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_case8 x43 x45 x47 x49 <= P_id_case8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_or_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_or x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_gen_tag_bounded : forall x42, (0 <= x42) ->0 <= P_id_gen_tag x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_new_bounded : forall x42, (0 <= x42) ->0 <= P_id_record_new x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_a_bounded : 0 <= P_id_a . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_release_lock_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_release_lock x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_istops_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_istops x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_add_pending x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_release_bounded : 0 <= P_id_release . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainable_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainable x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_imp_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_imp x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_stack_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_stack x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_map_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker_bounded : 0 <= P_id_locker . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case4_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case4 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_int_bounded : forall x42, (0 <= x42) ->0 <= P_id_int x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_push x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_add_pending_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_add_pending x44 x43 x42. 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_locker2_check_availables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_availables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_F_bounded : 0 <= P_id_F . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_update_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) -> (0 <= x44) ->(0 <= x45) ->0 <= P_id_record_update x45 x44 x43 x42. 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_gen_modtageq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_gen_modtageq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_undefined_bounded : 0 <= P_id_undefined . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nocalls_bounded : 0 <= P_id_nocalls . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_remove_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_remove_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_resource_bounded : 0 <= P_id_resource . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case6_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case6 x45 x44 x43 x42. 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 x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_if x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pops_bounded : forall x42, (0 <= x42) ->0 <= P_id_pops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_map_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_map_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ok_bounded : 0 <= P_id_ok . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case5_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case5 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuple_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_tuple x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_member_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_member x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_s_bounded : forall x42, (0 <= x42) ->0 <= P_id_s x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_delete_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_delete x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_T_bounded : 0 <= P_id_T . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case9_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case9 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_extract_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_extract x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excl_bounded : 0 <= P_id_excl . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case2_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case2 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_nil_bounded : 0 <= P_id_nil . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eqc_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eqc x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case0_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_case0 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_request_bounded : 0 <= P_id_request . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_check_available_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_check_available x43 x42. 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 x42, (0 <= x42) ->0 <= P_id_not x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pushs_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_pushs x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_promote_pending_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_promote_pending x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_mcrlrecord_bounded : 0 <= P_id_mcrlrecord . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_obtainables_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_obtainables x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_cons_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_cons x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_push1_bounded : forall x44 x42 x46 x45 x43 x47, (0 <= x42) -> (0 <= x43) -> (0 <= x44) -> (0 <= x45) -> (0 <= x46) ->(0 <= x47) ->0 <= P_id_push1 x47 x46 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case1_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case1 x45 x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_element_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_element x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_adduniq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_locker2_adduniq x43 x42. 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 x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_and x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_empty_bounded : 0 <= P_id_empty . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_record_updates_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_record_updates x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_lock_bounded : 0 <= P_id_lock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_excllock_bounded : 0 <= P_id_excllock . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pid_bounded : forall x42, (0 <= x42) ->0 <= P_id_pid x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_calls_bounded : forall x44 x42 x43, (0 <= x42) ->(0 <= x43) ->(0 <= x44) ->0 <= P_id_calls x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_subtract_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_subtract x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tag_bounded : 0 <= P_id_tag . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_equal_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_equal x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_eq_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_eq x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tops_bounded : forall x42, (0 <= x42) ->0 <= P_id_tops x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_locker2_claim_lock_bounded : forall x44 x42 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->0 <= P_id_locker2_claim_lock x44 x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_pending_bounded : 0 <= P_id_pending . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_andt_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_andt x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_tuplenil_bounded : forall x42, (0 <= x42) ->0 <= P_id_tuplenil x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_append_bounded : forall x42 x43, (0 <= x42) ->(0 <= x43) ->0 <= P_id_append x43 x42. Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_0_bounded : 0 <= P_id_0 . Proof. intros . cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_case8_bounded : forall x44 x42 x45 x43, (0 <= x42) -> (0 <= x43) ->(0 <= x44) ->(0 <= x45) ->0 <= P_id_case8 x45 x44 x43 x42. 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_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8. Lemma measure_equation : forall t, measure t = match t with | (algebra.Alg.Term algebra.F.id_or (x43::x42::nil)) => P_id_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_a nil) => P_id_a | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43::x42::nil)) => P_id_locker2_release_lock (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43::x42::nil)) => P_id_eqt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43::x42::nil)) => P_id_istops (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43::x42::nil)) => P_id_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_release nil) => P_id_release | (algebra.Alg.Term algebra.F.id_locker2_obtainable (x43:: x42::nil)) => P_id_locker2_obtainable (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43::x42::nil)) => P_id_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_stack (x43::x42::nil)) => P_id_stack (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_locker2_map_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker nil) => P_id_locker | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_int (x42::nil)) => P_id_int (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_push (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43::x42::nil)) => P_id_locker2_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_true nil) => P_id_true | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43::x42::nil)) => P_id_locker2_check_availables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_F nil) => P_id_F | (algebra.Alg.Term algebra.F.id_eqs (x43::x42::nil)) => P_id_eqs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44:: x43::x42::nil)) => P_id_record_update (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_false nil) => P_id_false | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43:: x42::nil)) => P_id_gen_modtageq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_undefined nil) => P_id_undefined | (algebra.Alg.Term algebra.F.id_nocalls nil) => P_id_nocalls | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43::x42::nil)) => P_id_locker2_remove_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_resource nil) => P_id_resource | (algebra.Alg.Term algebra.F.id_case6 (x45::x44::x43:: x42::nil)) => P_id_case6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43::x42::nil)) => P_id_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_pops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43::x42::nil)) => P_id_locker2_map_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_ok nil) => P_id_ok | (algebra.Alg.Term algebra.F.id_case5 (x45::x44::x43:: x42::nil)) => P_id_case5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuple (x43::x42::nil)) => P_id_tuple (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43::x42::nil)) => P_id_member (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_s (x42::nil)) => P_id_s (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43::x42::nil)) => P_id_delete (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_T nil) => P_id_T | (algebra.Alg.Term algebra.F.id_case9 (x45::x44::x43:: x42::nil)) => P_id_case9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44:: x43::x42::nil)) => P_id_record_extract (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_excl nil) => P_id_excl | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_case2 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_nil nil) => P_id_nil | (algebra.Alg.Term algebra.F.id_eqc (x43::x42::nil)) => P_id_eqc (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_case0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_request nil) => P_id_request | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43::x42::nil)) => P_id_locker2_check_available (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_not (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43::x42::nil)) => P_id_pushs (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43::x42::nil)) => P_id_locker2_promote_pending (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_mcrlrecord nil) => P_id_mcrlrecord | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43::x42::nil)) => P_id_locker2_obtainables (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_cons (x43::x42::nil)) => P_id_cons (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46::x45:: x44::x43::x42::nil)) => P_id_push1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44::x43:: x42::nil)) => P_id_case1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43::x42::nil)) => P_id_element (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43:: x42::nil)) => P_id_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43::x42::nil)) => P_id_and (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_empty nil) => P_id_empty | (algebra.Alg.Term algebra.F.id_record_updates (x44:: x43::x42::nil)) => P_id_record_updates (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_lock nil) => P_id_lock | (algebra.Alg.Term algebra.F.id_excllock nil) => P_id_excllock | (algebra.Alg.Term algebra.F.id_pid (x42::nil)) => P_id_pid (measure x42) | (algebra.Alg.Term algebra.F.id_calls (x44::x43:: x42::nil)) => P_id_calls (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43::x42::nil)) => P_id_subtract (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tag nil) => P_id_tag | (algebra.Alg.Term algebra.F.id_equal (x43::x42::nil)) => P_id_equal (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43::x42::nil)) => P_id_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_tops (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_claim_lock (x44:: x43::x42::nil)) => P_id_locker2_claim_lock (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pending nil) => P_id_pending | (algebra.Alg.Term algebra.F.id_andt (x43::x42::nil)) => P_id_andt (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tuplenil (x42::nil)) => P_id_tuplenil (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43::x42::nil)) => P_id_append (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_0 nil) => P_id_0 | (algebra.Alg.Term algebra.F.id_case8 (x45::x44::x43:: x42::nil)) => P_id_case8 (measure x45) (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. Qed. Definition P_id_GEN_MODTAGEQ (x42:Z) (x43:Z) := 0. Definition P_id_ELEMENT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_eq (x42:Z) (x43:Z) := 0. Definition P_id_CASE9 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_REMOVE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_record_new (x42:Z) := 0. Definition P_id_CASE6 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_LOCKER2_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_if (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_MEMBER (x42:Z) (x43:Z) := 0. Definition P_id_CASE5 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_UPDATE (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_Marked_not (x42:Z) := 0. Definition P_id_ISTOPS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_ADD_PENDING (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_or (x42:Z) (x43:Z) := 0. Definition P_id_DELETE (x42:Z) (x43:Z) := 0. Definition P_id_CASE0 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_imp (x42:Z) (x43:Z) := 0. Definition P_id_EQT (x42:Z) (x43:Z) := 0. Definition P_id_PUSHS (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_RELEASE_LOCK (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_OBTAINABLES (x42:Z) (x43:Z) := 0. Definition P_id_RECORD_UPDATES (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_adduniq (x42:Z) (x43:Z) := 0. Definition P_id_EQS (x42:Z) (x43:Z) := 0. Definition P_id_SUBTRACT (x42:Z) (x43:Z) := 0. Definition P_id_Marked_gen_tag (x42:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLES (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_CLAIM_LOCK (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_case4 (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_PUSH1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) (x46:Z) (x47: Z) := 0. Definition P_id_APPEND (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_CHECK_AVAILABLE (x42:Z) (x43:Z) := 0. Definition P_id_LOCKER2_MAP_PROMOTE_PENDING (x42:Z) (x43:Z) := 0. Definition P_id_Marked_pops (x42:Z) := 0. Definition P_id_EQC (x42:Z) (x43:Z) := 1* x42 + 1* x43. Definition P_id_CASE1 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_CASE8 (x42:Z) (x43:Z) (x44:Z) (x45:Z) := 0. Definition P_id_RECORD_EXTRACT (x42:Z) (x43:Z) (x44:Z) := 0. Definition P_id_Marked_locker2_map_add_pending (x42:Z) (x43:Z) (x44: Z) := 0. Definition P_id_AND (x42:Z) (x43:Z) := 0. Definition P_id_Marked_tops (x42:Z) := 0. Definition P_id_CASE2 (x42:Z) (x43:Z) (x44:Z) := 0. Lemma P_id_GEN_MODTAGEQ_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_GEN_MODTAGEQ x43 x45 <= P_id_GEN_MODTAGEQ x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ELEMENT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_ELEMENT x43 x45 <= P_id_ELEMENT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_eq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_eq x43 x45 <= P_id_Marked_eq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE9_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE9 x43 x45 x47 x49 <= P_id_CASE9 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_REMOVE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_REMOVE_PENDING x43 x45 <= P_id_LOCKER2_REMOVE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_record_new_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_record_new x43 <= P_id_Marked_record_new x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE6_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE6 x43 x45 x47 x49 <= P_id_CASE6 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_if_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_if x43 x45 x47 <= P_id_Marked_if x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH x43 x45 x47 <= P_id_PUSH x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_MEMBER_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_MEMBER x43 x45 <= P_id_MEMBER x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE5_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE5 x43 x45 x47 x49 <= P_id_CASE5 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATE_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATE x43 x45 x47 x49 <= P_id_RECORD_UPDATE x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_not_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_not x43 <= P_id_Marked_not x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_ISTOPS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_ISTOPS x43 x45 <= P_id_ISTOPS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_ADD_PENDING_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_ADD_PENDING x43 x45 x47 <= P_id_LOCKER2_ADD_PENDING x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_or_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_or x43 x45 <= P_id_Marked_or x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_DELETE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_DELETE x43 x45 <= P_id_DELETE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE0_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE0 x43 x45 x47 <= P_id_CASE0 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_imp_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_imp x43 x45 <= P_id_Marked_imp x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQT x43 x45 <= P_id_EQT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSHS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_PUSHS x43 x45 <= P_id_PUSHS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_RELEASE_LOCK_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_RELEASE_LOCK x43 x45 <= P_id_LOCKER2_RELEASE_LOCK x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_OBTAINABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_OBTAINABLES x43 x45 <= P_id_LOCKER2_OBTAINABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_UPDATES_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_UPDATES x43 x45 x47 <= P_id_RECORD_UPDATES x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_adduniq_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_adduniq x43 x45 <= P_id_Marked_locker2_adduniq x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQS_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQS x43 x45 <= P_id_EQS x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_SUBTRACT_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_SUBTRACT x43 x45 <= P_id_SUBTRACT x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_gen_tag_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_gen_tag x43 <= P_id_Marked_gen_tag x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLES_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLES x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLES x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_CLAIM_LOCK x43 x45 x47 <= P_id_LOCKER2_MAP_CLAIM_LOCK x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_case4_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_case4 x43 x45 x47 <= P_id_Marked_case4 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_PUSH1_monotonic : forall x48 x52 x44 x50 x42 x46 x49 x53 x45 x51 x43 x47, (0 <= x53)/\ (x53 <= x52) -> (0 <= x51)/\ (x51 <= x50) -> (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_PUSH1 x43 x45 x47 x49 x51 x53 <= P_id_PUSH1 x42 x44 x46 x48 x50 x52. Proof. intros x53 x52 x51 x50 x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. intros [H_9 H_8]. intros [H_11 H_10]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_APPEND_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_APPEND x43 x45 <= P_id_APPEND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_CHECK_AVAILABLE_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_CHECK_AVAILABLE x43 x45 <= P_id_LOCKER2_CHECK_AVAILABLE x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_LOCKER2_MAP_PROMOTE_PENDING x43 x45 <= P_id_LOCKER2_MAP_PROMOTE_PENDING x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_pops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_pops x43 <= P_id_Marked_pops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_EQC_monotonic : forall x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_EQC x43 x45 <= P_id_EQC x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE1_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE1 x43 x45 x47 x49 <= P_id_CASE1 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE8_monotonic : forall x48 x44 x42 x46 x49 x45 x43 x47, (0 <= x49)/\ (x49 <= x48) -> (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE8 x43 x45 x47 x49 <= P_id_CASE8 x42 x44 x46 x48. Proof. intros x49 x48 x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. intros [H_7 H_6]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_RECORD_EXTRACT_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_RECORD_EXTRACT x43 x45 x47 <= P_id_RECORD_EXTRACT x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_locker2_map_add_pending_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_Marked_locker2_map_add_pending x43 x45 x47 <= P_id_Marked_locker2_map_add_pending x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. 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 x44 x42 x45 x43, (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) ->P_id_AND x43 x45 <= P_id_AND x42 x44. Proof. intros x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_Marked_tops_monotonic : forall x42 x43, (0 <= x43)/\ (x43 <= x42) ->P_id_Marked_tops x43 <= P_id_Marked_tops x42. Proof. intros x43 x42. intros [H_1 H_0]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Lemma P_id_CASE2_monotonic : forall x44 x42 x46 x45 x43 x47, (0 <= x47)/\ (x47 <= x46) -> (0 <= x45)/\ (x45 <= x44) -> (0 <= x43)/\ (x43 <= x42) -> P_id_CASE2 x43 x45 x47 <= P_id_CASE2 x42 x44 x46. Proof. intros x47 x46 x45 x44 x43 x42. intros [H_1 H_0]. intros [H_3 H_2]. intros [H_5 H_4]. cbv beta iota zeta delta - [Zle Zlt Zplus Zmult] ;intuition; (auto with zarith)||(repeat (translate_vars );prove_ineq ). Qed. Definition marked_measure := InterpZ.marked_measure 0 P_id_or P_id_gen_tag P_id_record_new P_id_a P_id_locker2_release_lock P_id_eqt P_id_istops P_id_locker2_map_add_pending P_id_release P_id_locker2_obtainable P_id_imp P_id_stack P_id_locker2_map_promote_pending P_id_locker P_id_case4 P_id_int P_id_push P_id_locker2_add_pending P_id_true P_id_locker2_check_availables P_id_F P_id_eqs P_id_record_update P_id_false P_id_gen_modtageq P_id_undefined P_id_nocalls P_id_locker2_remove_pending P_id_resource P_id_case6 P_id_if P_id_pops P_id_locker2_map_claim_lock P_id_ok P_id_case5 P_id_tuple P_id_member P_id_s P_id_delete P_id_T P_id_case9 P_id_record_extract P_id_excl P_id_case2 P_id_nil P_id_eqc P_id_case0 P_id_request P_id_locker2_check_available P_id_not P_id_pushs P_id_locker2_promote_pending P_id_mcrlrecord P_id_locker2_obtainables P_id_cons P_id_push1 P_id_case1 P_id_element P_id_locker2_adduniq P_id_and P_id_empty P_id_record_updates P_id_lock P_id_excllock P_id_pid P_id_calls P_id_subtract P_id_tag P_id_equal P_id_eq P_id_tops P_id_locker2_claim_lock P_id_pending P_id_andt P_id_tuplenil P_id_append P_id_0 P_id_case8 P_id_GEN_MODTAGEQ P_id_ELEMENT P_id_Marked_eq P_id_CASE9 P_id_LOCKER2_REMOVE_PENDING P_id_Marked_record_new P_id_CASE6 P_id_LOCKER2_PROMOTE_PENDING P_id_Marked_if P_id_PUSH P_id_MEMBER P_id_CASE5 P_id_RECORD_UPDATE P_id_Marked_not P_id_ISTOPS P_id_LOCKER2_ADD_PENDING P_id_Marked_or P_id_DELETE P_id_CASE0 P_id_Marked_imp P_id_EQT P_id_PUSHS P_id_LOCKER2_RELEASE_LOCK P_id_LOCKER2_OBTAINABLES P_id_RECORD_UPDATES P_id_Marked_locker2_adduniq P_id_EQS P_id_SUBTRACT P_id_Marked_gen_tag P_id_LOCKER2_CHECK_AVAILABLES P_id_LOCKER2_MAP_CLAIM_LOCK P_id_Marked_case4 P_id_PUSH1 P_id_APPEND P_id_LOCKER2_CHECK_AVAILABLE P_id_LOCKER2_MAP_PROMOTE_PENDING P_id_Marked_pops P_id_EQC P_id_CASE1 P_id_CASE8 P_id_RECORD_EXTRACT P_id_Marked_locker2_map_add_pending P_id_AND P_id_Marked_tops P_id_CASE2. Lemma marked_measure_equation : forall t, marked_measure t = match t with | (algebra.Alg.Term algebra.F.id_gen_modtageq (x43::x42::nil)) => P_id_GEN_MODTAGEQ (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_element (x43:: x42::nil)) => P_id_ELEMENT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eq (x43:: x42::nil)) => P_id_Marked_eq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case9 (x45::x44:: x43::x42::nil)) => P_id_CASE9 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_remove_pending (x43:: x42::nil)) => P_id_LOCKER2_REMOVE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_new (x42::nil)) => P_id_Marked_record_new (measure x42) | (algebra.Alg.Term algebra.F.id_case6 (x45::x44:: x43::x42::nil)) => P_id_CASE6 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_if (x44::x43:: x42::nil)) => P_id_Marked_if (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push (x44::x43:: x42::nil)) => P_id_PUSH (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_member (x43:: x42::nil)) => P_id_MEMBER (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case5 (x45::x44:: x43::x42::nil)) => P_id_CASE5 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_update (x45::x44::x43::x42::nil)) => P_id_RECORD_UPDATE (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_not (x42::nil)) => P_id_Marked_not (measure x42) | (algebra.Alg.Term algebra.F.id_istops (x43:: x42::nil)) => P_id_ISTOPS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_add_pending (x44::x43:: x42::nil)) => P_id_LOCKER2_ADD_PENDING (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_or (x43:: x42::nil)) => P_id_Marked_or (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_delete (x43:: x42::nil)) => P_id_DELETE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case0 (x44::x43:: x42::nil)) => P_id_CASE0 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_imp (x43:: x42::nil)) => P_id_Marked_imp (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqt (x43:: x42::nil)) => P_id_EQT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pushs (x43:: x42::nil)) => P_id_PUSHS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_release_lock (x43:: x42::nil)) => P_id_LOCKER2_RELEASE_LOCK (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_obtainables (x43:: x42::nil)) => P_id_LOCKER2_OBTAINABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_updates (x44::x43::x42::nil)) => P_id_RECORD_UPDATES (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_adduniq (x43::x42::nil)) => P_id_Marked_locker2_adduniq (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_eqs (x43:: x42::nil)) => P_id_EQS (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_subtract (x43:: x42::nil)) => P_id_SUBTRACT (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_gen_tag (x42::nil)) => P_id_Marked_gen_tag (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_availables (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLES (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_claim_lock (x44::x43:: x42::nil)) => P_id_LOCKER2_MAP_CLAIM_LOCK (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case4 (x44::x43:: x42::nil)) => P_id_Marked_case4 (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_push1 (x47::x46:: x45::x44::x43::x42::nil)) => P_id_PUSH1 (measure x47) (measure x46) (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_append (x43:: x42::nil)) => P_id_APPEND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_check_available (x43:: x42::nil)) => P_id_LOCKER2_CHECK_AVAILABLE (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_promote_pending (x43:: x42::nil)) => P_id_LOCKER2_MAP_PROMOTE_PENDING (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_pops (x42::nil)) => P_id_Marked_pops (measure x42) | (algebra.Alg.Term algebra.F.id_eqc (x43:: x42::nil)) => P_id_EQC (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case1 (x45::x44:: x43::x42::nil)) => P_id_CASE1 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_case8 (x45::x44:: x43::x42::nil)) => P_id_CASE8 (measure x45) (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_record_extract (x44::x43::x42::nil)) => P_id_RECORD_EXTRACT (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_locker2_map_add_pending (x44::x43:: x42::nil)) => P_id_Marked_locker2_map_add_pending (measure x44) (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_and (x43:: x42::nil)) => P_id_AND (measure x43) (measure x42) | (algebra.Alg.Term algebra.F.id_tops (x42::nil)) => P_id_Marked_tops (measure x42) | (algebra.Alg.Term algebra.F.id_case2 (x44::x43:: x42::nil)) => P_id_CASE2 (measure x44) (measure x43) (measure x42) | _ => 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_or_monotonic;assumption. intros ;apply P_id_gen_tag_monotonic;assumption. intros ;apply P_id_record_new_monotonic;assumption. intros ;apply P_id_locker2_release_lock_monotonic;assumption. intros ;apply P_id_eqt_monotonic;assumption. intros ;apply P_id_istops_monotonic;assumption. intros ;apply P_id_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainable_monotonic;assumption. intros ;apply P_id_imp_monotonic;assumption. intros ;apply P_id_stack_monotonic;assumption. intros ;apply P_id_locker2_map_promote_pending_monotonic;assumption. intros ;apply P_id_case4_monotonic;assumption. intros ;apply P_id_int_monotonic;assumption. intros ;apply P_id_push_monotonic;assumption. intros ;apply P_id_locker2_add_pending_monotonic;assumption. intros ;apply P_id_locker2_check_availables_monotonic;assumption. intros ;apply P_id_eqs_monotonic;assumption. intros ;apply P_id_record_update_monotonic;assumption. intros ;apply P_id_gen_modtageq_monotonic;assumption. intros ;apply P_id_locker2_remove_pending_monotonic;assumption. intros ;apply P_id_case6_monotonic;assumption. intros ;apply P_id_if_monotonic;assumption. intros ;apply P_id_pops_monotonic;assumption. intros ;apply P_id_locker2_map_claim_lock_monotonic;assumption. intros ;apply P_id_case5_monotonic;assumption. intros ;apply P_id_tuple_monotonic;assumption. intros ;apply P_id_member_monotonic;assumption. intros ;apply P_id_s_monotonic;assumption. intros ;apply P_id_delete_monotonic;assumption. intros ;apply P_id_case9_monotonic;assumption. intros ;apply P_id_record_extract_monotonic;assumption. intros ;apply P_id_case2_monotonic;assumption. intros ;apply P_id_eqc_monotonic;assumption. intros ;apply P_id_case0_monotonic;assumption. intros ;apply P_id_locker2_check_available_monotonic;assumption. intros ;apply P_id_not_monotonic;assumption. intros ;apply P_id_pushs_monotonic;assumption. intros ;apply P_id_locker2_promote_pending_monotonic;assumption. intros ;apply P_id_locker2_obtainables_monotonic;assumption. intros ;apply P_id_cons_monotonic;assumption. intros ;apply P_id_push1_monotonic;assumption. intros ;apply P_id_case1_monotonic;assumption. intros ;apply P_id_element_monotonic;assumption. intros ;apply P_id_locker2_adduniq_monotonic;assumption. intros ;apply P_id_and_monotonic;assumption. intros ;apply P_id_record_updates_monotonic;assumption. intros ;apply P_id_pid_monotonic;assumption. intros ;apply P_id_calls_monotonic;assumption. intros ;apply P_id_subtract_monotonic;assumption. intros ;apply P_id_equal_monotonic;assumption. intros ;apply P_id_eq_monotonic;assumption. intros ;apply P_id_tops_monotonic;assumption. intros ;apply P_id_locker2_claim_lock_monotonic;assumption. intros ;apply P_id_andt_monotonic;assumption. intros ;apply P_id_tuplenil_monotonic;assumption. intros ;apply P_id_append_monotonic;assumption. intros ;apply P_id_case8_monotonic;assumption. intros ;apply P_id_or_bounded;assumption. intros ;apply P_id_gen_tag_bounded;assumption. intros ;apply P_id_record_new_bounded;assumption. intros ;apply P_id_a_bounded;assumption. intros ;apply P_id_locker2_release_lock_bounded;assumption. intros ;apply P_id_eqt_bounded;assumption. intros ;apply P_id_istops_bounded;assumption. intros ;apply P_id_locker2_map_add_pending_bounded;assumption. intros ;apply P_id_release_bounded;assumption. intros ;apply P_id_locker2_obtainable_bounded;assumption. intros ;apply P_id_imp_bounded;assumption. intros ;apply P_id_stack_bounded;assumption. intros ;apply P_id_locker2_map_promote_pending_bounded;assumption. intros ;apply P_id_locker_bounded;assumption. intros ;apply P_id_case4_bounded;assumption. intros ;apply P_id_int_bounded;assumption. intros ;apply P_id_push_bounded;assumption. intros ;apply P_id_locker2_add_pending_bounded;assumption. intros ;apply P_id_true_bounded;assumption. intros ;apply P_id_locker2_check_availables_bounded;assumption. intros ;apply P_id_F_bounded;assumption. intros ;apply P_id_eqs_bounded;assumption. intros ;apply P_id_record_update_bounded;assumption. intros ;apply P_id_false_bounded;assumption. intros ;apply P_id_gen_modtageq_bounded;assumption. intros ;apply P_id_undefined_bounded;assumption. intros ;apply P_id_nocalls_bounded;assumption. intros ;apply P_id_locker2_remove_pending_bounded;assumption. intros ;apply P_id_resource_bounded;assumption. intros ;apply P_id_case6_bounded;assumption. intros ;apply P_id_if_bounded;assumption. intros ;apply P_id_pops_bounded;assumption. intros ;apply P_id_locker2_map_claim_lock_bounded;assumption. intros ;apply P_id_ok_bounded;assumption. intros ;apply P_id_case5_bounded;assumption. intros ;apply P_id_tuple_bounded;assumption. intros ;apply P_id_member_bounded;assumption. intros ;apply P_id_s_bounded;assumption. intros ;apply P_id_delete_bounded;assumption. intros ;apply P_id_T_bounded;assumption. intros ;apply P_id_case9_bounded;assumption. intros ;apply P_id_record_extract_bounded;assumption. intros ;apply P_id_excl_bounded;assumption. intros ;apply P_id_case2_bounded;assumption. intros ;apply P_id_nil_bounded;assumption. intros ;apply P_id_eqc_bounded;assumption. intros ;apply P_id_case0_bounded;assumption. intros ;apply P_id_request_bounded;assumption. intros ;apply P_id_locker2_check_available_bounded;assumption. intros ;apply P_id_not_bounded;assumption. intros ;apply P_id_pushs_bounded;assumption. intros ;apply P_id_locker2_promote_pending_bounded;assumption. intros ;apply P_id_mcrlrecord_bounded;assumption. intros ;apply P_id_locker2_obtainables_bounded;assumption. intros ;apply P_id_cons_bounded;assumption. intros ;apply P_id_push1_bounded;assumption. intros ;apply P_id_case1_bounded;assumption. intros ;apply P_id_element_bounded;assumption. intros ;apply P_id_locker2_adduniq_bounded;assumption. intros ;apply P_id_and_bounded;assumption. intros ;apply P_id_empty_bounded;assumption. intros ;apply P_id_record_updates_bounded;assumption. intros ;apply P_id_lock_bounded;assumption. intros ;apply P_id_excllock_bounded;assumption. intros ;apply P_id_pid_bounded;assumption. intros ;apply P_id_calls_bounded;assumption. intros ;apply P_id_subtract_bounded;assumption. intros ;apply P_id_tag_bounded;assumption. intros ;apply P_id_equal_bounded;assumption. intros ;apply P_id_eq_bounded;assumption. intros ;apply P_id_tops_bounded;assumption. intros ;apply P_id_locker2_claim_lock_bounded;assumption. intros ;apply P_id_pending_bounded;assumption. intros ;apply P_id_andt_bounded;assumption. intros ;apply P_id_tuplenil_bounded;assumption. intros ;apply P_id_append_bounded;assumption. intros ;apply P_id_0_bounded;assumption. intros ;apply P_id_case8_bounded;assumption. apply rules_monotonic. intros ;apply P_id_GEN_MODTAGEQ_monotonic;assumption. intros ;apply P_id_ELEMENT_monotonic;assumption. intros ;apply P_id_Marked_eq_monotonic;assumption. intros ;apply P_id_CASE9_monotonic;assumption. intros ;apply P_id_LOCKER2_REMOVE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_record_new_monotonic;assumption. intros ;apply P_id_CASE6_monotonic;assumption. intros ;apply P_id_LOCKER2_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_if_monotonic;assumption. intros ;apply P_id_PUSH_monotonic;assumption. intros ;apply P_id_MEMBER_monotonic;assumption. intros ;apply P_id_CASE5_monotonic;assumption. intros ;apply P_id_RECORD_UPDATE_monotonic;assumption. intros ;apply P_id_Marked_not_monotonic;assumption. intros ;apply P_id_ISTOPS_monotonic;assumption. intros ;apply P_id_LOCKER2_ADD_PENDING_monotonic;assumption. intros ;apply P_id_Marked_or_monotonic;assumption. intros ;apply P_id_DELETE_monotonic;assumption. intros ;apply P_id_CASE0_monotonic;assumption. intros ;apply P_id_Marked_imp_monotonic;assumption. intros ;apply P_id_EQT_monotonic;assumption. intros ;apply P_id_PUSHS_monotonic;assumption. intros ;apply P_id_LOCKER2_RELEASE_LOCK_monotonic;assumption. intros ;apply P_id_LOCKER2_OBTAINABLES_monotonic;assumption. intros ;apply P_id_RECORD_UPDATES_monotonic;assumption. intros ;apply P_id_Marked_locker2_adduniq_monotonic;assumption. intros ;apply P_id_EQS_monotonic;assumption. intros ;apply P_id_SUBTRACT_monotonic;assumption. intros ;apply P_id_Marked_gen_tag_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLES_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_CLAIM_LOCK_monotonic;assumption. intros ;apply P_id_Marked_case4_monotonic;assumption. intros ;apply P_id_PUSH1_monotonic;assumption. intros ;apply P_id_APPEND_monotonic;assumption. intros ;apply P_id_LOCKER2_CHECK_AVAILABLE_monotonic;assumption. intros ;apply P_id_LOCKER2_MAP_PROMOTE_PENDING_monotonic;assumption. intros ;apply P_id_Marked_pops_monotonic;assumption. intros ;apply P_id_EQC_monotonic;assumption. intros ;apply P_id_CASE1_monotonic;assumption. intros ;apply P_id_CASE8_monotonic;assumption. intros ;apply P_id_RECORD_EXTRACT_monotonic;assumption. intros ;apply P_id_Marked_locker2_map_add_pending_monotonic;assumption. intros ;apply P_id_AND_monotonic;assumption. intros ;apply P_id_Marked_tops_monotonic;assumption. intros ;apply P_id_CASE2_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_54. 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_54. Definition wf_DP_R_xml_0_scc_54 := WF_DP_R_xml_0_scc_54.wf. Lemma acc_DP_R_xml_0_scc_54 : forall x y, (DP_R_xml_0_scc_54 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_54). 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_53; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_49; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_4; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_3; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| (eapply Hrec; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))))))). apply wf_DP_R_xml_0_scc_54. 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_53; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_52; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_51; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_50; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_49; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_48; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_47; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_46; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_45; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_44; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_43; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_42; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_41; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_40; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_39; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_38; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_37; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_36; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_35; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_34; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_33; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_32; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_31; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_30; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_29; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_28; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_27; econstructor (eassumption)||(algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_26; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_25; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_24; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_23; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_22; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_21; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_20; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_19; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_18; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_17; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_16; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_15; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_14; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_13; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_12; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_11; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_10; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_9; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_8; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_7; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_6; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_5; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_4; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_3; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_2; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_1; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_non_scc_0; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_54; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_53; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_52; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_51; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_50; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_49; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_48; econstructor (eassumption)|| (algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_47; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ((eapply acc_DP_R_xml_0_scc_46; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| (( eapply acc_DP_R_xml_0_scc_45; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_44; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_43; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_42; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_41; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_40; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_39; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_38; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_37; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_36; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_35; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_34; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_33; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_32; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_31; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_30; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_29; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_28; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_27; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_26; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_25; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_24; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_23; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_22; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_21; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_20; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_19; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_18; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_17; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_16; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_15; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_14; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_13; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_12; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_11; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_10; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_9; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_8; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_7; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_6; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_5; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_4; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_3; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_2; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_1; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( eapply acc_DP_R_xml_0_scc_0; econstructor ( eassumption)|| ( algebra.Alg_ext.star_refl' ))|| ( ( R_xml_0_deep_rew.impossible_star_reduction_R_xml_0 )|| ( fail)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))). Qed. End WF_DP_R_xml_0. Definition wf_H := WF_DP_R_xml_0.wf. Lemma wf : well_founded (algebra.EQT.one_step R_xml_0_deep_rew.R_xml_0_rules). Proof. apply ddp.dp_criterion. apply R_xml_0_deep_rew.R_xml_0_non_var. apply R_xml_0_deep_rew.R_xml_0_reg. intros ; apply (ddp.constructor_defined_dec _ _ R_xml_0_deep_rew.R_xml_0_rules_included). refine (Inclusion.wf_incl _ _ _ _ wf_H). intros x y H. destruct (R_xml_0_dp_step_spec H) as [f [l1 [l2 [H1 [H2 H3]]]]]. destruct (ddp.dp_list_complete _ _ R_xml_0_deep_rew.R_xml_0_rules_included _ _ H3) as [x' [y' [sigma [h1 [h2 h3]]]]]. clear H3. subst. vm_compute in h3|-. let e := type of h3 in (dp_concl_tac h2 h3 ltac:(fun _ => idtac) e). Qed. End WF_R_xml_0_deep_rew. (* *** Local Variables: *** *** coq-prog-name: "coqtop" *** *** coq-prog-args: ("-emacs-U" "-I" "$COCCINELLE/examples" "-I" "$COCCINELLE/term_algebra" "-I" "$COCCINELLE/term_orderings" "-I" "$COCCINELLE/basis" "-I" "$COCCINELLE/list_extensions" "-I" "$COCCINELLE/examples/cime_trace/") *** *** compile-command: "coqc -I $COCCINELLE/term_algebra -I $COCCINELLE/term_orderings -I $COCCINELLE/basis -I $COCCINELLE/list_extensions -I $COCCINELLE/examples/cime_trace/ -I $COCCINELLE/examples/ c_output/strat/tpdb-5.0___TRS___Cime___mucrl1.trs/a3pat.v" *** *** End: *** *)