diff --git a/thys/Lambda_Free_EPO/Lambda_Free_EPO.thy b/thys/Lambda_Free_EPO/Lambda_Free_EPO.thy --- a/thys/Lambda_Free_EPO/Lambda_Free_EPO.thy +++ b/thys/Lambda_Free_EPO/Lambda_Free_EPO.thy @@ -1,1247 +1,1232 @@ (* Title: The Embedding Path Order for Lambda-Free Higher-Order Terms Author: Alexander Bentkamp , 2018 Maintainer: Alexander Bentkamp *) section \The Embedding Path Order for Lambda-Free Higher-Order Terms\ theory Lambda_Free_EPO imports Chop abbrevs ">t" = ">\<^sub>t" and "\t" = "\\<^sub>t" begin text \ This theory defines the embedding path order for \\\-free higher-order terms. \ subsection \Setup\ locale epo = ground_heads "(>\<^sub>s)" arity_sym arity_var for gt_sym :: "'s \ 's \ bool" (infix ">\<^sub>s" 50) and arity_sym :: "'s \ enat" and arity_var :: "'v \ enat" + fixes extf :: "'s \ (('s, 'v) tm \ ('s, 'v) tm \ bool) \ ('s, 'v) tm list \ ('s, 'v) tm list \ bool" assumes extf_ext_trans_before_irrefl: "ext_trans_before_irrefl (extf f)" and extf_ext_compat_list: "ext_compat_list (extf f)" assumes extf_ext_compat_snoc: "ext_compat_snoc (extf f)" assumes extf_ext_compat_cons: "ext_compat_cons (extf f)" assumes extf_ext_snoc: "ext_snoc (extf f)" assumes extf_min_empty: "\ extf f gt [] ss" (* TODO: seperate definition? *) begin lemma extf_ext_trans: "ext_trans (extf f)" by (rule ext_trans_before_irrefl.axioms(1)[OF extf_ext_trans_before_irrefl]) lemma extf_ext: "ext (extf f)" by (rule ext_trans.axioms(1)[OF extf_ext_trans]) lemmas extf_mono_strong = ext.mono_strong[OF extf_ext] lemmas extf_mono = ext.mono[OF extf_ext, mono] lemmas extf_map = ext.map[OF extf_ext] lemmas extf_trans = ext_trans.trans[OF extf_ext_trans] lemmas extf_irrefl_from_trans = ext_trans_before_irrefl.irrefl_from_trans[OF extf_ext_trans_before_irrefl] lemmas extf_compat_list = ext_compat_list.compat_list[OF extf_ext_compat_list] lemmas extf_snoc = ext_snoc.snoc[OF extf_ext_snoc] lemmas extf_compat_append_right = ext_compat_snoc.compat_append_right[OF extf_ext_compat_snoc] lemmas extf_compat_append_left = ext_compat_cons.compat_append_left[OF extf_ext_compat_cons] lemma extf_ext_insert_arg: "extf f gt (xs @ z # ys) (xs @ ys)" using extf_compat_append_left extf_compat_append_right extf_snoc[of f gt Nil z] by fastforce subsection \Inductive Definitions\ definition chkchop :: "(('s, 'v) tm \ ('s, 'v) tm \ bool) \ ('s, 'v) tm \ ('s, 'v) tm \ bool" where [simp]: "chkchop gt t s \ is_Hd s \ gt t (chop s)" definition chkchop_same :: "(('s, 'v) tm \ ('s, 'v) tm \ bool) \ ('s, 'v) tm \ ('s, 'v) tm \ bool" where [simp]: "chkchop_same gt t s \ (if is_Var (head t) then is_Hd t \ chkchop gt (chop t) s else chkchop gt t s)" lemma chkchop_mono[mono]: "gt \ gt' \ chkchop gt \ chkchop gt'" using chkchop_def by blast lemma chkchop_same_mono[mono]: "gt \ gt' \ chkchop_same gt \ chkchop_same gt'" using chkchop_same_def by fastforce inductive gt :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix ">\<^sub>t" 50) where gt_chop: "is_App t \ chop t >\<^sub>t s \ chop t = s \ t >\<^sub>t s" | gt_diff: "head t >\<^sub>h\<^sub>d head s \ is_Sym (head s) \ chkchop (>\<^sub>t) t s \ t >\<^sub>t s" | gt_same: "head t = head s \ chkchop_same (>\<^sub>t) t s \ (\f \ ground_heads (head t). extf f (>\<^sub>t) (args t) (args s)) \ t >\<^sub>t s" abbreviation ge :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix "\\<^sub>t" 50) where "t \\<^sub>t s \ t >\<^sub>t s \ t = s" inductive gt_chop :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_chopI: "is_App t \ chop t \\<^sub>t s \ gt_chop t s" inductive gt_diff :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_diffI: "head t >\<^sub>h\<^sub>d head s \ is_Sym (head s) \ chkchop (>\<^sub>t) t s \ gt_diff t s" inductive gt_same :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_sameI: "head t = head s \ chkchop_same (>\<^sub>t) t s \ (\f \ ground_heads (head t). extf f (>\<^sub>t) (args t) (args s)) \ gt_same t s" lemma gt_iff_chop_diff_same: "t >\<^sub>t s \ gt_chop t s \ gt_diff t s \ gt_same t s" by (subst gt.simps) (auto simp: gt_chop.simps gt_diff.simps gt_same.simps) subsection \Transitivity\ lemma t_gt_chop_t: "is_App t \ t >\<^sub>t chop t" by (simp add: gt_chop) lemma gt_imp_vars: "t >\<^sub>t s \ vars t \ vars s" proof (simp only: atomize_imp, rule measure_induct_rule[of "\(t, s). hsize t + hsize s" "\(t, s). t >\<^sub>t s \ vars t \ vars s" "(t, s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric]) fix t s assume ih: "\ta sa. hsize ta + hsize sa < hsize t + hsize s \ ta >\<^sub>t sa \ vars ta \ vars sa" and t_gt_s: "t >\<^sub>t s" show "vars t \ vars s" using t_gt_s proof cases case (gt_chop) thus ?thesis using ih by (metis add_mono_thms_linordered_field(1) le_supI1 order_refl hsize_chop_lt vars_chop) next case gt_diff show ?thesis proof (cases s) case Hd thus ?thesis using gt_diff(2) by (metis empty_iff hd.collapse(2) hd.simps(18) subsetI tm.sel(1) tm.simps(17)) next case (App s1 s2) have "vars (chop s) \ vars t" using ih using App chkchop_def local.gt_diff(3) nat_add_left_cancel_less hsize_chop_lt tm.disc(2) by blast thus ?thesis using App le_sup_iff local.gt_diff(2) tm.disc(2) vars_chop by (metis empty_iff hd.collapse(2) hd.simps(18) subsetI) qed next case gt_same thus ?thesis proof (cases "head t") case (Var x) then show ?thesis proof (cases t) case (Hd _) then show ?thesis using gt_same extf_min_empty[of _ "(>\<^sub>t)" "args s"] by simp next case (App t1 t2) then show ?thesis proof (cases s) case (Hd _) then show ?thesis using local.gt_same(1) vars_head_subseteq by fastforce next case (App s1 s2) then have "chop t >\<^sub>t chop s" by (metis Var args.simps(1) chkchop_def chkchop_same_def epo.extf_min_empty epo_axioms gt_hd_def gt_hd_irrefl hd.disc(1) local.gt_same(2) local.gt_same(3) tm.collapse(1) tm.disc(2)) then have "vars (chop s) \ vars (chop t)" using ih[OF _ \chop t >\<^sub>t chop s\] by (metis App add_mono_thms_linordered_field(5) args_Nil_iff_is_Hd extf_min_empty gt_hd_def gt_hd_irrefl local.gt_same(3) hsize_chop_lt tm.disc(2)) then show ?thesis using gt_same(1) vars_chop[of t] vars_chop[of s] by (metis App args_Nil_iff_is_Hd extf_min_empty gt_hd_def gt_hd_irrefl le_sup_iff local.gt_same(3) order_refl sup.coboundedI1 tm.disc(2)) qed qed next case (Sym f) then have "chkchop (>\<^sub>t) t s" using gt_same chkchop_same_def by auto then show ?thesis proof (cases s) case (Hd _) then show ?thesis using local.gt_same(1) vars_head_subseteq by force next case (App s1 s2) then show ?thesis unfolding chkchop_def using vars_chop ih[of t "chop s"] by (metis \chkchop (>\<^sub>t) t s\ chkchop_def le_sup_iff local.gt_same(1) nat_add_left_cancel_less hsize_chop_lt tm.disc(2) vars_head_subseteq) qed qed qed qed lemma gt_trans: "u >\<^sub>t t \ t >\<^sub>t s \ u >\<^sub>t s" proof (simp only: atomize_imp, rule measure_induct_rule[of "\(u, t, s). {#hsize u, hsize t, hsize s#}" "\(u, t, s). u >\<^sub>t t \ t >\<^sub>t s \ u >\<^sub>t s" "(u, t, s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric]) fix u t s assume ih: "\ua ta sa. {#hsize ua, hsize ta, hsize sa#} < {#hsize u, hsize t, hsize s#} \ ua >\<^sub>t ta \ ta >\<^sub>t sa \ ua >\<^sub>t sa" and u_gt_t: "u >\<^sub>t t" and t_gt_s: "t >\<^sub>t s" have u_gt_s_if_ui: "chop u \\<^sub>t t \ u >\<^sub>t s" if ui_in: "is_App u" using ih[of "chop u" t s, simplified] t_gt_s gt_chop hsize_chop_lt ui_in by blast show "u >\<^sub>t s" using t_gt_s proof cases case gt_chop have u_gt_s_if_chk_u_t: ?thesis if chk_u_t: "chkchop (>\<^sub>t) u t" using ih[of u "chop u" s] gt_chop chk_u_t by (metis add_mset_lt_left_lt add_mset_lt_right_lt chkchop_def ih hsize_chop_lt) show ?thesis by (metis args_Nil_iff_is_Hd chkchop_def chkchop_same_def epo.extf_min_empty epo_axioms gt.simps gt_hd_def gt_hd_irrefl u_gt_s_if_chk_u_t u_gt_s_if_ui u_gt_t) next case gt_diff_t_s: gt_diff show ?thesis using u_gt_t proof cases case gt_chop then show ?thesis using u_gt_s_if_ui by blast next case gt_diff_u_t: gt_diff have "head u >\<^sub>h\<^sub>d head s" using gt_diff_u_t(1) gt_diff_t_s(1) by (auto intro: gt_hd_trans) thus ?thesis using add_mset_lt_left_lt add_mset_lt_right_lt chkchop_def gt_diff gt_diff_t_s(3) ih hsize_chop_lt u_gt_t by (metis gt_diff_t_s(2)) next case gt_same_u_t: gt_same have "head u >\<^sub>h\<^sub>d head s" using gt_diff_t_s(1) gt_same_u_t(1) by auto thus ?thesis using add_mset_lt_left_lt add_mset_lt_right_lt chkchop_def gt_diff gt_diff_t_s(3) ih hsize_chop_lt u_gt_t by (metis gt_diff_t_s(2)) qed next case gt_same_t_s: gt_same show ?thesis using u_gt_t proof cases case gt_chop then show ?thesis using u_gt_s_if_ui by linarith next case gt_diff_u_t: gt_diff have "head u >\<^sub>h\<^sub>d head s" using gt_diff_u_t(1) gt_same_t_s(1) by simp thus ?thesis using add_mset_lt_left_lt add_mset_lt_right_lt chkchop_def gt_diff gt_same_t_s ih hsize_chop_lt u_gt_t by (metis chkchop_same_def gt_diff_u_t(2)) next case gt_same_u_t: gt_same have hd_u_s: "head u = head s" using gt_same_u_t(1) gt_same_t_s(1) by simp let ?S = "set (args u) \ set (args t) \ set (args s)" have gt_trans_args: "\ua \ ?S. \ta \ ?S. \sa \ ?S. ua >\<^sub>t ta \ ta >\<^sub>t sa \ ua >\<^sub>t sa" proof clarify fix sa ta ua assume ua_in: "ua \ ?S" and ta_in: "ta \ ?S" and sa_in: "sa \ ?S" and ua_gt_ta: "ua >\<^sub>t ta" and ta_gt_sa: "ta >\<^sub>t sa" show "ua >\<^sub>t sa" by (auto intro!: ih[OF Max_lt_imp_lt_mset ua_gt_ta ta_gt_sa]) (meson ua_in ta_in sa_in Un_iff max.strict_coboundedI1 max.strict_coboundedI2 hsize_in_args)+ qed have "\f \ ground_heads (head u). extf f (>\<^sub>t) (args u) (args s)" proof (clarify, rule extf_trans[OF _ _ _ gt_trans_args]) fix f assume f_in_grounds: "f \ ground_heads (head u)" show "extf f (>\<^sub>t) (args u) (args t)" using f_in_grounds gt_same_u_t(3) by blast show "extf f (>\<^sub>t) (args t) (args s)" using f_in_grounds gt_same_t_s(3) unfolding gt_same_u_t(1) by blast qed auto have "chkchop_same (>\<^sub>t) u s" proof (cases "head u") case (Var x) then show ?thesis proof (cases u) case (Hd _) then show ?thesis using Var by auto next case (App u1 u2) then have "chop u >\<^sub>t chop t" by (metis Var args.simps(1) chkchop_def chkchop_same_def epo.extf_min_empty epo_axioms gt_hd_def gt_hd_irrefl gt_same_t_s(3) gt_same_u_t(2) hd.disc(1) tm.collapse(1) tm.disc(2)) then show ?thesis proof (cases t) case (Hd _) then show ?thesis using extf_min_empty gt_same_t_s(3) by auto next case t_App: (App t1 t2) then have "is_App s \ chop t >\<^sub>t chop s" using gt_same_t_s unfolding chkchop_same_def unfolding chkchop_def using Var hd_u_s by auto then have "chkchop (>\<^sub>t) (chop u) s" unfolding chkchop_def using ih[of "chop u" "chop t" "chop s"] by (metis App \chop u >\<^sub>t chop t\ t_App add_mset_lt_lt_le less_imp_le mset_lt_single_iff hsize_chop_lt tm.disc(2)) then show ?thesis unfolding chkchop_same_def using Var by auto qed qed next case (Sym f) have "chkchop (>\<^sub>t) u s" proof (cases s) case (Hd _) then show ?thesis by simp next case (App s1 s2) then have "t >\<^sub>t chop s" using Sym gt_same_t_s(1) gt_same_t_s(2) hd_u_s by auto then have "u >\<^sub>t chop s" using ih[of u t "chop s"] by (metis App add_mset_lt_right_lt mset_lt_single_iff hsize_chop_lt tm.disc(2) u_gt_t) then show ?thesis unfolding chkchop_def by blast qed then show ?thesis by (simp add: Sym) qed thus ?thesis using \\f\local.ground_heads (head u). extf f (>\<^sub>t) (args u) (args s)\ gt_same hd_u_s by blast qed qed qed subsection \Irreflexivity\ theorem gt_irrefl: "\ s >\<^sub>t s" proof (standard, induct s rule: measure_induct_rule[of hsize]) case (less s) note ih = this(1) and s_gt_s = this(2) show False using s_gt_s proof cases case gt_chop then show False using ih[of "chop s"] by (metis gt.gt_chop gt_trans hsize_chop_lt) next case gt_diff thus False by (cases "head s") (auto simp: gt_hd_irrefl) next case gt_same note in_grounds = this(3) obtain si where si_in_args: "si \ set (args s)" and si_gt_si: "si >\<^sub>t si" using in_grounds by (metis (full_types) all_not_in_conv extf_irrefl_from_trans ground_heads_nonempty gt_trans) have "hsize si < hsize s" by (rule hsize_in_args[OF si_in_args]) thus False by (rule ih[OF _ si_gt_si]) qed qed lemma gt_antisym: "t >\<^sub>t s \ \ s >\<^sub>t t" using gt_irrefl gt_trans by blast subsection "Compatibility with Embedding Relation" (* TODO: move? *) lemma nth_drop_lemma: assumes "length xs = length ys" and "k \ length xs" and "\i. i < length xs \ i \ k \ xs ! i = ys ! i" shows "drop k xs = drop k ys" using assms proof (induct arbitrary:k rule:list_induct2) case Nil then show ?case by simp next case (Cons x xs y ys) then show ?case proof (cases k) case 0 then have "x # xs = y # ys" by (metis Cons.hyps(1) Cons.prems(2) leI length_Cons not_less_zero nth_equalityI) then show ?thesis by blast next case (Suc m) then have "drop m xs = drop m ys" by (metis Cons.hyps(2) Cons.prems(1) Cons.prems(2) Suc_le_mono Suc_mono length_Cons nth_Cons_Suc) then show ?thesis by (simp add: Suc) qed qed lemma gt_embedding_step_property: assumes "t \\<^sub>e\<^sub>m\<^sub>b s" shows "t >\<^sub>t s" using assms apply(simp only: atomize_imp) apply (rule measure_induct_rule[of "\(t, s). hsize t + hsize s" "\(t, s). t \\<^sub>e\<^sub>m\<^sub>b s \ t >\<^sub>t s" "(t, s)", simplified prod.case]) proof(simp only: split_paired_all prod.case atomize_imp[symmetric]) fix s t :: "('s,'v) tm" assume "t \\<^sub>e\<^sub>m\<^sub>b s" and ih: "\tt ss. hsize tt + hsize ss < hsize t + hsize s \ tt \\<^sub>e\<^sub>m\<^sub>b ss \ tt >\<^sub>t ss" have "is_App t" by (metis \t \\<^sub>e\<^sub>m\<^sub>b s\ emb_step_at_is_App emb_step_equiv) obtain p d where "emb_step_at p d t = s" "position_of t (p @ [d])" using \t \\<^sub>e\<^sub>m\<^sub>b s\ emb_step_equiv by (metis position_if_emb_step_at) define q where q_rep_t: "q = replicate (num_args (fun t)) Left" show "t >\<^sub>t s" proof (cases "list_all (\x. x = Left) p") case True show ?thesis proof (cases d) text \Embedding removes an argument i\ case Left define i where "i = num_args t - Suc (length p)" then have " head t = head s" "i < num_args t" "args s = take i (args t) @ drop (Suc i) (args t)" using emb_step_at_remove_arg Left True \emb_step_at p d t = s\ \position_of t (p @ [d])\ by metis+ have "is_App s \ chop t \\<^sub>e\<^sub>m\<^sub>b chop s" proof (cases "p = q") case True assume "is_App s" show ?thesis proof - have "Suc (num_args (fun s)) = num_args (fun t)" by (metis One_nat_def Suc_num_args True \args s = take i (args t) @ drop (Suc i) (args t)\ \is_App s\ \is_App t\ append_self_conv2 cancel_comm_monoid_add_class.diff_cancel diff_Suc_1 i_def length_drop length_replicate q_rep_t take_eq_Nil) have "emb_step_at (replicate (num_args (fun s)) Left) Right (chop t) = emb_step_at (replicate (num_args (fun s)) Left) Right t" using merge_emb_step_at[of "replicate (num_args (fun s)) Left" Right Nil Right t, unfolded append_Nil2 opp_simps(1) replicate_append_same] by (metis \Suc (num_args (fun s)) = num_args (fun t)\ \is_App t\ chop_emb_step_at replicate_Suc) then have "emb_step_at (replicate (num_args (fun s)) dir.Left) dir.Right (chop t) = chop s" unfolding chop_emb_step_at[OF \is_App s\] using merge_emb_step_at[of "replicate (num_args (fun s)) Left" Right Nil Left t, unfolded append_Nil2 opp_simps(1) replicate_append_same] by (metis Left True \Suc (num_args (fun s)) = num_args (fun t)\ \emb_step_at p d t = s\ q_rep_t replicate_Suc) then show "chop t \\<^sub>e\<^sub>m\<^sub>b chop s" by (metis \is_App s\ \is_App t\ \t \\<^sub>e\<^sub>m\<^sub>b s\ emb_step_equiv emb_step_hsize nat_neq_iff hsize_chop) qed next case False assume "is_App s" have p_rep: "p = replicate (length p) Left" by (metis (full_types) \list_all (\x. x = Left) p\ list_all_iff replicate_length_same) have length_p:"length p < num_args t" using no_position_replicate_num_args \position_of t (p @ [d])\ replicate_add[of "num_args t" "length p - num_args t" Left] p_rep q_rep_t by (metis Left add_diff_inverse_nat replicate_app_Cons_same replicate_append_same shallower_pos) then have "length p \ length q" using Suc_num_args \is_App t\ q_rep_t by fastforce then have "length p < length q" using False le_neq_implies_less p_rep q_rep_t by fastforce then have "take (Suc (length p)) q = p @ [Left]" by (metis (no_types, lifting) \length p \ length q\ length_replicate min.orderE nth_replicate p_rep q_rep_t take_Suc_conv_app_nth take_replicate) then obtain q' where "q = p @ [Left] @ q'" by (metis append.assoc append_take_drop_id) have "Suc (num_args (fun s)) = num_args (fun t)" by (metis (no_types, lifting) Cons_nth_drop_Suc Suc_num_args \args s = take i (args t) @ drop (Suc i) (args t)\ \i < num_args t\ \is_App s\ \is_App t\ add_Suc_right append_take_drop_id diff_Suc_1 length_Cons length_append) then have "chop s = emb_step_at p dir.Left (chop t)" using swap_nested_emb_step_at[of p q' Right Left t] chop_emb_step_at[OF \is_App s\] chop_emb_step_at[OF \is_App t\] by (metis (no_types, lifting) Cons_replicate_eq Left \emb_step_at p d t = s\ \q = p @ [dir.Left] @ q'\ append.assoc append_Cons diff_Suc_1 p_rep q_rep_t replicate_append_same) then show "chop t \\<^sub>e\<^sub>m\<^sub>b chop s" by (metis Left \is_App t\ \position_of t (p @ [d])\ \q = p @ [dir.Left] @ q'\ chop_emb_step_at emb_step_at_if_position pos_emb_step_at_nested q_rep_t) qed then have "chkchop_same (>\<^sub>t) t s" proof (cases "is_Var (head t)") case True then show ?thesis unfolding chkchop_same_def chkchop_def using ih[of "chop t" "chop s"] add_less_mono hsize_chop_lt \is_App s \ chop t \\<^sub>e\<^sub>m\<^sub>b chop s\ by metis next case False then show ?thesis unfolding chkchop_same_def chkchop_def using \is_App t\ add_less_mono gt_chop ih hsize_chop_lt \is_App s \ chop t \\<^sub>e\<^sub>m\<^sub>b chop s\ by metis qed have "\f\local.ground_heads (head t). extf f (>\<^sub>t) (args t) (args s)" using extf_ext_insert_arg[of _ _ "take i (args t)" "args t ! i" "drop (Suc i) (args t)"] using \args s = take i (args t) @ drop (Suc i) (args t)\ \i < num_args t\ id_take_nth_drop by fastforce then show ?thesis using gt_same \head t = head s\ \chkchop_same (>\<^sub>t) t s\ by blast next text \Embedding chops and might remove arguments from the left\ case Right show ?thesis using emb_step_at_chop by (metis Right True \emb_step_at p d t = s\ \is_App t\ \position_of t (p @ [d])\ add_Suc gt_chop ih less_Suc_eq hsize_chop) qed next text \Embedding operates under one of the arguments\ case False have "num_args t = num_args s" using emb_step_under_args_num_args[OF False] by (metis (no_types) \\t d. num_args (emb_step_at p d t) = num_args t\ \emb_step_at p d t = s\) then have "is_App s" using \num_args t = num_args s\ by (metis args_Nil_iff_is_Hd length_0_conv \is_App t\) have q_rep_s: "q = replicate (num_args (fun s)) Left" by (metis q_rep_t \is_App s\ \is_App t\ \num_args t = num_args s\ args.simps(2) butlast_snoc length_butlast tm.collapse(2)) have "chop t \\<^sub>e\<^sub>m\<^sub>b chop s" proof (cases "take (num_args (fun t)) p = q") case True have "num_args (fun t) < length p" proof (rule ccontr) assume "\ num_args (fun t) < length p" then have "num_args (fun t) = length p" "q = p" using True q_rep_t by auto then show False using False using list_all_length q_rep_t by fastforce qed have "p ! (num_args (fun t)) = Right" proof (rule ccontr) assume "p ! num_args (fun t) \ Right" then have "p ! num_args (fun t) = Left" using dir.exhaust by blast then have "take (num_args t) p = replicate (num_args t) Left" using True Suc_num_args[OF \is_App t\] q_rep_t take_Suc_conv_app_nth[of "num_args (fun t)" p] by (metis \num_args (fun t) < length p\ replicate_Suc replicate_append_same) then show False by (metis \num_args t = num_args s\ \num_args t = num_args s\ \position_of t (p @ [d])\ append.assoc append_eq_Cons_conv append_take_drop_id no_position_replicate_num_args shallower_pos) qed then obtain q' where "p = q @ [Right] @ q'" by (metis Cons_nth_drop_Suc True \num_args (fun t) < length p\ append_Cons append_Nil append_eq_conv_conj length_replicate q_rep_t) have "emb_step_at (q @ q') d (chop t) = chop s" unfolding chop_emb_step_at[OF \is_App t\] chop_emb_step_at[OF \is_App s\] using swap_nested_emb_step_at[of q q' d Right t, unfolded] \emb_step_at p d t = s\ \p = q @ [Right] @ q'\ q_rep_t q_rep_s by auto moreover have "chop t \ chop s" by (metis \is_App s\ \is_App t\ \t \\<^sub>e\<^sub>m\<^sub>b s\ emb_step_hsize nat_less_le hsize_chop) ultimately show "chop t \\<^sub>e\<^sub>m\<^sub>b chop s" using emb_step_equiv by blast next case False then have takepq: "take (length q) p \ q" using q_rep_t by auto have takeqp: "length p \ length q \ take (length p) q \ p" using \\ list_all (\x. x = Left) p\[unfolded list_all_length] using diff_diff_cancel take_replicate length_replicate nth_replicate q_rep_s by metis have "chop s = emb_step_at p d (chop t)" using swap_disjunct_emb_step_at[of p q Right d t, OF takeqp takepq, unfolded \emb_step_at p d t = s\] using \is_App s\ \is_App t\ chop_emb_step_at q_rep_s q_rep_t by (simp add: chop_emb_step_at) then show ?thesis by (metis False \is_App t\ \position_of t (p @ [d])\ chop_emb_step_at emb_step_at_if_position length_replicate nat_le_linear pos_emb_step_at_disjunct q_rep_t take_all takeqp) qed then have gt1:"chkchop_same (>\<^sub>t) t s" proof (cases "head t") case (Var _) then have "chop t >\<^sub>t chop s" using ih[of "chop t" "chop s"] by (meson \chop t \\<^sub>e\<^sub>m\<^sub>b chop s\ \is_App s\ \is_App t\ add_strict_mono hsize_chop_lt) then show ?thesis unfolding chkchop_same_def by (simp add: Var) next case (Sym _) then show ?thesis unfolding chkchop_same_def using \chop t \\<^sub>e\<^sub>m\<^sub>b chop s\ \is_App t\ add_Suc add_Suc_shift chkchop_def gt_trans ih less_Suc_eq hsize_chop t_gt_chop_t by (metis (no_types, lifting)) qed have gt2:"head t = head s" by (metis emb_step_under_args_head False \emb_step_at p d t = s\) have gt3:"\f\local.ground_heads (head t). extf f (>\<^sub>t) (args t) (args s)" proof (rule ballI) fix f assume "f\local.ground_heads (head t)" obtain i where i_def: "i < num_args t" "args t ! i \\<^sub>e\<^sub>m\<^sub>b args (emb_step_at p d t) ! i" "\j. j < num_args t \ i \ j \ args t ! j = args (emb_step_at p d t) ! j" using emb_step_under_args_emb_step[of p t d] using False \position_of t (p @ [d])\ by blast have compat_list1: "args t ! i >\<^sub>t args s ! i" by (metis \num_args t = num_args s\ \emb_step_at p d t = s\ add_less_mono i_def(1) i_def(2) ih nth_mem hsize_in_args) have compat_list2: "args t ! i \ args s ! i" using emb_step_equiv i_def(2) \emb_step_at p d t = s\ by blast have argst:"args t = take i (args t) @ args t ! i # drop (Suc i) (args t)" by (simp add: Cons_nth_drop_Suc i_def(1)) have argss:"args s = take i (args t) @ args (emb_step_at p d t) ! i # drop (Suc i) (args t)" proof - have "take i (args t) = take i (args s)" apply (rule nth_take_lemma) using \num_args t = num_args s\ i_def(1) i_def(3)[unfolded \emb_step_at p d t = s\] by auto moreover have "drop (Suc i) (args t) = drop (Suc i) (args s)" apply (rule nth_drop_lemma) apply (simp add: \num_args t = num_args s\) apply (simp add: Suc_le_eq i_def(1)) using Suc_n_not_le_n \emb_step_at p d t = s\ i_def(3) by blast ultimately show ?thesis using \emb_step_at p d t = s\ \num_args t = num_args s\ i_def(1) id_take_nth_drop by auto qed show "extf f (>\<^sub>t) (args t) (args s)" using extf_compat_list[of "args t ! i" "args (emb_step_at p d t) ! i" gt f "take i (args t)" "drop (Suc i) (args t)"] using \emb_step_at p d t = s\ argss argst compat_list1 compat_list2 by force qed show ?thesis using gt_same using gt1 gt2 gt3 by blast qed qed lemma gt_embedding_property: assumes "t \\<^sub>e\<^sub>m\<^sub>b s" "t \ s" shows "t >\<^sub>t s" using assms proof (induction) case (refl t) then show ?case by simp next case (step t u s) then show ?case using gt_embedding_step_property gt_trans by blast qed subsection \Subterm Property\ theorem gt_proper_sub: "proper_sub s t \ t >\<^sub>t s" using gt_embedding_property sub_emb by blast lemma gt_emb_fun: "App s t >\<^sub>t s" and gt_emb_arg: "App s t >\<^sub>t t" by (simp_all add: gt_embedding_step_property left right) subsection \Compatibility with Contexts\ lemma gt_fun_imp: "fun t >\<^sub>t s \ t >\<^sub>t s" by (metis emb_step_fun gt_embedding_step_property gt_trans tm.disc(2) tm.exhaust_sel tm.sel(3)) lemma gt_arg_imp: "arg t >\<^sub>t s \ t >\<^sub>t s" by (metis emb_step_arg gt_embedding_step_property gt_trans tm.disc(2) tm.exhaust_sel tm.sel(5)) lemma gt_compat_fun: assumes "t' >\<^sub>t t" shows "App s t' >\<^sub>t App s t" using assms apply (simp only:atomize_imp) proof (induction rule:measure_induct_rule[of "\(t, s). hsize t + hsize s" "\(t, s). t' >\<^sub>t t \ App s t' >\<^sub>t App s t" "(t,s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric]) fix t s ::"('s, 'v) tm" assume ih:"\ta sa. hsize ta + hsize sa < hsize t + hsize s \ t' >\<^sub>t ta \ App sa t' >\<^sub>t App sa ta" and t'_gt_t:"t' >\<^sub>t t" have t'_ne_t: "t' \ t" using gt_antisym t'_gt_t by blast have extf_args_single: "\f \ ground_heads (head s). extf f (>\<^sub>t) (args s @ [t']) (args s @ [t])" by (simp add: extf_compat_list t'_gt_t t'_ne_t) show "App s t' >\<^sub>t App s t" proof (rule gt_same) show "head (App s t') = head (App s t)" by simp show "\f\local.ground_heads (head (App s t')). extf f (>\<^sub>t) (args (App s t')) (args (App s t))" by (simp add: extf_args_single) have 0: "chop (App s t') >\<^sub>t chop (App s t)" proof (cases s) case (Hd _) then show ?thesis using chop_App_Hd by (simp add: chop_App_Hd t'_gt_t) next case (App s1 s2) then show ?thesis using ih[of t "chop s"] chop_fun by (metis nat_add_left_cancel_less hsize_chop_lt t'_gt_t tm.disc(2) tm.sel(4) tm.sel(6)) qed show "chkchop_same (>\<^sub>t) (App s t') (App s t)" proof (cases "is_Var (head (App s t'))") case True then show ?thesis unfolding chkchop_same_def chkchop_def using True 0 by auto next case False have "App s t' >\<^sub>t chop (App s t)" using 0 by (simp add: gt_chop) then show ?thesis unfolding chkchop_same_def chkchop_def using False by auto qed qed qed theorem gt_compat_arg: shows "s' >\<^sub>t s \ t' \\<^sub>t t \ App s' t' >\<^sub>t App s t" proof (simp only:atomize_imp,induction rule:measure_induct[of "\(s',s,t). hsize s' + hsize s + hsize t" "\(s',s,t). s' >\<^sub>t s \ t' \\<^sub>t t \ App s' t' >\<^sub>t App s t" "(s',s,t)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric] atomize_all[symmetric]) fix s' s t assume ih:"\ab ac ba. hsize ab + hsize ac + hsize ba < hsize s' + hsize s + hsize t \ ab >\<^sub>t ac \ t' \\<^sub>t ba \ App ab t' >\<^sub>t App ac ba" and "s' >\<^sub>t s" and "t' \\<^sub>t t" { fix s''::"('s,'v) tm" assume hsize_s'':"hsize s'' \ hsize s'" assume chkchop_s'_s: "chkchop (>\<^sub>t) s'' s" then have "chkchop (>\<^sub>t) (App s'' t') (App s t)" proof (cases "is_Hd s") case True then show ?thesis using chkchop_s'_s unfolding chkchop_def by (metis \t' \\<^sub>t t\ chop_App_Hd gt_arg_imp gt_emb_arg tm.sel(6)) next case False then show ?thesis using chkchop_s'_s unfolding chkchop_def using ih[of s'' "chop s" t] hsize_s'' by (metis \t' \\<^sub>t t\ add_less_mono add_mono_thms_linordered_field(1) chop_fun le_eq_less_or_eq nat_add_left_cancel_less hsize_chop_lt tm.sel(4) tm.sel(6)) qed } note chkchop_compat_arg = this show "App s' t' >\<^sub>t App s t " using \s' >\<^sub>t s\ proof (cases rule:gt.cases) case gt_chop then show ?thesis proof (cases "t = t'") case True show ?thesis using True gt.gt_chop[of "App s' t'" "App s t"] gt_chop chkchop_compat_arg[of "chop s'"] by (metis add_strict_right_mono chop_fun ih hsize_chop_lt tm.disc(2) tm.sel(4) tm.sel(6)) next case False then have "t' >\<^sub>t t" using \t' \\<^sub>t t\ by blast have "App s' t' >\<^sub>t App (chop s') t'" by (simp add: context_left emb_step_chop gt_embedding_step_property local.gt_chop(1)) moreover have "... >\<^sub>t App s t" using ih[of "chop s'" s t] using \t' >\<^sub>t t\ gt_compat_fun local.gt_chop(1) local.gt_chop(2) hsize_chop_lt by fastforce ultimately show ?thesis using gt_trans by blast qed next case gt_diff then show ?thesis using chkchop_compat_arg gt.gt_diff by auto next case gt_same have hd_s'_eq_s: "head s' = head s" by (simp add: local.gt_same(1)) { fix f assume f_gh: "f \ ground_heads (head s)" have f_s_args: "extf f (>\<^sub>t) (args s') (args s)" using local.gt_same(3) f_gh by (simp add: hd_s'_eq_s) have f_compat_snoc: "\xs ys x. extf f (>\<^sub>t) ys xs \ extf f (>\<^sub>t) (ys @ [x]) (xs @ [x])" by (simp add: extf_compat_append_right) have f_st_args2: "extf f (>\<^sub>t) (args (App s' t)) (args (App s t))" by (simp add: f_compat_snoc f_s_args) have 0:"\z\UNIV. \y\UNIV. \x\UNIV. z >\<^sub>t y \ y >\<^sub>t x \ z >\<^sub>t x" using gt_trans by blast then have f_trans:"\xs ys zs. extf f (>\<^sub>t) zs ys \ extf f (>\<^sub>t) ys xs \ extf f (>\<^sub>t) zs xs" using extf_trans[of _ UNIV, unfolded lists_UNIV, OF UNIV_I UNIV_I UNIV_I 0] by metis have "extf f (>\<^sub>t) (args (App s' t')) (args (App s t))" proof (cases "t' = t") case True then show ?thesis using f_st_args2 by metis next case False have f_st_args1: "extf f (>\<^sub>t) (args (App s' t')) (args (App s' t))" using extf_compat_list \t' \\<^sub>t t\ False by simp then show ?thesis using f_trans f_st_args1 f_st_args2 by metis qed } note extf_cond = this have "chkchop_same (>\<^sub>t) (App s' t') (App s t)" unfolding chkchop_same_def using args.simps(1) chop_fun chkchop_compat_arg[of "chop s'", unfolded le_eq_less_or_eq] chkchop_compat_arg[of s'] chkchop_def chkchop_same_def hsize_chop_lt epo.extf_min_empty[OF epo_axioms] gt.gt_same gt_antisym hd_s'_eq_s head_App leI less_irrefl_nat local.gt_same(2) local.gt_same(3) tm.collapse(1) tm.sel(4) tm.sel(6) by metis then show ?thesis using extf_cond gt.gt_same hd_s'_eq_s by auto qed qed theorem gt_compat_fun_strong: assumes t'_gt_t: "t' >\<^sub>t t" shows "apps s (t' # us) >\<^sub>t apps s (t # us)" proof (induct us rule: rev_induct) case Nil then show ?case by (simp add: gt_compat_fun t'_gt_t) next case (snoc x xs) then show ?case unfolding App_apps[symmetric] append_Cons[symmetric] using gt_compat_arg by blast qed theorem gt_or_eq_compat_App: "s' \\<^sub>t s \ t' \\<^sub>t t \ App s' t' \\<^sub>t App s t" using gt_compat_fun gt_compat_arg by blast theorem gt_compat_App: shows "s' \\<^sub>t s \ t' >\<^sub>t t \ App s' t' >\<^sub>t App s t" using gt_compat_fun gt_compat_arg by blast subsection "Stability under Substitutions" (* TODO: move *) lemma extf_map2: assumes "\y\set ys \ set xs. \x\set ys \ set xs. y >\<^sub>t x \ (h y) >\<^sub>t (h x)" "extf f (>\<^sub>t) ys xs" shows "extf f (>\<^sub>t) (map h ys) (map h xs)" apply (rule extf_map[of "set ys \ set xs" ys xs "(>\<^sub>t)" h f]) apply simp apply (simp add: in_listsI) apply (simp add: in_listsI) using gt_antisym apply blast using gt_trans apply blast by (simp add: assms)+ -(* TODO: move ? *) -lemma less_multiset_doubletons: - assumes - "y < t \ y < s" - "x < t \ x < s" - shows - "{# y, x#} < {# t, s#}" - unfolding less_multiset\<^sub>D\<^sub>M -proof (intro exI) - let ?X = "{# t, s#}" - let ?Y = "{#y, x#}" - show "?X \ {#} \ ?X \# {#t, s#} \ {#y, x#} = {#t, s#} - ?X + ?Y \ (\k. k \# ?Y \ (\a. a \# ?X \ k < a))" - using add_eq_conv_diff assms(1) assms(2) by auto -qed - theorem gt_sus: assumes \_wary: "wary_subst \" assumes ghd: "\x. ground_heads (Var x) = UNIV" (* This condition is only needed for gt_same, not for gt_diff ! *) shows "t >\<^sub>t s \ subst \ t >\<^sub>t subst \ s" proof (simp only:atomize_imp,induction rule:measure_induct[of "\(t,s). {# hsize t, hsize s #}" "\(t,s). t >\<^sub>t s \ subst \ t >\<^sub>t subst \ s" "(t,s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric] atomize_all[symmetric]) fix t s assume ih:"\tt ss. {# hsize tt, hsize ss #} < {# hsize t, hsize s #} \ tt >\<^sub>t ss \ subst \ tt >\<^sub>t subst \ ss" and "t >\<^sub>t s" show "subst \ t >\<^sub>t subst \ s" using \t >\<^sub>t s\ proof (cases) case t_gt_s_chop: gt_chop then show ?thesis using emb_step_subst emb_step_chop[OF t_gt_s_chop(1)] gt_embedding_step_property emb_step_hsize gt_trans ih[of "chop t" s] by (metis add_mset_lt_left_lt) next case t_gt_s_diff: gt_diff have gt_diff1: "head (subst \ t) >\<^sub>h\<^sub>d head (subst \ s)" by (meson assms gt_hd_def subsetCE t_gt_s_diff(1) wary_subst_ground_heads) have gt_diff2: "is_Sym (head (subst \ s))" by (metis ground_imp_subst_iden hd.collapse(2) hd.simps(18) head_subst t_gt_s_diff(2) tm.sel(1) tm.simps(17)) have gt_diff3: "chkchop (>\<^sub>t) (subst \ t) (subst \ s)" proof (cases s) case (Hd _) then show ?thesis using t_gt_s_diff unfolding chkchop_def by (metis ground_imp_subst_iden hd.collapse(2) hd.simps(18) tm.disc(1) tm.sel(1) tm.simps(17)) next case s_App: (App s1 s2) then show ?thesis using t_gt_s_diff unfolding chkchop_def using ih[of t "chop s"] chop_subst_Sym hsize_chop_lt tm.disc(2) by (metis add_mset_lt_left_lt add_mset_lt_right_lt) qed show ?thesis using gt_diff gt_diff1 gt_diff2 gt_diff3 by blast next case t_gt_s_same: gt_same have gt_same1: "head (subst \ t) = head (subst \ s)" by (simp add: t_gt_s_same(1)) have extf_map_ts:"\f\ground_heads (head t). extf f (>\<^sub>t) (map (subst \) (args t)) (map (subst \) (args s))" proof - have ih_args: "\y\set (args t) \ set (args s). \x\set (args t) \ set (args s). y >\<^sub>t x \ subst \ y >\<^sub>t subst \ x" by (metis Un_iff less_multiset_doubletons hsize_in_args ih) have "\f\ground_heads (head t). extf f (>\<^sub>t) (args t) (args s)" using ghd t_gt_s_same(3) by metis then show ?thesis using extf_map[of "set (args t) \ set (args s)" "args t" "args s" gt "subst \"] using gt_irrefl gt_trans ih_args by blast qed show ?thesis proof (cases "head t") case (Var x1) then have "is_Var (head t)" by simp { fix u :: "('s, 'v) tm" assume "ground_heads (head u) \ ground_heads (head t)" "hsize u \ hsize (subst \ (Hd (head t)))" then have "apps u (map (subst \) (args t)) >\<^sub>t apps u (map (subst \) (args s))" proof (induct "hsize u" arbitrary:u rule:less_induct) case less then show ?case proof (cases u) case u_Hd: (Hd _) then have "args u = []" by simp then show ?thesis proof (cases s) case s_Hd: (Hd _) show ?thesis apply (rule gt_same) using extf_map_ts args_Nil_iff_is_Hd s_Hd u_Hd \args u = []\ less by fastforce+ next case s_App: (App _ _) then have "is_App t" by (metis args_Nil_iff_is_Hd extf_min_empty gt_hd_def gt_hd_irrefl t_gt_s_same(3)) have "chop t >\<^sub>t chop s" using \is_App t\ \is_Var (head t)\ s_App t_gt_s_same(2) by auto then have "subst \ (chop t) >\<^sub>t subst \ (chop s)" using ih by (metis \is_App t\ less_multiset_doubletons s_App hsize_chop_lt tm.disc(2)) define ut where "ut = apps u (map (subst \) (args t))" define us where "us = apps u (map (subst \) (args s))" have 0:"\ss. args (apps u ss) = ss" using \args u = []\ by simp have chop_us: "chop us = subst \ (chop s)" unfolding chop_def subst_apps us_def 0 using hd_map by (metis (no_types, lifting) args_Nil_iff_is_Hd map_tl s_App tm.disc(2)) have chop_ut: "chop ut = subst \ (chop t)" unfolding chop_def subst_apps ut_def 0 using \is_App t\ by (simp add: args_Nil_iff_is_Hd hd_map map_tl) have "head ut = head us" by (simp add: us_def ut_def) moreover have "chkchop_same (>\<^sub>t) ut us" unfolding chkchop_def chkchop_same_def by (metis "0" UNIV_witness \is_Var (head t)\ \subst \ (chop t) >\<^sub>t subst \ (chop s)\ args.simps(1) chop_us chop_ut extf_map_ts extf_min_empty ghd gt_chop is_Var_def tm.collapse(1) ut_def) moreover have "\f\local.ground_heads (head ut). extf f (>\<^sub>t) (args ut) (args us)" using extf_map_ts less us_def ut_def using "0" by auto ultimately show "ut >\<^sub>t us" using gt_same by blast qed next case u_app: (App _ _) let ?ut = "apps u (map (subst \) (args t))" let ?us = "apps u (map (subst \) (args s))" have 1:"head ?ut = head ?ut" by simp have "apps (chop u) (map (subst \) (args t)) >\<^sub>t apps (chop u) (map (subst \) (args s))" using less.hyps[of "chop u"] hsize_chop_lt by (metis Var dual_order.trans ghd less.prems(2) less_or_eq_imp_le subset_UNIV tm.disc(2) u_app) then have "chop ?ut >\<^sub>t chop ?us" by (simp add: chop_apps u_app) then have 2:"chkchop_same (>\<^sub>t) ?ut ?us" unfolding chkchop_same_def chkchop_def by (metis UNIV_I \is_Var (head t)\ args_Nil_iff_is_Hd args_apps extf_compat_append_left extf_map_ts extf_min_empty ghd gt_chop is_Var_def) have 3:"\f\local.ground_heads (head ?ut). extf f (>\<^sub>t) (args ?ut) (args ?us)" using extf_compat_append_left extf_map_ts less.prems(1) by auto show ?thesis using gt_same 1 2 3 by simp qed qed } note inner_induction = this show ?thesis using inner_induction[of "subst \ (Hd (head t))", unfolded subst_apps[symmetric]] by (metis Var ghd order_refl subset_UNIV t_gt_s_same(1) tm_collapse_apps) next case (Sym _) then have "is_Sym (head (subst \ t))" "head (subst \ t) = head t" by simp_all then have "chkchop_same (>\<^sub>t) t s" using t_gt_s_same unfolding chkchop_same_def chkchop_def using Sym by metis then have gt_same2: "chkchop_same (>\<^sub>t) (subst \ t) (subst \ s)" unfolding chkchop_same_def chkchop_def using ih[of t "chop s"] by (metis (no_types, lifting) Sym \head (subst \ t) = head t\ \is_Sym (head (subst \ t))\ add_mset_commute add_mset_lt_left_lt chop_subst_Sym ground_imp_subst_iden hd.simps(18) hsize_chop_lt t_gt_s_same(1) tm.collapse(1) tm.simps(17)) have gt_same3: "\f\local.ground_heads (head (subst \ t)). extf f (>\<^sub>t) (args (subst \ t)) (args (subst \ s))" using \head (subst \ t) = head t\ extf_compat_append_left extf_map_ts t_gt_s_same(1) by auto show ?thesis using gt_same gt_same1 gt_same2 gt_same3 by blast qed qed qed subsection \Totality on Ground Terms\ theorem gt_total_ground: assumes extf_total: "\f. ext_total (extf f)" shows "ground t \ ground s \ t >\<^sub>t s \ s >\<^sub>t t \ t = s" proof (simp only: atomize_imp, rule measure_induct_rule[of "\(t, s). {# hsize t, hsize s #}" "\(t, s). ground t \ ground s \ t >\<^sub>t s \ s >\<^sub>t t \ t = s" "(t, s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric]) fix t s :: "('s, 'v) tm" assume ih: "\ta sa. {# hsize ta, hsize sa #} < {# hsize t, hsize s #} \ ground ta \ ground sa \ ta >\<^sub>t sa \ sa >\<^sub>t ta \ ta = sa" and gr_t: "ground t" and gr_s: "ground s" let ?case = "t >\<^sub>t s \ s >\<^sub>t t \ t = s" have "chkchop (>\<^sub>t) t s \ s >\<^sub>t t" unfolding chkchop_def tm.case_eq_if using ih[of t "chop s"] by (metis (no_types, lifting) add_mset_commute add_mset_lt_left_lt gr_s gr_t ground_chop gt_chop hsize_chop_lt) moreover have "chkchop (>\<^sub>t) s t \ t >\<^sub>t s" unfolding chkchop_def tm.case_eq_if using ih[of "chop t" s] by (metis add_mset_lt_left_lt gr_s gr_t ground_chop gt_chop.intros gt_iff_chop_diff_same hsize_chop_lt) moreover { assume chkembs_t_s: "chkchop (>\<^sub>t) t s" and chkembs_s_t: "chkchop (>\<^sub>t) s t" obtain g where g: "head t = Sym g" using gr_t by (metis ground_head hd.collapse(2)) obtain f where f: "head s = Sym f" using gr_s by (metis ground_head hd.collapse(2)) { assume g_gt_f: "g >\<^sub>s f" have "t >\<^sub>t s" using chkembs_t_s f g g_gt_f gt_diff gt_sym_imp_hd by auto } moreover { assume f_gt_g: "f >\<^sub>s g" have "s >\<^sub>t t" using chkembs_s_t f f_gt_g g gt_diff gt_sym_imp_hd by auto } moreover { assume g_eq_f: "g = f" hence hd_t: "head t = head s" using g f by auto let ?ts = "args t" let ?ss = "args s" have gr_ts: "\ta \ set ?ts. ground ta" using ground_args[OF _ gr_t] by blast have gr_ss: "\sa \ set ?ss. ground sa" using ground_args[OF _ gr_s] by blast { assume ts_eq_ss: "?ts = ?ss" have "t = s" using hd_t ts_eq_ss by (rule tm_expand_apps) } moreover { assume ts_gt_ss: "extf g (>\<^sub>t) ?ts ?ss" have "t >\<^sub>t s" using chkembs_t_s g gt_same hd_t ts_gt_ss by auto } moreover { assume ss_gt_ts: "extf g (>\<^sub>t) ?ss ?ts" have "s >\<^sub>t t" using chkembs_s_t f g_eq_f gt_same hd_t ss_gt_ts by auto } ultimately have ?case using ih gr_ss gr_ts - ext_total.total[OF extf_total, rule_format, of "set ?ts \ set ?ss" "set ?ts \ set ?ss" "(>\<^sub>t)" ?ts ?ss g] - using epo.less_multiset_doubletons epo_axioms hsize_in_args in_listsI by (metis Un_iff) + ext_total.total[OF extf_total, rule_format, of "set ?ts \ set ?ss" "(>\<^sub>t)" ?ts ?ss g] + using less_multiset_doubletons epo_axioms hsize_in_args in_listsI by (metis Un_iff) } ultimately have ?case using gt_sym_total by blast } ultimately show ?case by fast qed subsection \Well-foundedness\ abbreviation gtg :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix ">\<^sub>t\<^sub>g" 50) where "(>\<^sub>t\<^sub>g) \ \t s. ground t \ t >\<^sub>t s" theorem gt_wf: assumes ghd_UNIV: "\x. ground_heads_var x = UNIV" assumes extf_wf: "\f. ext_wf (extf f)" shows "wfP (\s t. t >\<^sub>t s)" proof - have ground_wfP: "wfP (\s t. t >\<^sub>t\<^sub>g s)" unfolding wfP_iff_no_inf_chain proof assume "\f. inf_chain (>\<^sub>t\<^sub>g) f" then obtain t where t_bad: "bad (>\<^sub>t\<^sub>g) t" unfolding inf_chain_def bad_def by blast let ?ff = "worst_chain (>\<^sub>t\<^sub>g) (\t s. hsize t > hsize s)" let ?U_of = "\i. {u. (?ff i) \\<^sub>e\<^sub>m\<^sub>b u}" note wf_sz = wf_app[OF wellorder_class.wf, of hsize, simplified] define U where "U = (\i. ?U_of i)" have gr: "\i. ground (?ff i)" using worst_chain_bad[OF wf_sz t_bad, unfolded inf_chain_def] by fast have gr_u: "\u. u \ U \ ground u" unfolding U_def using gr ground_emb by fastforce have "\ bad (>\<^sub>t\<^sub>g) u" if u_in: "u \ ?U_of i" for u i proof let ?ti = "?ff i" assume u_bad: "bad (>\<^sub>t\<^sub>g) u" have sz_u: "hsize u < hsize ?ti" using emb_hsize_neq u_in by blast show False proof (cases i) case 0 thus False using sz_u min_worst_chain_0[OF wf_sz u_bad] by simp next case Suc hence "?ff (i - 1) >\<^sub>t ?ff i" using worst_chain_pred[OF wf_sz t_bad] by simp moreover have "?ff i >\<^sub>t u" using gt_embedding_property u_in by blast ultimately have "?ff (i - 1) >\<^sub>t u" by (rule gt_trans) thus False using Suc sz_u min_worst_chain_Suc[OF wf_sz u_bad] gr by fastforce qed qed hence u_good: "\u. u \ U \ \ bad (>\<^sub>t\<^sub>g) u" unfolding U_def by blast have bad_diff_same: "inf_chain (\t s. ground t \ (gt_diff t s \ gt_same t s)) ?ff" unfolding inf_chain_def proof (intro allI conjI) fix i show "ground (?ff i)" by (rule gr) have gt: "?ff i >\<^sub>t ?ff (Suc i)" using worst_chain_pred[OF wf_sz t_bad] by blast have "\ gt_chop (?ff i) (?ff (Suc i))" proof assume a: "gt_chop (?ff i) (?ff (Suc i))" then have "chop (?ff i) \ ?U_of i" by (metis (mono_tags, lifting) emb_step_chop emb_step_is_emb gt_chop gt_chop.cases gt_irrefl mem_Collect_eq) then have uij_in:"chop (?ff i) \ U" unfolding U_def by fast have "\n. ?ff n >\<^sub>t ?ff (Suc n)" by (rule worst_chain_pred[OF wf_sz t_bad, THEN conjunct2]) hence uij_gt_i_plus_3: "chop (?ff i) >\<^sub>t ?ff (Suc (Suc i))" using gt_trans by (metis (mono_tags, lifting) a gt_chop.cases) have "inf_chain (>\<^sub>t\<^sub>g) (\j. if j = 0 then chop (?ff i) else ?ff (Suc (i + j)))" unfolding inf_chain_def by (auto intro!: gr gr_u[OF uij_in] uij_gt_i_plus_3 worst_chain_pred[OF wf_sz t_bad]) hence "bad (>\<^sub>t\<^sub>g) (chop (?ff i))" unfolding bad_def by fastforce thus False using u_good[OF uij_in] by sat qed thus "gt_diff (?ff i) (?ff (Suc i)) \ gt_same (?ff i) (?ff (Suc i))" using gt unfolding gt_iff_chop_diff_same by sat qed have "wf {(s, t). ground s \ ground t \ sym (head t) >\<^sub>s sym (head s)}" using gt_sym_wf unfolding wfP_def wf_iff_no_infinite_down_chain by fast moreover have "{(s, t). ground t \ gt_diff t s} \ {(s, t). ground s \ ground t \ sym (head t) >\<^sub>s sym (head s)}" proof (clarsimp, intro conjI) fix s t assume gr_t: "ground t" and gt_diff_t_s: "gt_diff t s" thus gr_s: "ground s" using gt_iff_chop_diff_same gt_imp_vars by fastforce show "sym (head t) >\<^sub>s sym (head s)" using gt_diff_t_s ground_head[OF gr_s] ground_head[OF gr_t] by (cases; cases "head s"; cases "head t") (auto simp: gt_hd_def) qed ultimately have wf_diff: "wf {(s, t). ground t \ gt_diff t s}" by (rule wf_subset) have diff_O_same: "{(s, t). ground t \ gt_diff t s} O {(s, t). ground t \ gt_same t s} \ {(s, t). ground t \ gt_diff t s}" unfolding gt_diff.simps gt_same.simps by clarsimp (metis chkchop_def chkchop_same_def gt_same gt_trans) have diff_same_as_union: "{(s, t). ground t \ (gt_diff t s \ gt_same t s)} = {(s, t). ground t \ gt_diff t s} \ {(s, t). ground t \ gt_same t s}" by auto obtain k where bad_same: "inf_chain (\t s. ground t \ gt_same t s) (\i. ?ff (i + k))" using wf_infinite_down_chain_compatible[OF wf_diff _ diff_O_same, of ?ff] bad_diff_same unfolding inf_chain_def diff_same_as_union[symmetric] by auto hence hd_sym: "\i. is_Sym (head (?ff (i + k)))" unfolding inf_chain_def by (simp add: ground_head) define f where "f = sym (head (?ff k))" have hd_eq_f: "head (?ff (i + k)) = Sym f" for i proof (induct i) case 0 thus ?case by (auto simp: f_def hd.collapse(2)[OF hd_sym, of 0, simplified]) next case (Suc ia) thus ?case using bad_same unfolding inf_chain_def gt_same.simps by simp qed let ?gtu = "\t s. t \ U \ t >\<^sub>t s" thm UnionI CollectI have "t \ set (args (?ff i)) \ t \ U" for t i unfolding U_def apply (rule UnionI[of "?U_of i"]) using arg_emb CollectI arg_emb hsize_in_args by fast+ moreover have "\i. extf f (>\<^sub>t\<^sub>g) (args (?ff (i + k))) (args (?ff (Suc i + k)))" using bad_same hd_eq_f unfolding inf_chain_def gt_same.simps f_def hd.collapse(2)[OF ground_head, OF gr] using extf_mono_strong[of _ _ "(>\<^sub>t)" "(\t s. ground t \ t >\<^sub>t s)" ] ground_hd_in_ground_heads by (metis (no_types, lifting) ground_args) ultimately have "\i. extf f ?gtu (args (?ff (i + k))) (args (?ff (Suc i + k)))" using extf_mono_strong[of _ _ "(\t s. ground t \ t >\<^sub>t s)" "\t s. t \ U \ t >\<^sub>t s"] unfolding U_def by blast hence "inf_chain (extf f ?gtu) (\i. args (?ff (i + k)))" unfolding inf_chain_def by blast hence nwf_ext: "\ wfP (\xs ys. extf f ?gtu ys xs)" unfolding wfP_iff_no_inf_chain by fast have gtu_le_gtg: "?gtu \ (>\<^sub>t\<^sub>g)" by (auto intro!: gr_u) have "wfP (\s t. ?gtu t s)" unfolding wfP_iff_no_inf_chain proof (intro notI, elim exE) fix f assume bad_f: "inf_chain ?gtu f" hence bad_f0: "bad ?gtu (f 0)" by (rule inf_chain_bad) have "f 0 \ U" using bad_f unfolding inf_chain_def by blast hence good_f0: "\ bad ?gtu (f 0)" using u_good bad_f inf_chain_bad inf_chain_subset[OF _ gtu_le_gtg] by blast show False using bad_f0 good_f0 by sat qed hence wf_ext: "wfP (\xs ys. extf f ?gtu ys xs)" by (rule ext_wf.wf[OF extf_wf, rule_format]) show False using nwf_ext wf_ext by blast qed let ?subst = "subst grounding_\" have "wfP (\s t. ?subst t >\<^sub>t\<^sub>g ?subst s)" by (rule wfP_app[OF ground_wfP]) hence "wfP (\s t. ?subst t >\<^sub>t ?subst s)" by (simp add: ground_grounding_\) thus ?thesis using gt_sus ghd_UNIV ground_heads.simps(1) wary_grounding_\ wfP_eq_minimal by (metis (no_types, lifting)) qed end end diff --git a/thys/Lambda_Free_KBOs/Lambda_Free_KBO_Std.thy b/thys/Lambda_Free_KBOs/Lambda_Free_KBO_Std.thy --- a/thys/Lambda_Free_KBOs/Lambda_Free_KBO_Std.thy +++ b/thys/Lambda_Free_KBOs/Lambda_Free_KBO_Std.thy @@ -1,1268 +1,1269 @@ (* Title: The Graceful Standard Knuth-Bendix Order for Lambda-Free Higher-Order Terms Author: Heiko Becker , 2016 Author: Jasmin Blanchette , 2016 Author: Uwe Waldmann , 2016 Author: Daniel Wand , 2016 Maintainer: Jasmin Blanchette *) section \The Graceful Standard Knuth--Bendix Order for Lambda-Free Higher-Order Terms\ theory Lambda_Free_KBO_Std imports Lambda_Free_KBO_Util abbrevs ">t" = ">\<^sub>t" and "\t" = "\\<^sub>t" begin text \ This theory defines the standard version of the graceful Knuth--Bendix order for \\\-free higher-order terms. Standard means that one symbol is allowed to have a weight of 0. \ subsection \Setup\ locale kbo_std = kbo_std_basis _ _ arity_sym arity_var wt_sym for arity_sym :: "'s \ enat" and arity_var :: "'v \ enat" and wt_sym :: "'s \ nat" begin subsection \Weights\ primrec wt :: "('s, 'v) tm \ nat" where "wt (Hd \) = (LEAST w. \f \ ground_heads \. w = wt_sym f + the_enat (\ * arity_sym f))" | "wt (App s t) = (wt s - \) + wt t" lemma wt_Hd_Sym: "wt (Hd (Sym f)) = wt_sym f + the_enat (\ * arity_sym f)" by simp lemma exists_wt_sym: "\f \ ground_heads \. wt (Hd \) = wt_sym f + the_enat (\ * arity_sym f)" by (auto intro: Least_in_nonempty_set_imp_ex) lemma wt_le_wt_sym: "f \ ground_heads \ \ wt (Hd \) \ wt_sym f + the_enat (\ * arity_sym f)" using not_le_imp_less not_less_Least by fastforce lemma enat_the_enat_\_times_arity_sym[simp]: "enat (the_enat (\ * arity_sym f)) = \ * arity_sym f" using arity_sym_ne_infinity_if_\_gt_0 imult_is_infinity zero_enat_def by fastforce lemma wt_arg_le: "wt (arg s) \ wt s" by (cases s) auto lemma wt_ge_\: "wt s \ \" by (induct s, metis exists_wt_sym of_nat_eq_enat le_diff_conv of_nat_id wt_sym_ge, simp add: add_increasing) lemma wt_ge_\: "wt s \ \" by (meson \_le_\ order.trans enat_ord_simps(1) wt_ge_\) lemma wt_gt_\_if_superunary: "arity_hd (head s) > 1 \ wt s > \" proof (induct s) case \: (Hd \) obtain g where g_in_grs: "g \ ground_heads \" and wt_\: "wt (Hd \) = wt_sym g + the_enat (\ * arity_sym g)" using exists_wt_sym by blast have "arity_hd \ > 1" using \ by auto hence ary_g: "arity_sym g > 1" using ground_heads_arity[OF g_in_grs] by simp show ?case proof (cases "\ = 0") case True thus ?thesis by (metis \_gt_0 gr0I leD wt_ge_\) next case \_ne_0: False hence ary_g_ninf: "arity_sym g \ \" using arity_sym_ne_infinity_if_\_gt_0 by blast hence "\ < the_enat (enat \ * arity_sym g)" using \_ne_0 ary_g by (cases "arity_sym g") (auto simp: one_enat_def) thus ?thesis unfolding wt_\ by simp qed next case (App s t) thus ?case using wt_ge_\[of t] by force qed lemma wt_App_\: "wt (App s t) = wt t \ wt s = \" by (simp add: order.antisym wt_ge_\) lemma wt_App_ge_fun: "wt (App s t) \ wt s" by (metis diff_le_mono2 wt_ge_\ le_diff_conv wt.simps(2)) lemma wt_hd_le: "wt (Hd (head s)) \ wt s" by (induct s, simp) (metis head_App leD le_less_trans not_le_imp_less wt_App_ge_fun) lemma wt_\_imp_\_eq_\: "wt s = \ \ \ = \" by (metis \_le_\ le_antisym wt_ge_\) lemma wt_ge_arity_head_if_\_gt_0: assumes \_gt_0: "\ > 0" shows "wt s \ arity_hd (head s)" proof (induct s) case (Hd \) obtain f where f_in_\: "f \ ground_heads \" and wt_\: "wt (Hd \) = wt_sym f + the_enat (\ * arity_sym f)" using exists_wt_sym by blast have "arity_sym f \ arity_hd \" by (rule ground_heads_arity[OF f_in_\]) hence "the_enat (\ * arity_sym f) \ arity_hd \" using \_gt_0 by (metis One_nat_def Suc_ile_eq dual_order.trans enat_ord_simps(2) enat_the_enat_\_times_arity_sym i0_lb mult.commute mult.right_neutral mult_left_mono one_enat_def) thus ?case unfolding wt_\ by (metis add.left_neutral add_mono le_iff_add plus_enat_simps(1) tm.sel(1)) next case App thus ?case by (metis dual_order.trans enat_ord_simps(1) head_App wt_App_ge_fun) qed lemma wt_ge_num_args_if_\_eq_0: assumes \_eq_0: "\ = 0" shows "wt s \ num_args s" by (induct s, simp_all, metis (no_types) \_eq_0 \_gt_0 wt_\_imp_\_eq_\ add_le_same_cancel1 le_0_eq le_trans minus_nat.diff_0 not_gr_zero not_less_eq_eq) lemma wt_ge_num_args: "wary s \ wt s \ num_args s" using wt_ge_arity_head_if_\_gt_0 wt_ge_num_args_if_\_eq_0 by (meson order.trans enat_ord_simps(1) neq0_conv wary_num_args_le_arity_head) subsection \Inductive Definitions\ inductive gt :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix ">\<^sub>t" 50) where gt_wt: "vars_mset t \# vars_mset s \ wt t > wt s \ t >\<^sub>t s" | gt_unary: "wt t = wt s \ \ head t \\\<^sub>h\<^sub>d head s \ num_args t = 1 \ (\f \ ground_heads (head t). arity_sym f = 1 \ wt_sym f = 0) \ arg t >\<^sub>t s \ arg t = s \ t >\<^sub>t s" | gt_diff: "vars_mset t \# vars_mset s \ wt t = wt s \ head t >\<^sub>h\<^sub>d head s \ t >\<^sub>t s" | gt_same: "vars_mset t \# vars_mset s \ wt t = wt s \ head t = head s \ (\f \ ground_heads (head t). extf f (>\<^sub>t) (args t) (args s)) \ t >\<^sub>t s" abbreviation ge :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix "\\<^sub>t" 50) where "t \\<^sub>t s \ t >\<^sub>t s \ t = s" inductive gt_wt :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_wtI: "vars_mset t \# vars_mset s \ wt t > wt s \ gt_wt t s" inductive gt_diff :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_diffI: "vars_mset t \# vars_mset s \ wt t = wt s \ head t >\<^sub>h\<^sub>d head s \ gt_diff t s" inductive gt_unary :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_unaryI: "wt t = wt s \ \ head t \\\<^sub>h\<^sub>d head s \ num_args t = 1 \ (\f \ ground_heads (head t). arity_sym f = 1 \ wt_sym f = 0) \ arg t \\<^sub>t s \ gt_unary t s" inductive gt_same :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_sameI: "vars_mset t \# vars_mset s \ wt t = wt s \ head t = head s \ (\f \ ground_heads (head t). extf f (>\<^sub>t) (args t) (args s)) \ gt_same t s" lemma gt_iff_wt_unary_diff_same: "t >\<^sub>t s \ gt_wt t s \ gt_unary t s \ gt_diff t s \ gt_same t s" by (subst gt.simps) (auto simp: gt_wt.simps gt_unary.simps gt_diff.simps gt_same.simps) lemma gt_imp_vars_mset: "t >\<^sub>t s \ vars_mset t \# vars_mset s" by (induct rule: gt.induct) (auto intro: subset_mset.order.trans) lemma gt_imp_vars: "t >\<^sub>t s \ vars t \ vars s" using set_mset_mono[OF gt_imp_vars_mset] by simp subsection \Irreflexivity\ theorem gt_irrefl: "wary s \ \ s >\<^sub>t s" proof (induct "size s" arbitrary: s rule: less_induct) case less note ih = this(1) and wary_s = this(2) show ?case proof assume s_gt_s: "s >\<^sub>t s" show False using s_gt_s proof (cases rule: gt.cases) case gt_same then obtain f where f: "extf f (>\<^sub>t) (args s) (args s)" by fastforce thus False using wary_s ih by (metis wary_args extf_irrefl size_in_args) qed (auto simp: comp_hd_def gt_hd_irrefl) qed qed subsection \Transitivity\ lemma gt_imp_wt_ge: "t >\<^sub>t s \ wt t \ wt s" by (induct rule: gt.induct) auto lemma not_extf_gt_nil_singleton_if_\_eq_\: assumes wary_s: "wary s" and \_eq_\: "\ = \" shows "\ extf f (>\<^sub>t) [] [s]" proof assume nil_gt_s: "extf f (>\<^sub>t) [] [s]" note s_gt_nil = extf_singleton_nil_if_\_eq_\[OF \_eq_\, of f gt s] have "\ extf f (>\<^sub>t) [] []" by (rule extf_irrefl) simp moreover have "extf f (>\<^sub>t) [] []" using extf_trans_from_irrefl[of "{s}", OF _ _ _ _ _ _ nil_gt_s s_gt_nil] gt_irrefl[OF wary_s] by fastforce ultimately show False by sat qed lemma gt_sub_arg: "wary (App s t) \ App s t >\<^sub>t t" proof (induct t arbitrary: s rule: measure_induct_rule[of size]) case (less t) note ih = this(1) and wary_st = this(2) { assume wt_st: "wt (App s t) = wt t" hence \_eq_\: "\ = \" using wt_App_\ wt_\_imp_\_eq_\ by metis hence \_gt_0: "\ > 0" using \_gt_0 by simp have wt_s: "wt s = \" by (rule wt_App_\[OF wt_st]) have wary_t: "wary t" and nargs_lt: "num_args s < arity_hd (head s)" using wary_st wary.simps by blast+ have ary_hd_s: "arity_hd (head s) = 1" by (metis One_nat_def arity.wary_AppE dual_order.order_iff_strict eSuc_enat enat_defs(1) enat_defs(2) ileI1 linorder_not_le not_iless0 wary_st wt_gt_\_if_superunary wt_s) hence nargs_s: "num_args s = 0" by (metis enat_ord_simps(2) less_one nargs_lt one_enat_def) have "s = Hd (head s)" by (simp add: Hd_head_id nargs_s) then obtain f where f_in: "f \ ground_heads (head s)" and wt_f_etc: "wt_sym f + the_enat (\ * arity_sym f) = \" using exists_wt_sym wt_s by fastforce have ary_f_1: "arity_sym f = 1" proof - have ary_f_ge_1: "arity_sym f \ 1" using ary_hd_s f_in ground_heads_arity by fastforce hence "enat \ * arity_sym f = \" using wt_f_etc by (metis enat_ord_simps(1) enat_the_enat_\_times_arity_sym le_add2 le_antisym mult.right_neutral mult_left_mono zero_le) thus ?thesis using \_gt_0 by (cases "arity_sym f") (auto simp: one_enat_def) qed hence wt_f_0: "wt_sym f = 0" using wt_f_etc by simp { assume hd_s_ncmp_t: "\ head s \\\<^sub>h\<^sub>d head t" have ?case by (rule gt_unary[OF wt_st]) (auto simp: hd_s_ncmp_t nargs_s intro: f_in ary_f_1 wt_f_0) } moreover { assume hd_s_gt_t: "head s >\<^sub>h\<^sub>d head t" have ?case by (rule gt_diff) (auto simp: hd_s_gt_t wt_s[folded \_eq_\]) } moreover { assume "head t >\<^sub>h\<^sub>d head s" hence False using ary_f_1 exists_wt_sym f_in gt_hd_def gt_sym_antisym unary_wt_sym_0_gt wt_f_0 by blast } moreover { assume hd_t_eq_s: "head t = head s" hence nargs_t_le: "num_args t \ 1" using ary_hd_s wary_num_args_le_arity_head[OF wary_t] by (simp add: one_enat_def) have extf: "extf f (>\<^sub>t) [t] (args t)" for f proof (cases "args t") case Nil thus ?thesis by (simp add: extf_singleton_nil_if_\_eq_\[OF \_eq_\]) next case args_t: (Cons ta ts) hence ts: "ts = []" using ary_hd_s[folded hd_t_eq_s] wary_num_args_le_arity_head[OF wary_t] nargs_t_le by simp have ta: "ta = arg t" by (metis apps.simps(1) apps.simps(2) args_t tm.sel(6) tm_collapse_apps ts) hence t: "t = App (fun t) ta" by (metis args.simps(1) args_t not_Cons_self2 tm.exhaust_sel ts) have "t >\<^sub>t ta" by (rule ih[of ta "fun t", folded t, OF _ wary_t]) (metis ta size_arg_lt t tm.disc(2)) thus ?thesis unfolding args_t ts by (metis extf_singleton gt_irrefl wary_t) qed have ?case by (rule gt_same) (auto simp: hd_t_eq_s wt_s[folded \_eq_\] length_0_conv[THEN iffD1, OF nargs_s] extf) } ultimately have ?case unfolding comp_hd_def by metis } thus ?case using gt_wt by fastforce qed lemma gt_arg: "wary s \ is_App s \ s >\<^sub>t arg s" by (cases s) (auto intro: gt_sub_arg) theorem gt_trans: "wary u \ wary t \ wary s \ u >\<^sub>t t \ t >\<^sub>t s \ u >\<^sub>t s" proof (simp only: atomize_imp, rule measure_induct_rule[of "\(u, t, s). {#size u, size t, size s#}" "\(u, t, s). wary u \ wary t \ wary s \ u >\<^sub>t t \ t >\<^sub>t s \ u >\<^sub>t s" "(u, t, s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric]) fix u t s assume ih: "\ua ta sa. {#size ua, size ta, size sa#} < {#size u, size t, size s#} \ wary ua \ wary ta \ wary sa \ ua >\<^sub>t ta \ ta >\<^sub>t sa \ ua >\<^sub>t sa" and wary_u: "wary u" and wary_t: "wary t" and wary_s: "wary s" and u_gt_t: "u >\<^sub>t t" and t_gt_s: "t >\<^sub>t s" have "vars_mset u \# vars_mset t" and "vars_mset t \# vars_mset s" using u_gt_t t_gt_s by (auto simp: gt_imp_vars_mset) hence vars_u_s: "vars_mset u \# vars_mset s" by auto have wt_u_ge_t: "wt u \ wt t" and wt_t_ge_s: "wt t \ wt s" using gt_imp_wt_ge u_gt_t t_gt_s by auto { assume wt_t_s: "wt t = wt s" and wt_u_t: "wt u = wt t" hence wt_u_s: "wt u = wt s" by simp have wary_arg_u: "wary (arg u)" by (rule wary_arg[OF wary_u]) have wary_arg_t: "wary (arg t)" by (rule wary_arg[OF wary_t]) have wary_arg_s: "wary (arg s)" by (rule wary_arg[OF wary_s]) have "u >\<^sub>t s" using t_gt_s proof cases case gt_unary_t_s: gt_unary have t_app: "is_App t" by (metis args_Nil_iff_is_Hd gt_unary_t_s(3) length_greater_0_conv less_numeral_extra(1)) have \_eq_\: "\ = \" using gt_unary_t_s(4) unary_wt_sym_0_imp_\_eq_\ by blast show ?thesis using u_gt_t proof cases case gt_unary_u_t: gt_unary have u_app: "is_App u" by (metis args_Nil_iff_is_Hd gt_unary_u_t(3) length_greater_0_conv less_numeral_extra(1)) hence arg_u_gt_s: "arg u >\<^sub>t s" using ih[of "arg u" t s] gt_unary_u_t(5) t_gt_s size_arg_lt wary_arg_u wary_s wary_t by force hence arg_u_ge_s: "arg u \\<^sub>t s" by sat { assume "size (arg u) < size t" hence ?thesis using ih[of u "arg u" s] arg_u_gt_s gt_arg by (simp add: u_app wary_arg_u wary_s wary_u) } moreover { assume "size (arg t) < size s" hence "u >\<^sub>t arg t" using ih[of u t "arg t"] args_Nil_iff_is_Hd gt_arg gt_unary_t_s(3) u_gt_t wary_t wary_u by force hence ?thesis using ih[of u "arg t" s] args_Nil_iff_is_Hd gt_unary_t_s(3,5) size_arg_lt wary_arg_t wary_s wary_u by force } moreover { assume sz_u_gt_t: "size u > size t" and sz_t_gt_s: "size t > size s" have wt_fun_u: "wt (fun u) = \" by (metis antisym gt_imp_wt_ge gt_unary_u_t(5) tm.collapse(2) u_app wt_App_\ wt_arg_le wt_t_s wt_u_s) have nargs_fun_u: "num_args (fun u) = 0" by (metis args.simps(1) gt_unary_u_t(3) list.size(3) one_arg_imp_Hd tm.collapse(2) u_app) { assume hd_u_eq_s: "head u = head s" hence ary_hd_s: "arity_hd (head s) = 1" using ground_heads_arity gt_unary_u_t(3,4) hd_u_eq_s one_enat_def wary_num_args_le_arity_head wary_u by fastforce have extf: "extf f (>\<^sub>t) (args u) (args s)" for f proof (cases "args s") case Nil thus ?thesis by (metis Hd_head_id \_eq_\ append_Nil args.simps(2) extf_singleton_nil_if_\_eq_\ gt_unary_u_t(3) head_fun length_greater_0_conv less_irrefl_nat nargs_fun_u tm.exhaust_sel zero_neq_one) next case args_s: (Cons sa ss) hence ss: "ss = []" by (cases s, simp, metis One_nat_def antisym_conv ary_hd_s diff_Suc_1 enat_ord_simps(1) le_add2 length_0_conv length_Cons list.size(4) one_enat_def wary_num_args_le_arity_head wary_s) have sa: "sa = arg s" by (metis apps.simps(1) apps.simps(2) args_s tm.sel(6) tm_collapse_apps ss) have s_app: "is_App s" using args_Nil_iff_is_Hd args_s by force have args_u: "args u = [arg u]" by (metis append_Nil args.simps(2) args_Nil_iff_is_Hd gt_unary_u_t(3) length_0_conv nargs_fun_u tm.collapse(2) zero_neq_one) have max_sz_u_t_s: "Max {size s, size t, size u} = size u" using sz_t_gt_s sz_u_gt_t by auto have max_sz_arg_u_t_arg_t: "Max {size (arg t), size t, size (arg u)} < size u" using size_arg_lt sz_u_gt_t t_app u_app by fastforce have "{#size (arg u), size t, size (arg t)#} < {#size u, size t, size s#}" using max_sz_arg_u_t_arg_t by (simp add: Max_lt_imp_lt_mset insert_commute max_sz_u_t_s) hence arg_u_gt_arg_t: "arg u >\<^sub>t arg t" using ih[OF _ wary_arg_u wary_t wary_arg_t] args_Nil_iff_is_Hd gt_arg gt_unary_t_s(3) gt_unary_u_t(5) wary_t by force have max_sz_arg_s_s_arg_t: "Max {size (arg s), size s, size (arg t)} < size u" using s_app t_app size_arg_lt sz_t_gt_s sz_u_gt_t by force have "{#size (arg t), size s, size (arg s)#} < {#size u, size t, size s#}" by (meson add_mset_lt_lt_lt less_trans mset_lt_single_iff s_app size_arg_lt sz_t_gt_s sz_u_gt_t t_app) hence arg_t_gt_arg_s: "arg t >\<^sub>t arg s" using ih[OF _ wary_arg_t wary_s wary_arg_s] gt_unary_t_s(5) gt_arg args_Nil_iff_is_Hd args_s wary_s by force have "arg u >\<^sub>t arg s" using ih[of "arg u" "arg t" "arg s"] arg_u_gt_arg_t arg_t_gt_arg_s by (simp add: add_mset_lt_le_lt less_imp_le_nat s_app size_arg_lt t_app u_app wary_arg_s wary_arg_t wary_arg_u) thus ?thesis unfolding args_u args_s ss sa by (metis extf_singleton gt_irrefl wary_arg_u) qed have ?thesis by (rule gt_same[OF vars_u_s wt_u_s hd_u_eq_s]) (simp add: extf) } moreover { assume "head u >\<^sub>h\<^sub>d head s" hence ?thesis by (rule gt_diff[OF vars_u_s wt_u_s]) } moreover { assume "head s >\<^sub>h\<^sub>d head u" hence False using gt_hd_def gt_hd_irrefl gt_sym_antisym gt_unary_u_t(4) unary_wt_sym_0_gt by blast } moreover { assume "\ head u \\\<^sub>h\<^sub>d head s" hence ?thesis by (rule gt_unary[OF wt_u_s _ gt_unary_u_t(3,4) arg_u_ge_s]) } ultimately have ?thesis unfolding comp_hd_def by sat } ultimately show ?thesis by (metis args_Nil_iff_is_Hd dual_order.strict_trans2 gt_unary_t_s(3) gt_unary_u_t(3) length_0_conv not_le_imp_less size_arg_lt zero_neq_one) next case gt_diff_u_t: gt_diff have False using gt_diff_u_t(3) gt_hd_def gt_hd_irrefl gt_sym_antisym gt_unary_t_s(4) unary_wt_sym_0_gt by blast thus ?thesis by sat next case gt_same_u_t: gt_same have hd_u_ncomp_s: "\ head u \\\<^sub>h\<^sub>d head s" by (rule gt_unary_t_s(2)[folded gt_same_u_t(3)]) have "num_args u \ 1" by (metis enat_ord_simps(1) ground_heads_arity gt_same_u_t(3) gt_unary_t_s(4) one_enat_def order_trans wary_num_args_le_arity_head wary_u) hence nargs_u: "num_args u = 1" by (cases "args u", metis Hd_head_id \_eq_\ append_Nil args.simps(2) gt_same_u_t(3,4) gt_unary_t_s(3,4) head_fun list.size(3) not_extf_gt_nil_singleton_if_\_eq_\ one_arg_imp_Hd tm.collapse(2)[OF t_app] wary_arg_t, simp) have "arg u >\<^sub>t arg t" by (metis extf_singleton[THEN iffD1] append_Nil args.simps args_Nil_iff_is_Hd comp_hd_def gt_hd_def gt_irrefl gt_same_u_t(3,4) gt_unary_t_s(2,3) head_fun length_0_conv nargs_u one_arg_imp_Hd t_app tm.collapse(2) u_gt_t wary_u) hence "arg u >\<^sub>t s" using ih[OF _ wary_arg_u wary_arg_t wary_s] gt_unary_t_s(5) by (metis add_mset_lt_left_lt add_mset_lt_lt_lt args_Nil_iff_is_Hd list.size(3) nargs_u size_arg_lt t_app zero_neq_one) hence arg_u_ge_s: "arg u \\<^sub>t s" by sat show ?thesis by (rule gt_unary[OF wt_u_s hd_u_ncomp_s nargs_u _ arg_u_ge_s]) (simp add: gt_same_u_t(3) gt_unary_t_s(4)) qed (simp add: wt_u_t) next case gt_diff_t_s: gt_diff show ?thesis using u_gt_t proof cases case gt_unary_u_t: gt_unary have "is_App u" by (metis args_Nil_iff_is_Hd gt_unary_u_t(3) length_greater_0_conv less_numeral_extra(1)) hence "arg u >\<^sub>t s" using ih[of "arg u" t s] gt_unary_u_t(5) t_gt_s size_arg_lt wary_arg_u wary_s wary_t by force hence arg_u_ge_s: "arg u \\<^sub>t s" by sat { assume "head u = head s" hence False using gt_diff_t_s(3) gt_unary_u_t(2) unfolding comp_hd_def by force } moreover { assume "head s >\<^sub>h\<^sub>d head u" hence False using gt_hd_def gt_hd_irrefl gt_sym_antisym gt_unary_u_t(4) unary_wt_sym_0_gt by blast } moreover { assume "head u >\<^sub>h\<^sub>d head s" hence ?thesis by (rule gt_diff[OF vars_u_s wt_u_s]) } moreover { assume "\ head u \\\<^sub>h\<^sub>d head s" hence ?thesis by (rule gt_unary[OF wt_u_s _ gt_unary_u_t(3,4) arg_u_ge_s]) } ultimately show ?thesis unfolding comp_hd_def by sat next case gt_diff_u_t: gt_diff have "head u >\<^sub>h\<^sub>d head s" using gt_diff_u_t(3) gt_diff_t_s(3) gt_hd_trans by blast thus ?thesis by (rule gt_diff[OF vars_u_s wt_u_s]) next case gt_same_u_t: gt_same have "head u >\<^sub>h\<^sub>d head s" using gt_diff_t_s(3) gt_same_u_t(3) by simp thus ?thesis by (rule gt_diff[OF vars_u_s wt_u_s]) qed (simp add: wt_u_t) next case gt_same_t_s: gt_same show ?thesis using u_gt_t proof cases case gt_unary_u_t: gt_unary have "is_App u" by (metis args_Nil_iff_is_Hd gt_unary_u_t(3) length_greater_0_conv less_numeral_extra(1)) hence "arg u >\<^sub>t s" using ih[of "arg u" t s] gt_unary_u_t(5) t_gt_s size_arg_lt wary_arg_u wary_s wary_t by force hence arg_u_ge_s: "arg u \\<^sub>t s" by sat have "\ head u \\\<^sub>h\<^sub>d head s" using gt_same_t_s(3) gt_unary_u_t(2) by simp thus ?thesis by (rule gt_unary[OF wt_u_s _ gt_unary_u_t(3,4) arg_u_ge_s]) next case gt_diff_u_t: gt_diff have "head u >\<^sub>h\<^sub>d head s" using gt_diff_u_t(3) gt_same_t_s(3) by simp thus ?thesis by (rule gt_diff[OF vars_u_s wt_u_s]) next case gt_same_u_t: gt_same have hd_u_s: "head u = head s" by (simp only: gt_same_t_s(3) gt_same_u_t(3)) let ?S = "set (args u) \ set (args t) \ set (args s)" have gt_trans_args: "\ua \ ?S. \ta \ ?S. \sa \ ?S. ua >\<^sub>t ta \ ta >\<^sub>t sa \ ua >\<^sub>t sa" proof clarify fix sa ta ua assume ua_in: "ua \ ?S" and ta_in: "ta \ ?S" and sa_in: "sa \ ?S" and ua_gt_ta: "ua >\<^sub>t ta" and ta_gt_sa: "ta >\<^sub>t sa" have wary_sa: "wary sa" and wary_ta: "wary ta" and wary_ua: "wary ua" using wary_args ua_in ta_in sa_in wary_u wary_t wary_s by blast+ show "ua >\<^sub>t sa" by (auto intro!: ih[OF Max_lt_imp_lt_mset wary_ua wary_ta wary_sa ua_gt_ta ta_gt_sa]) (meson ua_in ta_in sa_in Un_iff max.strict_coboundedI1 max.strict_coboundedI2 size_in_args)+ qed have "\f \ ground_heads (head u). extf f (>\<^sub>t) (args u) (args s)" by (clarify, rule extf_trans_from_irrefl[of ?S _ "args t", OF _ _ _ _ _ gt_trans_args]) (auto simp: gt_same_u_t(3,4) gt_same_t_s(4) wary_args wary_u wary_t wary_s gt_irrefl) thus ?thesis by (rule gt_same[OF vars_u_s wt_u_s hd_u_s]) qed (simp add: wt_u_t) qed (simp add: wt_t_s) } thus "u >\<^sub>t s" using vars_u_s wt_t_ge_s wt_u_ge_t by (force intro: gt_wt) qed lemma gt_antisym: "wary s \ wary t \ t >\<^sub>t s \ \ s >\<^sub>t t" using gt_irrefl gt_trans by blast subsection \Subterm Property\ lemma gt_sub_fun: "App s t >\<^sub>t s" proof (cases "wt (App s t) > wt s") case True thus ?thesis using gt_wt by simp next case False hence wt_st: "wt (App s t) = wt s" by (meson order.antisym not_le_imp_less wt_App_ge_fun) hence \_eq_\: "\ = \" by (metis add_diff_cancel_left' diff_diff_cancel wt_\_imp_\_eq_\ wt_ge_\ wt.simps(2)) have vars_st: "vars_mset (App s t) \# vars_mset s" by auto have hd_st: "head (App s t) = head s" by auto have extf: "\f \ ground_heads (head (App s t)). extf f (>\<^sub>t) (args (App s t)) (args s)" by (simp add: \_eq_\ extf_snoc_if_\_eq_\) show ?thesis by (rule gt_same[OF vars_st wt_st hd_st extf]) qed theorem gt_proper_sub: "wary t \ proper_sub s t \ t >\<^sub>t s" by (induct t) (auto intro: gt_sub_fun gt_sub_arg gt_trans sub.intros wary_sub) subsection \Compatibility with Functions\ theorem gt_compat_fun: assumes wary_t: "wary t" and t'_gt_t: "t' >\<^sub>t t" shows "App s t' >\<^sub>t App s t" proof - have vars_st': "vars_mset (App s t') \# vars_mset (App s t)" by (simp add: t'_gt_t gt_imp_vars_mset) show ?thesis proof (cases "wt t' > wt t") case True hence wt_st': "wt (App s t') > wt (App s t)" by (simp only: wt.simps) show ?thesis by (rule gt_wt[OF vars_st' wt_st']) next case False hence "wt t' = wt t" using t'_gt_t gt_imp_wt_ge order.not_eq_order_implies_strict by fastforce hence wt_st': "wt (App s t') = wt (App s t)" by (simp only: wt.simps) have head_st': "head (App s t') = head (App s t)" by simp have extf: "\f. extf f (>\<^sub>t) (args s @ [t']) (args s @ [t])" using t'_gt_t by (metis extf_compat_list gt_irrefl[OF wary_t]) show ?thesis by (rule gt_same[OF vars_st' wt_st' head_st']) (simp add: extf) qed qed subsection \Compatibility with Arguments\ theorem gt_compat_arg: assumes wary_s't: "wary (App s' t)" and s'_gt_s: "s' >\<^sub>t s" shows "App s' t >\<^sub>t App s t" proof - have vars_s't: "vars_mset (App s' t) \# vars_mset (App s t)" by (simp add: s'_gt_s gt_imp_vars_mset) show ?thesis using s'_gt_s proof cases case gt_wt_s'_s: gt_wt have "wt (App s' t) > wt (App s t)" by (simp add: wt_ge_\) (metis add_diff_assoc add_less_cancel_right gt_wt_s'_s(2) wt_ge_\) thus ?thesis by (rule gt_wt[OF vars_s't]) next case gt_unary_s'_s: gt_unary have False by (metis ground_heads_arity gt_unary_s'_s(3) gt_unary_s'_s(4) leD one_enat_def wary_AppE wary_s't) thus ?thesis by sat next case _: gt_diff thus ?thesis by (simp add: gt_diff) next case gt_same_s'_s: gt_same have wt_s't: "wt (App s' t) = wt (App s t)" by (simp add: gt_same_s'_s(2)) have hd_s't: "head (App s' t) = head (App s t)" by (simp add: gt_same_s'_s(3)) have "\f \ ground_heads (head (App s' t)). extf f (>\<^sub>t) (args (App s' t)) (args (App s t))" using gt_same_s'_s(4) by (auto intro: extf_compat_append_right) thus ?thesis by (rule gt_same[OF vars_s't wt_s't hd_s't]) qed qed subsection \Stability under Substitution\ definition extra_wt :: "('v \ ('s, 'v) tm) \ ('s, 'v) tm \ nat" where "extra_wt \ s = (\x \# vars_mset s. wt (\ x) - wt (Hd (Var x)))" lemma extra_wt_Var[simp]: "extra_wt \ (Hd (Var x)) = wt (\ x) - wt (Hd (Var x))" and extra_wt_Sym[simp]: "extra_wt \ (Hd (Sym f)) = 0" and extra_wt_App[simp]: "extra_wt \ (App s t) = extra_wt \ s + extra_wt \ t" unfolding extra_wt_def by simp+ lemma extra_wt_subseteq: assumes vars_s: "vars_mset t \# vars_mset s" shows "extra_wt \ t \ extra_wt \ s" proof (unfold extra_wt_def) let ?diff = "\v. wt (\ v) - wt (Hd (Var v))" have "vars_mset s + (vars_mset t - vars_mset s) = vars_mset t" using vars_s by (meson subset_mset.add_diff_inverse) hence "{#?diff v. v \# vars_mset t#} = {#?diff v. v \# vars_mset s#} + {#?diff v. v \# vars_mset t - vars_mset s#}" by (metis image_mset_union) thus "(\v \# vars_mset t. ?diff v) \ (\v \# vars_mset s. ?diff v)" by simp qed lemma wt_subst: assumes wary_\: "wary_subst \" and wary_s: "wary s" shows "wt (subst \ s) = wt s + extra_wt \ s" using wary_s proof (induct s rule: tm.induct) case \: (Hd \) show ?case proof (cases \) case x: (Var x) let ?\ = "head (\ x)" obtain g where g_in_grs_\: "g \ ground_heads ?\" and wt_\: "wt (Hd ?\) = wt_sym g + the_enat (\ * arity_sym g)" using exists_wt_sym by blast have "g \ ground_heads \" using x g_in_grs_\ wary_\ wary_subst_def by auto hence wt_\x_ge: "wt (\ x) \ wt (Hd \)" by (metis (full_types) dual_order.trans wt_le_wt_sym wt_\ wt_hd_le) thus ?thesis using x by (simp add: extra_wt_def) qed auto next case (App s t) note ih_s = this(1) and ih_t = this(2) and wary_st = this(3) have "wary s" using wary_st by (meson wary_AppE) hence "\n. extra_wt \ s + (wt s - \ + n) = wt (subst \ s) - \ + n" using ih_s by (metis (full_types) add_diff_assoc2 ab_semigroup_add_class.add_ac(1) add.left_commute wt_ge_\) hence "extra_wt \ s + (wt s + wt t - \ + extra_wt \ t) = wt (subst \ s) + wt (subst \ t) - \" using ih_t wary_st by (metis (no_types) add_diff_assoc2 ab_semigroup_add_class.add_ac(1) wary_AppE wt_ge_\) thus ?case by (simp add: wt_ge_\) qed theorem gt_subst: assumes wary_\: "wary_subst \" shows "wary t \ wary s \ t >\<^sub>t s \ subst \ t >\<^sub>t subst \ s" proof (simp only: atomize_imp, rule measure_induct_rule[of "\(t, s). {#size t, size s#}" "\(t, s). wary t \ wary s \ t >\<^sub>t s \ subst \ t >\<^sub>t subst \ s" "(t, s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric]) fix t s assume ih: "\ta sa. {#size ta, size sa#} < {#size t, size s#} \ wary ta \ wary sa \ ta >\<^sub>t sa \ subst \ ta >\<^sub>t subst \ sa" and wary_t: "wary t" and wary_s: "wary s" and t_gt_s: "t >\<^sub>t s" show "subst \ t >\<^sub>t subst \ s" proof (cases "wt (subst \ t) = wt (subst \ s)") case wt_\t_ne_\s: False have vars_s: "vars_mset t \# vars_mset s" by (simp add: t_gt_s gt_imp_vars_mset) hence vars_\s: "vars_mset (subst \ t) \# vars_mset (subst \ s)" by (rule vars_mset_subst_subseteq) have wt_t_ge_s: "wt t \ wt s" by (simp add: gt_imp_wt_ge t_gt_s) have "wt (subst \ t) > wt (subst \ s)" using wt_\t_ne_\s unfolding wt_subst[OF wary_\ wary_s] wt_subst[OF wary_\ wary_t] by (metis add_le_cancel_left add_less_le_mono extra_wt_subseteq order.not_eq_order_implies_strict vars_s wt_t_ge_s) thus ?thesis by (rule gt_wt[OF vars_\s]) next case wt_\t_eq_\s: True show ?thesis using t_gt_s proof cases case gt_wt hence False using wt_\t_eq_\s wary_s wary_t by (metis add_diff_cancel_right' diff_le_mono2 extra_wt_subseteq wt_subst leD wary_\) thus ?thesis by sat next case gt_unary have wary_\t: "wary (subst \ t)" by (simp add: wary_subst_wary wary_t wary_\) show ?thesis proof (cases t) case Hd hence False using gt_unary(3) by simp thus ?thesis by sat next case t: (App t1 t2) hence t2: "t2 = arg t" by simp hence wary_t2: "wary t2" using wary_t by blast show ?thesis proof (cases "t2 = s") case True moreover have "subst \ t >\<^sub>t subst \ t2" using gt_sub_arg wary_\t unfolding t by simp ultimately show ?thesis by simp next case t2_ne_s: False hence t2_gt_s: "t2 >\<^sub>t s" using gt_unary(5) t2 by blast have "subst \ t2 >\<^sub>t subst \ s" by (rule ih[OF _ wary_t2 wary_s t2_gt_s]) (simp add: t) thus ?thesis by (metis gt_sub_arg gt_trans subst.simps(2) t wary_\ wary_\t wary_s wary_subst_wary wary_t2) qed qed next case _: gt_diff note vars_s = this(1) and hd_t_gt_hd_s = this(3) have vars_\s: "vars_mset (subst \ t) \# vars_mset (subst \ s)" by (rule vars_mset_subst_subseteq[OF vars_s]) have "head (subst \ t) >\<^sub>h\<^sub>d head (subst \ s)" by (meson hd_t_gt_hd_s wary_subst_ground_heads gt_hd_def rev_subsetD wary_\) thus ?thesis by (rule gt_diff[OF vars_\s wt_\t_eq_\s]) next case _: gt_same note vars_s = this(1) and hd_s_eq_hd_t = this(3) and extf = this(4) have vars_\s: "vars_mset (subst \ t) \# vars_mset (subst \ s)" by (rule vars_mset_subst_subseteq[OF vars_s]) have hd_\t: "head (subst \ t) = head (subst \ s)" by (simp add: hd_s_eq_hd_t) { fix f assume f_in_grs: "f \ ground_heads (head (subst \ t))" let ?S = "set (args t) \ set (args s)" have extf_args_s_t: "extf f (>\<^sub>t) (args t) (args s)" using extf f_in_grs wary_subst_ground_heads wary_\ by blast have "extf f (>\<^sub>t) (map (subst \) (args t)) (map (subst \) (args s))" proof (rule extf_map[of ?S, OF _ _ _ _ _ _ extf_args_s_t]) show "\x \ ?S. \ subst \ x >\<^sub>t subst \ x" using gt_irrefl wary_t wary_s wary_args wary_\ wary_subst_wary by fastforce next show "\z \ ?S. \y \ ?S. \x \ ?S. subst \ z >\<^sub>t subst \ y \ subst \ y >\<^sub>t subst \ x \ subst \ z >\<^sub>t subst \ x" using gt_trans wary_t wary_s wary_args wary_\ wary_subst_wary by (metis Un_iff) next have sz_a: "\ta \ ?S. \sa \ ?S. {#size ta, size sa#} < {#size t, size s#}" by (fastforce intro: Max_lt_imp_lt_mset dest: size_in_args) show "\y \ ?S. \x \ ?S. y >\<^sub>t x \ subst \ y >\<^sub>t subst \ x" using ih sz_a size_in_args wary_t wary_s wary_args wary_\ wary_subst_wary by fastforce qed auto hence "extf f (>\<^sub>t) (args (subst \ t)) (args (subst \ s))" by (auto simp: hd_s_eq_hd_t intro: extf_compat_append_left) } hence "\f \ ground_heads (head (subst \ t)). extf f (>\<^sub>t) (args (subst \ t)) (args (subst \ s))" by blast thus ?thesis by (rule gt_same[OF vars_\s wt_\t_eq_\s hd_\t]) qed qed qed subsection \Totality on Ground Terms\ theorem gt_total_ground: - assumes - extf_total: "\f. ext_total (extf f)" and - gr_t: "ground t" and - gr_s: "ground s" - shows "t >\<^sub>t s \ s >\<^sub>t t \ t = s" - using gr_t gr_s -proof (induct t arbitrary: s rule: tm_induct_apps) - case t: (apps \ ts) - note ih = this(1) and gr_t = this(2) and gr_s = this(3) + assumes extf_total: "\f. ext_total (extf f)" + shows "ground t \ ground s \ t >\<^sub>t s \ s >\<^sub>t t \ t = s" +proof (simp only: atomize_imp, + rule measure_induct_rule[of "\(t, s). {# size t, size s #}" + "\(t, s). ground t \ ground s \ t >\<^sub>t s \ s >\<^sub>t t \ t = s" "(t, s)", simplified prod.case], + simp only: split_paired_all prod.case atomize_imp[symmetric]) + fix t s :: "('s, 'v) tm" + assume + ih: "\ta sa. {# size ta, size sa #} < {# size t, size s #} \ ground ta \ ground sa \ + ta >\<^sub>t sa \ sa >\<^sub>t ta \ ta = sa" and + gr_t: "ground t" and gr_s: "ground s" - let ?t = "apps (Hd \) ts" + let ?case = "t >\<^sub>t s \ s >\<^sub>t t \ t = s" have - vars_t: "vars_mset ?t \# vars_mset s" and - vars_s: "vars_mset s \# vars_mset ?t" + vars_t: "vars_mset t \# vars_mset s" and + vars_s: "vars_mset s \# vars_mset t" by (simp only: vars_mset_empty_iff[THEN iffD2, OF gr_s] vars_mset_empty_iff[THEN iffD2, OF gr_t])+ show ?case - proof (cases "wt ?t = wt s") + proof (cases "wt t = wt s") case False moreover { - assume "wt ?t > wt s" - hence "?t >\<^sub>t s" + assume "wt t > wt s" + hence "t >\<^sub>t s" by (rule gt_wt[OF vars_t]) } moreover { - assume "wt s > wt ?t" - hence "s >\<^sub>t ?t" + assume "wt s > wt t" + hence "s >\<^sub>t t" by (rule gt_wt[OF vars_s]) } ultimately show ?thesis by linarith next case wt_t: True note wt_s = wt_t[symmetric] - show ?thesis - proof (cases s rule: tm_exhaust_apps) - case s: (apps \ ss) - obtain g where \: "\ = Sym g" - by (metis ground_head[OF gr_t] hd.collapse(2) head_apps tm.sel(1)) - obtain f where \: "\ = Sym f" - using s by (metis ground_head[OF gr_s] hd.collapse(2) head_apps tm.sel(1)) - - { - assume g_gt_f: "g >\<^sub>s f" - have "?t >\<^sub>t s" - by (rule gt_diff[OF vars_t wt_t]) (simp add: \ \ s g_gt_f gt_hd_def) - } - moreover - { - assume f_gt_g: "f >\<^sub>s g" - have "s >\<^sub>t ?t" - by (rule gt_diff[OF vars_s wt_s]) (simp add: \ \ s f_gt_g gt_hd_def) - } - moreover - { - assume g_eq_f: "g = f" - hence hd_t: "head ?t = head s" - using \ \ t s by force - note hd_s = hd_t[symmetric] + obtain g where \: "head t = Sym g" + by (metis ground_head[OF gr_t] hd.collapse(2)) + obtain f where \: "head s = Sym f" + by (metis ground_head[OF gr_s] hd.collapse(2)) - have gr_ts: "\t \ set ts. ground t" - using gr_t by auto - have gr_ss: "\s \ set ss. ground s" - using gr_s s by auto + { + assume g_gt_f: "g >\<^sub>s f" + have "t >\<^sub>t s" + by (rule gt_diff[OF vars_t wt_t]) (simp add: \ \ g_gt_f gt_hd_def) + } + moreover + { + assume f_gt_g: "f >\<^sub>s g" + have "s >\<^sub>t t" + by (rule gt_diff[OF vars_s wt_s]) (simp add: \ \ f_gt_g gt_hd_def) + } + moreover + { + assume g_eq_f: "g = f" + hence hd_t: "head t = head s" + using \ \ by force + note hd_s = hd_t[symmetric] - have ?thesis - proof (cases "ts = ss") - case ts_eq_ss: True - show ?thesis - using s \ \ g_eq_f ts_eq_ss by blast - next - case False - hence "extf g (>\<^sub>t) ts ss \ extf g (>\<^sub>t) ss ts" - using ih gr_ss gr_ts - ext_total.total[OF extf_total, rule_format, of "set ts" "set ss" "(>\<^sub>t)" ts ss g] - by blast - moreover - { - assume extf: "extf g (>\<^sub>t) ts ss" - have "?t >\<^sub>t s" - by (rule gt_same[OF vars_t wt_t hd_t]) (simp add: extf \ s) - } - moreover - { - assume extf: "extf g (>\<^sub>t) ss ts" - have "s >\<^sub>t ?t" - by (rule gt_same[OF vars_s wt_s hd_s]) (simp add: extf[unfolded g_eq_f] \ s) - } - ultimately show ?thesis - by sat - qed - } - ultimately show ?thesis - using gt_sym_total by blast - qed + have gr_ts: "\t \ set (args t). ground t" + using gr_t ground_args by auto + have gr_ss: "\s \ set (args s). ground s" + using gr_s ground_args by auto + + let ?ts = "args t" + let ?ss = "args s" + + have ?thesis + proof (cases "?ts = ?ss") + case ts_eq_ss: True + show ?thesis + using \ \ g_eq_f ts_eq_ss by (simp add: tm_expand_apps) + next + case False + hence "extf g (>\<^sub>t) (args t) ?ss \ extf g (>\<^sub>t) ?ss ?ts" + using ih gr_ss gr_ts + ext_total.total[OF extf_total, rule_format, of "set ?ts \ set ?ss" "(>\<^sub>t)" ?ts ?ss g] + by (metis Un_commute Un_iff in_lists_iff_set less_multiset_doubletons size_in_args sup_ge2) + moreover + { + assume extf: "extf g (>\<^sub>t) ?ts ?ss" + have "t >\<^sub>t s" + by (rule gt_same[OF vars_t wt_t hd_t]) (simp add: extf \) + } + moreover + { + assume extf: "extf g (>\<^sub>t) ?ss ?ts" + have "s >\<^sub>t t" + by (rule gt_same[OF vars_s wt_s hd_s]) (simp add: extf[unfolded g_eq_f] \) + } + ultimately show ?thesis + by sat + qed + } + ultimately show ?thesis + using gt_sym_total by blast qed qed subsection \Well-foundedness\ abbreviation gtw :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix ">\<^sub>t\<^sub>w" 50) where "(>\<^sub>t\<^sub>w) \ \t s. wary t \ wary s \ t >\<^sub>t s" abbreviation gtwg :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix ">\<^sub>t\<^sub>w\<^sub>g" 50) where "(>\<^sub>t\<^sub>w\<^sub>g) \ \t s. ground t \ t >\<^sub>t\<^sub>w s" lemma ground_gt_unary: assumes gr_t: "ground t" shows "\ gt_unary t s" proof assume gt_unary_t_s: "gt_unary t s" hence "t >\<^sub>t s" using gt_iff_wt_unary_diff_same by blast hence gr_s: "ground s" using gr_t gt_imp_vars by blast have ngr_t_or_s: "\ ground t \ \ ground s" using gt_unary_t_s by cases (blast dest: ground_head not_comp_hd_imp_Var) show False using gr_t gr_s ngr_t_or_s by sat qed theorem gt_wf: "wfP (\s t. t >\<^sub>t\<^sub>w s)" proof - have ground_wfP: "wfP (\s t. t >\<^sub>t\<^sub>w\<^sub>g s)" unfolding wfP_iff_no_inf_chain proof assume "\f. inf_chain (>\<^sub>t\<^sub>w\<^sub>g) f" then obtain t where t_bad: "bad (>\<^sub>t\<^sub>w\<^sub>g) t" unfolding inf_chain_def bad_def by blast let ?ff = "worst_chain (>\<^sub>t\<^sub>w\<^sub>g) (\t s. size t > size s)" note wf_sz = wf_app[OF wellorder_class.wf, of size, simplified] have ffi_ground: "\i. ground (?ff i)" and ffi_wary: "\i. wary (?ff i)" using worst_chain_bad[OF wf_sz t_bad, unfolded inf_chain_def] by fast+ have "inf_chain (>\<^sub>t\<^sub>w\<^sub>g) ?ff" by (rule worst_chain_bad[OF wf_sz t_bad]) hence bad_wt_diff_same: "inf_chain (\t s. ground t \ (gt_wt t s \ gt_diff t s \ gt_same t s)) ?ff" unfolding inf_chain_def using gt_iff_wt_unary_diff_same ground_gt_unary by blast have wf_wt: "wf {(s, t). ground t \ gt_wt t s}" by (rule wf_subset[OF wf_app[of _ wt, OF wf_less]]) (auto simp: gt_wt.simps) have wt_O_diff_same: "{(s, t). ground t \ gt_wt t s} O {(s, t). ground t \ (gt_diff t s \ gt_same t s)} \ {(s, t). ground t \ gt_wt t s}" unfolding gt_wt.simps gt_diff.simps gt_same.simps by auto have wt_diff_same_as_union: "{(s, t). ground t \ (gt_wt t s \ gt_diff t s \ gt_same t s)} = {(s, t). ground t \ gt_wt t s} \ {(s, t). ground t \ (gt_diff t s \ gt_same t s)}" by auto obtain k1 where bad_diff_same: "inf_chain (\t s. ground t \ (gt_diff t s \ gt_same t s)) (\i. ?ff (i + k1))" using wf_infinite_down_chain_compatible[OF wf_wt _ wt_O_diff_same, of ?ff] bad_wt_diff_same unfolding inf_chain_def wt_diff_same_as_union[symmetric] by auto have "wf {(s, t). ground s \ ground t \ sym (head t) >\<^sub>s sym (head s)}" using gt_sym_wf unfolding wfP_def wf_iff_no_infinite_down_chain by fast moreover have "{(s, t). ground t \ gt_diff t s} \ {(s, t). ground s \ ground t \ sym (head t) >\<^sub>s sym (head s)}" proof (clarsimp, intro conjI) fix s t assume gr_t: "ground t" and gt_diff_t_s: "gt_diff t s" thus gr_s: "ground s" using gt_iff_wt_unary_diff_same gt_imp_vars by fastforce show "sym (head t) >\<^sub>s sym (head s)" using gt_diff_t_s by cases (simp add: gt_hd_def gr_s gr_t ground_hd_in_ground_heads) qed ultimately have wf_diff: "wf {(s, t). ground t \ gt_diff t s}" by (rule wf_subset) have diff_O_same: "{(s, t). ground t \ gt_diff t s} O {(s, t). ground t \ gt_same t s} \ {(s, t). ground t \ gt_diff t s}" unfolding gt_diff.simps gt_same.simps by auto have diff_same_as_union: "{(s, t). ground t \ (gt_diff t s \ gt_same t s)} = {(s, t). ground t \ gt_diff t s} \ {(s, t). ground t \ gt_same t s}" by auto obtain k2 where bad_same: "inf_chain (\t s. ground t \ gt_same t s) (\i. ?ff (i + k2))" using wf_infinite_down_chain_compatible[OF wf_diff _ diff_O_same, of "\i. ?ff (i + k1)"] bad_diff_same unfolding inf_chain_def diff_same_as_union[symmetric] by (auto simp: add.assoc) hence hd_sym: "\i. is_Sym (head (?ff (i + k2)))" unfolding inf_chain_def by (simp add: ground_head) define f where "f = sym (head (?ff k2))" have hd_eq_f: "head (?ff (i + k2)) = Sym f" for i unfolding f_def proof (induct i) case 0 thus ?case by (auto simp: hd.collapse(2)[OF hd_sym, of 0, simplified]) next case (Suc ia) thus ?case using bad_same unfolding inf_chain_def gt_same.simps by simp qed define max_args where "max_args = wt (?ff k2)" have wt_eq_max_args: "wt (?ff (i + k2)) = max_args" for i unfolding max_args_def proof (induct i) case (Suc ia) thus ?case using bad_same unfolding inf_chain_def gt_same.simps by simp qed auto have nargs_le_max_args: "num_args (?ff (i + k2)) \ max_args" for i unfolding wt_eq_max_args[of i, symmetric] by (rule wt_ge_num_args[OF ffi_wary]) let ?U_of = "\i. set (args (?ff (i + k2)))" define U where "U = (\i. ?U_of i)" have gr_u: "\u. u \ U \ ground u" unfolding U_def by (blast dest: ground_args[OF _ ffi_ground]) have wary_u: "\u. u \ U \ wary u" unfolding U_def by (blast dest: wary_args[OF _ ffi_wary]) have "\ bad (>\<^sub>t\<^sub>w\<^sub>g) u" if u_in: "u \ ?U_of i" for u i proof assume u_bad: "bad (>\<^sub>t\<^sub>w\<^sub>g) u" have sz_u: "size u < size (?ff (i + k2))" by (rule size_in_args[OF u_in]) show False proof (cases "i + k2") case 0 thus False using sz_u min_worst_chain_0[OF wf_sz u_bad] by simp next case Suc hence gt: "?ff (i + k2 - 1) >\<^sub>t\<^sub>w ?ff (i + k2)" using worst_chain_pred[OF wf_sz t_bad] by auto moreover have "?ff (i + k2) >\<^sub>t\<^sub>w u" using gt gt_proper_sub sub_args sz_u u_in wary_args by auto ultimately have "?ff (i + k2 - 1) >\<^sub>t\<^sub>w u" using gt_trans by blast thus False using Suc sz_u min_worst_chain_Suc[OF wf_sz u_bad] ffi_ground by fastforce qed qed hence u_good: "\u. u \ U \ \ bad (>\<^sub>t\<^sub>w\<^sub>g) u" unfolding U_def by blast let ?gtwu = "\t s. t \ U \ t >\<^sub>t\<^sub>w s" have gtwu_irrefl: "\x. \ ?gtwu x x" using gt_irrefl by auto have "\i j. \t \ set (args (?ff (i + k2))). \s \ set (args (?ff (j + k2))). t >\<^sub>t s \ t \ U \ t >\<^sub>t\<^sub>w s" using wary_u unfolding U_def by blast moreover have "\i. extf f (>\<^sub>t) (args (?ff (i + k2))) (args (?ff (Suc i + k2)))" using bad_same hd_eq_f unfolding inf_chain_def gt_same.simps by auto ultimately have "\i. extf f ?gtwu (args (?ff (i + k2))) (args (?ff (Suc i + k2)))" by (rule extf_mono_strong) hence "inf_chain (extf f ?gtwu) (\i. args (?ff (i + k2)))" unfolding inf_chain_def by blast hence nwf_ext: "\ wfP (\xs ys. length ys \ max_args \ length xs \ max_args \ extf f ?gtwu ys xs)" unfolding inf_chain_def wfP_def wf_iff_no_infinite_down_chain using nargs_le_max_args by fast have gtwu_le_gtwg: "?gtwu \ (>\<^sub>t\<^sub>w\<^sub>g)" by (auto intro!: gr_u) have "wfP (\s t. ?gtwu t s)" unfolding wfP_iff_no_inf_chain proof (intro notI, elim exE) fix f assume bad_f: "inf_chain ?gtwu f" hence bad_f0: "bad ?gtwu (f 0)" by (rule inf_chain_bad) hence "f 0 \ U" using bad_f unfolding inf_chain_def by blast hence "\ bad (>\<^sub>t\<^sub>w\<^sub>g) (f 0)" using u_good by blast hence "\ bad ?gtwu (f 0)" using bad_f inf_chain_bad inf_chain_subset[OF _ gtwu_le_gtwg] by blast thus False using bad_f0 by sat qed hence wf_ext: "wfP (\xs ys. length ys \ max_args \ length xs \ max_args \ extf f ?gtwu ys xs)" using extf_wf_bounded[of ?gtwu] gtwu_irrefl by blast show False using nwf_ext wf_ext by blast qed let ?subst = "subst grounding_\" have "wfP (\s t. ?subst t >\<^sub>t\<^sub>w\<^sub>g ?subst s)" by (rule wfP_app[OF ground_wfP]) hence "wfP (\s t. ?subst t >\<^sub>t\<^sub>w ?subst s)" by (simp add: ground_grounding_\) thus ?thesis by (auto intro: wfP_subset wary_subst_wary[OF wary_grounding_\] gt_subst[OF wary_grounding_\]) qed end end diff --git a/thys/Lambda_Free_KBOs/Lambda_Free_TKBO_Coefs.thy b/thys/Lambda_Free_KBOs/Lambda_Free_TKBO_Coefs.thy --- a/thys/Lambda_Free_KBOs/Lambda_Free_TKBO_Coefs.thy +++ b/thys/Lambda_Free_KBOs/Lambda_Free_TKBO_Coefs.thy @@ -1,1997 +1,1999 @@ (* Title: The Graceful Transfinite Knuth-Bendix Order with Subterm Coefficients for Lambda-Free Higher-Order Terms Author: Heiko Becker , 2016 Author: Jasmin Blanchette , 2016 Author: Uwe Waldmann , 2016 Author: Daniel Wand , 2016 Maintainer: Jasmin Blanchette *) section \The Graceful Transfinite Knuth--Bendix Order with Subterm Coefficients for Lambda-Free Higher-Order Terms\ theory Lambda_Free_TKBO_Coefs imports Lambda_Free_KBO_Util Nested_Multisets_Ordinals.Signed_Syntactic_Ordinal abbrevs "=p" = "=\<^sub>p" and ">p" = ">\<^sub>p" and "\p" = "\\<^sub>p" and ">t" = ">\<^sub>t" and "\t" = "\\<^sub>t" and "!h" = "\<^sub>h" begin text \ This theory defines the graceful transfinite Knuth--Bendix order (KBO) with subterm coefficients for \\\-free higher-order terms. The proof was developed by copying that of the standard KBO and generalizing it along two axes:\ subterm coefficients and ordinals. Both features complicate the definitions and proofs substantially. \ subsection \Setup\ hide_const (open) Complex.arg locale tkbo_coefs = kbo_std_basis _ _ arity_sym arity_var wt_sym for arity_sym :: "'s \ enat" and arity_var :: "'v \ enat" and wt_sym :: "'s \ hmultiset" + fixes coef_sym :: "'s \ nat \ hmultiset" assumes coef_sym_gt_0: "coef_sym f i > 0" begin abbreviation \\<^sub>h :: hmultiset where "\\<^sub>h \ of_nat \" abbreviation \\<^sub>h :: hmultiset where "\\<^sub>h \ of_nat \" abbreviation arity_sym\<^sub>h :: "'s \ hmultiset" where "arity_sym\<^sub>h f \ hmset_of_enat (arity_sym f)" abbreviation arity_var\<^sub>h :: "'v \ hmultiset" where "arity_var\<^sub>h f \ hmset_of_enat (arity_var f)" abbreviation arity_hd\<^sub>h :: "('s, 'v) hd \ hmultiset" where "arity_hd\<^sub>h f \ hmset_of_enat (arity_hd f)" abbreviation arity\<^sub>h :: "('s, 'v) tm \ hmultiset" where "arity\<^sub>h s \ hmset_of_enat (arity s)" lemma arity\<^sub>h_conv: "arity\<^sub>h s = arity_hd\<^sub>h (head s) - of_nat (num_args s)" unfolding arity_def by simp lemma arity\<^sub>h_App[simp]: "arity\<^sub>h (App s t) = arity\<^sub>h s - 1" by (simp add: one_enat_def) lemmas wary_App\<^sub>h[intro] = wary_App[folded of_nat_lt_hmset_of_enat_iff] lemmas wary_AppE\<^sub>h = wary_AppE[folded of_nat_lt_hmset_of_enat_iff] lemmas wary_num_args_le_arity_head\<^sub>h = wary_num_args_le_arity_head[folded of_nat_le_hmset_of_enat_iff] lemmas wary_apps\<^sub>h = wary_apps[folded of_nat_le_hmset_of_enat_iff] lemmas wary_cases_apps\<^sub>h[consumes 1, case_names apps] = wary_cases_apps[folded of_nat_le_hmset_of_enat_iff] lemmas ground_heads_arity\<^sub>h = ground_heads_arity[folded hmset_of_enat_le] lemmas some_ground_head_arity\<^sub>h = some_ground_head_arity[folded hmset_of_enat_le] lemmas \\<^sub>h_gt_0 = \_gt_0[folded of_nat_less_hmset, unfolded of_nat_0] lemmas \\<^sub>h_le_\\<^sub>h = \_le_\[folded of_nat_le_hmset] lemmas arity_hd\<^sub>h_lt_\_if_\\<^sub>h_gt_0 = arity_hd_ne_infinity_if_\_gt_0 [folded of_nat_less_hmset, unfolded of_nat_0, folded hmset_of_enat_lt_iff_ne_infinity] lemma wt_sym_ge\<^sub>h: "wt_sym f \ \\<^sub>h - \\<^sub>h * arity_sym\<^sub>h f" proof - have "of_nat (the_enat (of_nat \ * arity_sym f)) = \\<^sub>h * arity_sym\<^sub>h f" by (cases "arity_sym f", simp add: of_nat_eq_enat, metis arity_sym_ne_infinity_if_\_gt_0 gr_zeroI mult_eq_0_iff of_nat_0 the_enat_0) thus ?thesis using wt_sym_ge[unfolded of_nat_minus_hmset] by metis qed lemmas unary_wt_sym_0_gt\<^sub>h = unary_wt_sym_0_gt[folded hmset_of_enat_inject, unfolded hmset_of_enat_1] lemmas unary_wt_sym_0_imp_\\<^sub>h_eq_\\<^sub>h = unary_wt_sym_0_imp_\_eq_\ [folded of_nat_inject_hmset, unfolded of_nat_0] lemmas extf_ext_snoc_if_\\<^sub>h_eq_\\<^sub>h = extf_ext_snoc_if_\_eq_\[folded of_nat_inject_hmset] lemmas extf_snoc_if_\\<^sub>h_eq_\\<^sub>h = ext_snoc.snoc[OF extf_ext_snoc_if_\\<^sub>h_eq_\\<^sub>h] lemmas arity_sym\<^sub>h_lt_\_if_\\<^sub>h_gt_0 = arity_sym_ne_infinity_if_\_gt_0 [folded of_nat_less_hmset hmset_of_enat_lt_iff_ne_infinity, unfolded of_nat_0] lemmas arity_var\<^sub>h_lt_\_if_\\<^sub>h_gt_0 = arity_var_ne_infinity_if_\_gt_0 [folded of_nat_less_hmset hmset_of_enat_lt_iff_ne_infinity, unfolded of_nat_0] lemmas arity\<^sub>h_lt_\_if_\\<^sub>h_gt_0 = arity_ne_infinity_if_\_gt_0 [folded of_nat_less_hmset hmset_of_enat_lt_iff_ne_infinity, unfolded of_nat_0] lemmas warywary_subst_subst\<^sub>h_conv = wary_subst_def[folded hmset_of_enat_le] lemmas extf_singleton_nil_if_\\<^sub>h_eq_\\<^sub>h = extf_singleton_nil_if_\_eq_\[folded of_nat_inject_hmset] lemma arity_sym\<^sub>h_if_\\<^sub>h_gt_0_E: assumes \_gt_0: "\\<^sub>h > 0" obtains n where "arity_sym\<^sub>h f = of_nat n" using arity_sym\<^sub>h_lt_\_if_\\<^sub>h_gt_0 assms lt_\_imp_ex_of_nat by blast lemma arity_var\<^sub>h_if_\\<^sub>h_gt_0_E: assumes \_gt_0: "\\<^sub>h > 0" obtains n where "arity_var\<^sub>h f = of_nat n" using arity_var\<^sub>h_lt_\_if_\\<^sub>h_gt_0 assms lt_\_imp_ex_of_nat by blast subsection \Weights and Subterm Coefficients\ abbreviation zhmset_of_tpoly :: "('a, hmultiset) tpoly \ ('a, zhmultiset) tpoly" where "zhmset_of_tpoly \ map_tpoly (\x. x) zhmset_of" abbreviation eval_ztpoly :: "('a \ zhmultiset) \ ('a, hmultiset) tpoly \ zhmultiset" where "eval_ztpoly A p \ eval_tpoly A (zhmset_of_tpoly p)" lemma eval_tpoly_eq_eval_ztpoly[simp]: "zhmset_of (eval_tpoly A p) = eval_ztpoly (\v. zhmset_of (A v)) p" by (induct p, simp_all add: zhmset_of_sum_list zhmset_of_prod_list o_def, simp_all cong: map_cong) definition min_ground_head :: "('s, 'v) hd \ 's" where "min_ground_head \ = (SOME f. f \ ground_heads \ \ (\g \ ground_heads \. wt_sym g + \\<^sub>h * arity_sym\<^sub>h g \ wt_sym f + \\<^sub>h * arity_sym\<^sub>h f))" datatype 'va pvar = PWt 'va | PCoef 'va nat primrec min_passign :: "'v pvar \ hmultiset" where "min_passign (PWt x) = wt_sym (min_ground_head (Var x))" | "min_passign (PCoef _ _) = 1" abbreviation min_zpassign :: "'v pvar \ zhmultiset" where "min_zpassign v \ zhmset_of (min_passign v)" lemma min_zpassign_simps[simp]: "min_zpassign (PWt x) = zhmset_of (wt_sym (min_ground_head (Var x)))" "min_zpassign (PCoef x i) = 1" by (simp_all add: zhmset_of_1) definition legal_passign :: "('v pvar \ hmultiset) \ bool" where "legal_passign A \ (\x. A x \ min_passign x)" definition legal_zpassign :: "('v pvar \ zhmultiset) \ bool" where "legal_zpassign A \ (\x. A x \ min_zpassign x)" lemma legal_min_passign: "legal_passign min_passign" unfolding legal_passign_def by simp lemma legal_min_zpassign: "legal_zpassign min_zpassign" unfolding legal_zpassign_def by simp lemma assign_ge_0[intro]: "legal_zpassign A \ A x \ 0" unfolding legal_zpassign_def by (auto intro: dual_order.trans) definition eq_tpoly :: "('v pvar, hmultiset) tpoly \ ('v pvar, hmultiset) tpoly \ bool" (infix "=\<^sub>p" 50) where "q =\<^sub>p p \ (\A. legal_zpassign A \ eval_ztpoly A q = eval_ztpoly A p)" definition ge_tpoly :: "('v pvar, hmultiset) tpoly \ ('v pvar, hmultiset) tpoly \ bool" (infix "\\<^sub>p" 50) where "q \\<^sub>p p \ (\A. legal_zpassign A \ eval_ztpoly A q \ eval_ztpoly A p)" definition gt_tpoly :: "('v pvar, hmultiset) tpoly \ ('v pvar, hmultiset) tpoly \ bool" (infix ">\<^sub>p" 50) where "q >\<^sub>p p \ (\A. legal_zpassign A \ eval_ztpoly A q > eval_ztpoly A p)" lemma gt_tpoly_imp_ge[intro]: "q >\<^sub>p p \ q \\<^sub>p p" unfolding ge_tpoly_def gt_tpoly_def by (simp add: le_less) lemma eq_tpoly_refl[simp]: "p =\<^sub>p p" unfolding eq_tpoly_def by simp lemma ge_tpoly_refl[simp]: "p \\<^sub>p p" unfolding ge_tpoly_def by simp lemma gt_tpoly_irrefl: "\ p >\<^sub>p p" unfolding gt_tpoly_def legal_zpassign_def by fast lemma eq_eq_tpoly_trans: "r =\<^sub>p q \ q =\<^sub>p p \ r =\<^sub>p p" and eq_ge_tpoly_trans: "r =\<^sub>p q \ q \\<^sub>p p \ r \\<^sub>p p" and eq_gt_tpoly_trans: "r =\<^sub>p q \ q >\<^sub>p p \ r >\<^sub>p p" and ge_eq_tpoly_trans: "r \\<^sub>p q \ q =\<^sub>p p \ r \\<^sub>p p" and ge_ge_tpoly_trans: "r \\<^sub>p q \ q \\<^sub>p p \ r \\<^sub>p p" and ge_gt_tpoly_trans: "r \\<^sub>p q \ q >\<^sub>p p \ r >\<^sub>p p" and gt_eq_tpoly_trans: "r >\<^sub>p q \ q =\<^sub>p p \ r >\<^sub>p p" and gt_ge_tpoly_trans: "r >\<^sub>p q \ q \\<^sub>p p \ r >\<^sub>p p" and gt_gt_tpoly_trans: "r >\<^sub>p q \ q >\<^sub>p p \ r >\<^sub>p p" unfolding eq_tpoly_def ge_tpoly_def gt_tpoly_def by (auto intro: order.trans less_trans less_le_trans le_less_trans)+ primrec coef_hd :: "('s, 'v) hd \ nat \ ('v pvar, hmultiset) tpoly" where "coef_hd (Var x) i = PVar (PCoef x i)" | "coef_hd (Sym f) i = PNum (coef_sym f i)" lemma coef_hd_gt_0: assumes legal: "legal_zpassign A" shows "eval_ztpoly A (coef_hd \ i) > 0" unfolding legal_zpassign_def proof (cases \) case (Var x1) thus ?thesis using legal[unfolded legal_zpassign_def, rule_format, of "PCoef x i" for x] by (auto simp: coef_sym_gt_0 zhmset_of_1 intro: dual_order.strict_trans1 zero_less_one) next case (Sym x2) thus ?thesis using legal[unfolded legal_zpassign_def, rule_format, of "PWt x" for x] by simp (metis coef_sym_gt_0 zhmset_of_0 zhmset_of_less) qed primrec coef :: "('s, 'v) tm \ nat \ ('v pvar, hmultiset) tpoly" where "coef (Hd \) i = coef_hd \ i" | "coef (App s _) i = coef s (i + 1)" lemma coef_apps[simp]: "coef (apps s ss) i = coef s (i + length ss)" by (induct ss arbitrary: s i) auto lemma coef_gt_0: "legal_zpassign A \ eval_ztpoly A (coef s i) > 0" by (induct s arbitrary: i) (auto intro: coef_hd_gt_0) lemma exists_min_ground_head: "\f. f \ ground_heads \ \ (\g \ ground_heads \. wt_sym g + \\<^sub>h * arity_sym\<^sub>h g \ wt_sym f + \\<^sub>h * arity_sym\<^sub>h f)" proof - let ?R = "{(f, g). f \ ground_heads \ \ g \ ground_heads \ \ wt_sym g + \\<^sub>h * arity_sym\<^sub>h g > wt_sym f + \\<^sub>h * arity_sym\<^sub>h f}" have wf_R: "wf ?R" using wf_app[of "{(M, N). M < N}" "\f. wt_sym f + \\<^sub>h * arity_sym\<^sub>h f", OF wf] by (auto intro: wf_subset) have "\f. f \ ground_heads \" by (meson ground_heads_nonempty subsetI subset_empty) thus ?thesis using wf_eq_minimal[THEN iffD1, OF wf_R] by force qed lemma min_ground_head_Sym[simp]: "min_ground_head (Sym f) = f" unfolding min_ground_head_def by auto lemma min_ground_head_in_ground_heads: "min_ground_head \ \ ground_heads \" unfolding min_ground_head_def using someI_ex[OF exists_min_ground_head] by blast lemma min_ground_head_min: "f \ ground_heads \ \ wt_sym f + \\<^sub>h * arity_sym\<^sub>h f \ wt_sym (min_ground_head \) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head \)" unfolding min_ground_head_def using someI_ex[OF exists_min_ground_head] by blast lemma min_ground_head_antimono: "ground_heads \ \ ground_heads \ \ wt_sym (min_ground_head \) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head \) \ wt_sym (min_ground_head \) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head \)" using min_ground_head_in_ground_heads min_ground_head_min by blast primrec wt0 :: "('s, 'v) hd \ ('v pvar, hmultiset) tpoly" where "wt0 (Var x) = PVar (PWt x)" | "wt0 (Sym f) = PNum (wt_sym f)" lemma wt0_ge_min_ground_head: "legal_zpassign A \ eval_ztpoly A (wt0 \) \ zhmset_of (wt_sym (min_ground_head \))" by (cases \, simp_all, metis legal_zpassign_def min_zpassign_simps(1)) lemma eval_ztpoly_nonneg: "legal_zpassign A \ eval_ztpoly A p \ 0" by (induct p) (auto cong: map_cong intro!: sum_list_nonneg prod_list_nonneg) lemma in_zip_imp_size_lt_apps: "(s, y) \ set (zip ss ys) \ size s < size (apps (Hd \) ss)" by (auto dest!: set_zip_leftD simp: size_in_args) function wt :: "('s, 'v) tm \ ('v pvar, hmultiset) tpoly" where "wt (apps (Hd \) ss) = PSum ([wt0 \, PNum (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss)))] @ map (\(s, i). PMult [coef_hd \ i, wt s]) (zip ss [0.. ('v pvar \ zhmultiset) \ ('s, 'v) hd \ ('s, 'v) tm list \ zhmultiset" where "wt_args i A \ ss = sum_list (map (eval_ztpoly A \ (\(s, i). PMult [coef_hd \ i, wt s])) (zip ss [i..) = PSum [wt0 \, PNum (\\<^sub>h * arity_sym\<^sub>h (min_ground_head \))]" by (rule wt.simps[of _ "[]", simplified]) lemma coef_hd_cong: "(\x \ vars_hd \. \i. A (PCoef x i) = B (PCoef x i)) \ eval_ztpoly A (coef_hd \ i) = eval_ztpoly B (coef_hd \ i)" by (cases \) auto lemma wt0_cong: assumes pwt_eq: "\x \ vars_hd \. A (PWt x) = B (PWt x)" shows "eval_ztpoly A (wt0 \) = eval_ztpoly B (wt0 \)" using pwt_eq by (cases \) auto lemma wt_cong: assumes "\x \ vars s. A (PWt x) = B (PWt x)" and "\x \ vars s. \i. A (PCoef x i) = B (PCoef x i)" shows "eval_ztpoly A (wt s) = eval_ztpoly B (wt s)" using assms proof (induct s rule: tm_induct_apps) case (apps \ ss) note ih = this(1) and pwt_eq = this(2) and pcoef_eq = this(3) have ih': "eval_ztpoly A (wt s) = eval_ztpoly B (wt s)" if s_in: "s \ set ss" for s proof (rule ih[OF s_in]) show "\x \ vars s. A (PWt x) = B (PWt x)" using pwt_eq s_in by force show "\x \ vars s. \i. A (PCoef x i) = B (PCoef x i)" using pcoef_eq s_in by force qed have wt0_eq: "eval_ztpoly A (wt0 \) = eval_ztpoly B (wt0 \)" by (rule wt0_cong) (simp add: pwt_eq) have coef_\_eq: "eval_ztpoly A (coef_hd \ i) = eval_ztpoly B (coef_hd \ i)" for i by (rule coef_hd_cong) (simp add: pcoef_eq) show ?case using ih' wt0_eq coef_\_eq by (auto dest!: set_zip_leftD intro!: arg_cong[of _ _ sum_list]) qed lemma ground_eval_ztpoly_wt_eq: "ground s \ eval_ztpoly A (wt s) = eval_ztpoly B (wt s)" by (rule wt_cong) auto lemma exists_wt_sym: assumes legal: "legal_zpassign A" shows "\f \ ground_heads \. eval_ztpoly A (wt (Hd \)) \ zhmset_of (wt_sym f + \\<^sub>h * arity_sym\<^sub>h f)" unfolding eq_tpoly_def proof (cases \) case Var thus ?thesis using legal[unfolded legal_zpassign_def] by simp (metis add_le_cancel_right ground_heads.simps(1) min_ground_head_in_ground_heads min_zpassign_simps(1) zhmset_of_plus) next case Sym thus ?thesis by (simp add: zhmset_of_plus) qed lemma wt_ge_\\<^sub>h: assumes legal: "legal_zpassign A" shows "eval_ztpoly A (wt s) \ zhmset_of \\<^sub>h" proof (induct s rule: tm_induct_apps) case (apps \ ss) note ih = this(1) { assume ss_eq_nil: "ss = []" have "\\<^sub>h \ wt_sym (min_ground_head \) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head \)" using wt_sym_ge\<^sub>h[of "min_ground_head \"] by (metis add_diff_cancel_left' leD leI le_imp_minus_plus_hmset le_minus_plus_same_hmset less_le_trans) hence "zhmset_of \\<^sub>h \ zhmset_of (wt_sym (min_ground_head \)) + zhmset_of (\\<^sub>h * arity_sym\<^sub>h (min_ground_head \))" by (metis zhmset_of_le zhmset_of_plus) also have "\ \ eval_tpoly A (map_tpoly (\x. x) zhmset_of (wt0 \)) + zhmset_of (\\<^sub>h * arity_sym\<^sub>h (min_ground_head \))" using wt0_ge_min_ground_head[OF legal] by simp finally have ?case using ss_eq_nil by simp } moreover { let ?arg_wt = "eval_tpoly A \ (map_tpoly (\x. x) zhmset_of \ (\(s, i). PMult [coef_hd \ i, wt s]))" assume ss_ne_nil: "ss \ []" hence "zhmset_of \\<^sub>h \ eval_tpoly A (map_tpoly (\x. x) zhmset_of (PMult [coef_hd \ 0, wt (hd ss)]))" by (simp add: ih coef_hd_gt_0[OF legal] nonneg_le_mult_right_mono_zhmset) also have "\ = hd (map ?arg_wt (zip ss [0.. \ sum_list (map ?arg_wt (zip ss [0.. \ eval_tpoly A (map_tpoly (\x. x) zhmset_of (wt0 \)) + (zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))) + sum_list (map ?arg_wt (zip ss [0.. eval_tpoly A (map_tpoly (\p. p) zhmset_of (wt0 \))" using legal eval_ztpoly_nonneg by blast then show ?thesis by (meson leD leI le_add_same_cancel2 less_le_trans zhmset_of_nonneg) qed finally have ?case by simp } ultimately show ?case by linarith qed lemma wt_args_ge_length_times_\\<^sub>h: assumes legal: "legal_zpassign A" shows "wt_args i A \ ss \ of_nat (length ss) * zhmset_of \\<^sub>h" unfolding wt_args_def by (rule sum_list_ge_length_times[unfolded wt_args_def, of "map (eval_ztpoly A \ (\(s, i). PMult [coef_hd \ i, wt s])) (zip ss [i..\<^sub>h) lemma wt_ge_\\<^sub>h: "legal_zpassign A \ eval_ztpoly A (wt s) \ zhmset_of \\<^sub>h" using \\<^sub>h_le_\\<^sub>h[folded zhmset_of_le] order.trans wt_ge_\\<^sub>h zhmset_of_le by blast lemma wt_gt_0: "legal_zpassign A \ eval_ztpoly A (wt s) > 0" using \\<^sub>h_gt_0[folded zhmset_of_less, unfolded zhmset_of_0] wt_ge_\\<^sub>h by (blast intro: less_le_trans) lemma wt_gt_\\<^sub>h_if_superunary: assumes legal: "legal_zpassign A" and superunary: "arity_hd\<^sub>h (head s) > 1" shows "eval_ztpoly A (wt s) > zhmset_of \\<^sub>h" proof (cases "\\<^sub>h = \\<^sub>h") case \_ne_\: False show ?thesis using order.not_eq_order_implies_strict[OF \_ne_\ \\<^sub>h_le_\\<^sub>h, folded zhmset_of_less] wt_ge_\\<^sub>h[OF legal] by (blast intro: less_le_trans) next case \_eq_\: True show ?thesis using superunary proof (induct s rule: tm_induct_apps) case (apps \ ss) have "arity_hd\<^sub>h \ > 1" using apps(2) by simp hence min_gr_ary: "arity_sym\<^sub>h (min_ground_head \) > 1" using ground_heads_arity\<^sub>h less_le_trans min_ground_head_in_ground_heads by blast have "zhmset_of \\<^sub>h < eval_ztpoly A (wt0 \) + zhmset_of (\\<^sub>h * arity_sym\<^sub>h (min_ground_head \))" unfolding \_eq_\ by (rule add_strict_increasing2[OF eval_ztpoly_nonneg[OF legal]], unfold zhmset_of_less, rule gt_0_lt_mult_gt_1_hmset[OF \\<^sub>h_gt_0 min_gr_ary]) also have "\ \ eval_ztpoly A (wt0 \) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))) + zhmset_of (of_nat (length ss) * \\<^sub>h)" by (auto simp: \\<^sub>h_gt_0 \_eq_\ zmset_of_le zhmset_of_plus[symmetric] algebra_simps simp del: ring_distribs simp: ring_distribs[symmetric]) (metis add.commute le_minus_plus_same_hmset) also have "\ \ eval_ztpoly A (wt0 \) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))) + wt_args 0 A \ ss" using wt_args_ge_length_times_\\<^sub>h[OF legal] by (simp add: zhmset_of_times of_nat_zhmset) finally show ?case by (simp add: wt_args_def add_ac(1) comp_def) qed qed lemma wt_App_plus_\\<^sub>h_ge: "eval_ztpoly A (wt (App s t)) + zhmset_of \\<^sub>h \ eval_ztpoly A (wt s) + eval_ztpoly A (coef s 0) * eval_ztpoly A (wt t)" proof (cases s rule: tm_exhaust_apps) case s: (apps \ ss) show ?thesis proof (cases "arity_sym\<^sub>h (min_ground_head \) = \") case ary_eq_\: True show ?thesis unfolding ary_eq_\ s App_apps wt.simps by (auto simp: diff_diff_add_hmset[symmetric] add.assoc) next case False show ?thesis unfolding s App_apps wt.simps by (simp add: algebra_simps zhmset_of_plus[symmetric] zmset_of_le, simp del: diff_diff_add_hmset add: add.commute[of 1] le_minus_plus_same_hmset distrib_left[of _ "1 :: hmultiset", unfolded mult.right_neutral, symmetric] diff_diff_add_hmset[symmetric]) qed qed lemma wt_App_fun_\\<^sub>h: assumes legal: "legal_zpassign A" and wt_st: "eval_ztpoly A (wt (App s t)) = eval_ztpoly A (wt t)" shows "eval_ztpoly A (wt s) = zhmset_of \\<^sub>h" proof - have "eval_ztpoly A (wt (App s t)) = eval_ztpoly A (wt t)" using wt_st by simp hence wt_s_t_le_\_t: "eval_ztpoly A (wt s) + eval_ztpoly A (coef s 0) * eval_ztpoly A (wt t) \ zhmset_of \\<^sub>h + eval_ztpoly A (wt t)" using wt_App_plus_\\<^sub>h_ge by (metis add.commute) also have "\ \ eval_ztpoly A (wt s) + eval_ztpoly A (wt t)" using wt_ge_\\<^sub>h[OF legal] by simp finally have "eval_ztpoly A (coef s 0) * eval_ztpoly A (wt t) \ eval_ztpoly A (wt t)" by simp hence "eval_ztpoly A (coef s 0) = 1" using eval_ztpoly_nonneg[OF legal] by (metis (no_types, lifting) coef_gt_0 dual_order.order_iff_strict leD legal mult_cancel_right1 nonneg_le_mult_right_mono_zhmset wt_gt_0) thus ?thesis using wt_s_t_le_\_t by (simp add: add.commute antisym wt_ge_\\<^sub>h[OF legal]) qed lemma wt_App_arg_\\<^sub>h: assumes legal: "legal_zpassign A" and wt_st: "eval_ztpoly A (wt (App s t)) = eval_ztpoly A (wt s)" shows "eval_ztpoly A (wt t) = zhmset_of \\<^sub>h" proof - have "eval_ztpoly A (wt (App s t)) + zhmset_of \\<^sub>h = eval_ztpoly A (wt s) + zhmset_of \\<^sub>h" using wt_st by simp hence "eval_ztpoly A (coef s 0) * eval_ztpoly A (wt t) \ zhmset_of \\<^sub>h" (is "?k * ?w \ _") by (metis add_le_cancel_left wt_App_plus_\\<^sub>h_ge) hence "?k * ?w = zhmset_of \\<^sub>h" using wt_ge_\\<^sub>h[OF legal] coef_gt_0[OF legal, unfolded zero_less_iff_1_le_hmset] by (simp add: antisym nonneg_le_mult_right_mono_zhmset) hence "?w \ zhmset_of \\<^sub>h" by (metis coef_gt_0[OF legal] dual_order.order_iff_strict eval_ztpoly_nonneg[OF legal] nonneg_le_mult_right_mono_zhmset) thus ?thesis by (simp add: antisym wt_ge_\\<^sub>h[OF legal]) qed lemma wt_App_ge_fun: "wt (App s t) \\<^sub>p wt s" unfolding ge_tpoly_def proof clarify fix A assume legal: "legal_zpassign A" have "zhmset_of \\<^sub>h \ eval_ztpoly A (coef s 0) * eval_ztpoly A (wt t)" by (simp add: coef_gt_0 legal nonneg_le_mult_right_mono_zhmset wt_ge_\\<^sub>h) hence "eval_ztpoly A (wt s) + zhmset_of \\<^sub>h \ eval_ztpoly A (wt (App s t)) + zhmset_of \\<^sub>h" by (metis add_le_cancel_right add_less_le_mono not_le wt_App_plus_\\<^sub>h_ge) thus "eval_ztpoly A (wt s) \ eval_ztpoly A (wt (App s t))" by simp qed lemma wt_App_ge_arg: "wt (App s t) \\<^sub>p wt t" unfolding ge_tpoly_def by (cases s rule: tm_exhaust_apps, simp, unfold App_apps wt.simps) (auto simp: comp_def coef_hd_gt_0 eval_ztpoly_nonneg nonneg_le_mult_right_mono_zhmset intro!: sum_list_nonneg eval_ztpoly_nonneg add_increasing) lemma wt_\\<^sub>h_imp_\\<^sub>h_eq_\\<^sub>h: assumes legal: "legal_zpassign A" and wt_s_eq_\: "eval_ztpoly A (wt s) = zhmset_of \\<^sub>h" shows "\\<^sub>h = \\<^sub>h" using \\<^sub>h_le_\\<^sub>h wt_ge_\\<^sub>h[OF legal, of s, unfolded wt_s_eq_\ zhmset_of_le] by (rule antisym) lemma wt_ge_vars: "wt t \\<^sub>p wt s \ vars t \ vars s" proof (induct s) case t: (Hd \) note wt_ge_\ = this(1) show ?case proof (cases \) case \: (Var x) { assume z_ni_t: "x \ vars t" let ?A = min_zpassign let ?B = "\v. if v = PWt x then eval_ztpoly ?A (wt t) + ?A v + 1 else ?A v" have legal_B: "legal_zpassign ?B" unfolding legal_zpassign_def by (auto simp: legal_min_zpassign intro!: add_increasing eval_ztpoly_nonneg) have eval_B_eq_A: "eval_ztpoly ?B (wt t) = eval_ztpoly ?A (wt t)" by (rule wt_cong) (auto simp: z_ni_t) have "eval_ztpoly ?B (wt (Hd (Var x))) > eval_ztpoly ?B (wt t)" by (auto simp: eval_B_eq_A zero_less_iff_1_le_zhmset zhmset_of_plus[symmetric] algebra_simps) hence False using wt_ge_\ \ unfolding ge_tpoly_def by (blast dest: leD intro: legal_B legal_min_zpassign) } thus ?thesis by (auto simp: \) qed simp next case (App s1 s2) note ih1 = this(1) and ih2 = this(2) and wt_t_ge_wt_s1s2 = this(3) have "vars s1 \ vars t" using ih1 wt_t_ge_wt_s1s2 wt_App_ge_fun order_trans unfolding ge_tpoly_def by blast moreover have "vars s2 \ vars t" using ih2 wt_t_ge_wt_s1s2 wt_App_ge_arg order_trans unfolding ge_tpoly_def by blast ultimately show ?case by simp qed lemma sum_coefs_ge_num_args_if_\\<^sub>h_eq_0: assumes legal: "legal_passign A" and \_eq_0: "\\<^sub>h = 0" and wary_s: "wary s" shows "sum_coefs (eval_tpoly A (wt s)) \ num_args s" proof (cases s rule: tm_exhaust_apps) case s: (apps \ ss) show ?thesis unfolding s proof (induct ss rule: rev_induct) case (snoc sa ss) note ih = this let ?Az = "\v. zhmset_of (A v)" have legalz: "legal_zpassign ?Az" using legal unfolding legal_passign_def legal_zpassign_def zhmset_of_le by assumption have "eval_ztpoly ?Az (coef_hd \ (length ss)) > 0" using legal coef_hd_gt_0 eval_tpoly_eq_eval_ztpoly by (simp add: coef_hd_gt_0[OF legalz]) hence k: "eval_tpoly A (coef_hd \ (length ss)) > 0" (is "?k > _") unfolding eval_tpoly_eq_eval_ztpoly[symmetric] zhmset_of_less[symmetric] zhmset_of_0 by assumption have "eval_ztpoly ?Az (wt sa) > 0" (is "?w > _") by (simp add: wt_gt_0[OF legalz]) hence w: "eval_tpoly A (wt sa) > 0" (is "?w > _") unfolding eval_tpoly_eq_eval_ztpoly[symmetric] zhmset_of_less[symmetric] zhmset_of_0 by assumption have "?k * ?w > 0" using k w by simp hence "sum_coefs (?k * ?w) > 0" by (rule sum_coefs_gt_0[THEN iffD2]) thus ?case using ih by (simp del: apps_append add: s \_eq_0) qed simp qed subsection \Inductive Definitions\ inductive gt :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix ">\<^sub>t" 50) where gt_wt: "wt t >\<^sub>p wt s \ t >\<^sub>t s" | gt_unary: "wt t \\<^sub>p wt s \ \ head t \\\<^sub>h\<^sub>d head s \ num_args t = 1 \ (\f \ ground_heads (head t). arity_sym f = 1 \ wt_sym f = 0) \ arg t >\<^sub>t s \ arg t = s \ t >\<^sub>t s" | gt_diff: "wt t \\<^sub>p wt s \ head t >\<^sub>h\<^sub>d head s \ t >\<^sub>t s" | gt_same: "wt t \\<^sub>p wt s \ head t = head s \ (\f \ ground_heads (head t). extf f (>\<^sub>t) (args t) (args s)) \ t >\<^sub>t s" abbreviation ge :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix "\\<^sub>t" 50) where "t \\<^sub>t s \ t >\<^sub>t s \ t = s" inductive gt_wt :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_wtI: "wt t >\<^sub>p wt s \ gt_wt t s" inductive gt_unary :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_unaryI: "wt t \\<^sub>p wt s \ \ head t \\\<^sub>h\<^sub>d head s \ num_args t = 1 \ (\f \ ground_heads (head t). arity_sym f = 1 \ wt_sym f = 0) \ arg t \\<^sub>t s \ gt_unary t s" inductive gt_diff :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_diffI: "wt t \\<^sub>p wt s \ head t >\<^sub>h\<^sub>d head s \ gt_diff t s" inductive gt_same :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_sameI: "wt t \\<^sub>p wt s \ head t = head s \ (\f \ ground_heads (head t). extf f (>\<^sub>t) (args t) (args s)) \ gt_same t s" lemma gt_iff_wt_unary_diff_same: "t >\<^sub>t s \ gt_wt t s \ gt_unary t s \ gt_diff t s \ gt_same t s" by (subst gt.simps) (auto simp: gt_wt.simps gt_unary.simps gt_diff.simps gt_same.simps) lemma gt_imp_wt: "t >\<^sub>t s \ wt t \\<^sub>p wt s" by (blast elim: gt.cases) lemma gt_imp_vars: "t >\<^sub>t s \ vars t \ vars s" by (erule wt_ge_vars[OF gt_imp_wt]) subsection \Irreflexivity\ theorem gt_irrefl: "wary s \ \ s >\<^sub>t s" proof (induct "size s" arbitrary: s rule: less_induct) case less note ih = this(1) and wary_s = this(2) show ?case proof assume s_gt_s: "s >\<^sub>t s" show False using s_gt_s proof (cases rule: gt.cases) case gt_same then obtain f where f: "extf f (>\<^sub>t) (args s) (args s)" by fastforce thus False using wary_s ih by (metis wary_args extf_irrefl size_in_args) qed (auto simp: comp_hd_def gt_tpoly_irrefl gt_hd_irrefl) qed qed subsection \Transitivity\ lemma not_extf_gt_nil_singleton_if_\\<^sub>h_eq_\\<^sub>h: assumes wary_s: "wary s" and \_eq_\: "\\<^sub>h = \\<^sub>h" shows "\ extf f (>\<^sub>t) [] [s]" proof assume nil_gt_s: "extf f (>\<^sub>t) [] [s]" note s_gt_nil = extf_singleton_nil_if_\\<^sub>h_eq_\\<^sub>h[OF \_eq_\, of f gt s] have "\ extf f (>\<^sub>t) [] []" by (rule extf_irrefl) simp moreover have "extf f (>\<^sub>t) [] []" using extf_trans_from_irrefl[of "{s}", OF _ _ _ _ _ _ nil_gt_s s_gt_nil] gt_irrefl[OF wary_s] by fastforce ultimately show False by sat qed lemma gt_sub_arg: "wary (App s t) \ App s t >\<^sub>t t" proof (induct t arbitrary: s rule: measure_induct_rule[of size]) case (less t) note ih = this(1) and wary_st = this(2) { fix A assume legal: "legal_zpassign A" and wt_st: "eval_ztpoly A (wt (App s t)) = eval_ztpoly A (wt t)" have \_eq_\: "\\<^sub>h = \\<^sub>h" using wt_App_fun_\\<^sub>h[OF legal] wt_\\<^sub>h_imp_\\<^sub>h_eq_\\<^sub>h[OF legal] wt_st by blast hence \_gt_0: "\\<^sub>h > 0" using \\<^sub>h_gt_0 by simp have wt_s: "eval_ztpoly A (wt s) = zhmset_of \\<^sub>h" by (rule wt_App_fun_\\<^sub>h[OF legal wt_st]) have wary_t: "wary t" by (rule wary_AppE\<^sub>h[OF wary_st]) have nargs_lt: "of_nat (num_args s) < arity_hd\<^sub>h (head s)" by (rule wary_AppE\<^sub>h[OF wary_st]) have ary_hd_s: "arity_hd\<^sub>h (head s) = 1" by (metis gr_implies_not_zero_hmset legal lt_1_iff_eq_0_hmset nargs_lt neq_iff wt_gt_\\<^sub>h_if_superunary wt_s) hence nargs_s: "num_args s = 0" by (metis less_one nargs_lt of_nat_1 of_nat_less_hmset) hence s_eq_hd: "s = Hd (head s)" by (simp add: Hd_head_id) obtain f where f_in: "f \ ground_heads (head s)" and wt_f_etc: "wt_sym f + \\<^sub>h * arity_sym\<^sub>h f = \\<^sub>h" proof - assume a: "\f. \f \ local.ground_heads (head s); wt_sym f + \\<^sub>h * arity_sym\<^sub>h f = \\<^sub>h\ \ thesis" have "\f. \\<^sub>h - \\<^sub>h * arity_sym\<^sub>h f \ wt_sym f" using wt_s by (metis legal wt_\\<^sub>h_imp_\\<^sub>h_eq_\\<^sub>h wt_sym_ge\<^sub>h) hence "\s. \ \\<^sub>h * arity_sym\<^sub>h s + wt_sym s < \\<^sub>h" by (metis add_diff_cancel_left' le_imp_minus_plus_hmset leD le_minus_plus_same_hmset less_le_trans) thus thesis using a wt_s s_eq_hd by (metis exists_wt_sym legal add.commute order.not_eq_order_implies_strict zhmset_of_le) qed have ary_f_1: "arity_sym f = 1" by (metis \_gt_0 add_diff_cancel_left' ary_hd_s diff_le_self_hmset dual_order.order_iff_strict f_in ground_heads_arity\<^sub>h gt_0_lt_mult_gt_1_hmset hmset_of_enat_1 hmset_of_enat_inject leD wt_f_etc) hence wt_f_0: "wt_sym f = 0" using wt_f_etc by simp { assume hd_s_ncmp_t: "\ head s \\\<^sub>h\<^sub>d head t" have ?case by (rule gt_unary[OF wt_App_ge_arg]) (auto simp: hd_s_ncmp_t nargs_s intro: f_in ary_f_1 wt_f_0) } moreover { assume hd_s_gt_t: "head s >\<^sub>h\<^sub>d head t" have ?case by (rule gt_diff[OF wt_App_ge_arg]) (simp add: hd_s_gt_t) } moreover { assume "head t >\<^sub>h\<^sub>d head s" hence False using ary_f_1 wt_f_0 f_in gt_hd_irrefl gt_sym_antisym unary_wt_sym_0_gt\<^sub>h hmset_of_enat_1 unfolding gt_hd_def by metis } moreover { assume hd_t_eq_s: "head t = head s" hence nargs_t_le: "num_args t \ 1" using ary_hd_s wary_num_args_le_arity_head\<^sub>h[OF wary_t] of_nat_le_hmset by fastforce have extf: "extf f (>\<^sub>t) [t] (args t)" for f proof (cases "args t") case Nil thus ?thesis by (simp add: extf_singleton_nil_if_\\<^sub>h_eq_\\<^sub>h[OF \_eq_\]) next case args_t: (Cons ta ts) hence ts: "ts = []" using ary_hd_s[folded hd_t_eq_s] wary_num_args_le_arity_head\<^sub>h[OF wary_t] of_nat_le_hmset nargs_t_le by simp have ta: "ta = arg t" by (metis apps.simps(1) apps.simps(2) args_t tm.sel(6) tm_collapse_apps ts) hence t: "t = App (fun t) ta" by (metis args.simps(1) args_t not_Cons_self2 tm.exhaust_sel ts) have "t >\<^sub>t ta" by (rule ih[of ta "fun t", folded t, OF _ wary_t]) (metis ta size_arg_lt t tm.disc(2)) thus ?thesis unfolding args_t ts by (metis extf_singleton gt_irrefl wary_t) qed have ?case by (rule gt_same[OF wt_App_ge_arg]) (simp_all add: hd_t_eq_s length_0_conv[THEN iffD1, OF nargs_s] extf) } ultimately have ?case unfolding comp_hd_def by metis } thus ?case using gt_wt by (metis ge_tpoly_def gt_tpoly_def wt_App_ge_arg order.not_eq_order_implies_strict) qed lemma gt_arg: "wary s \ is_App s \ s >\<^sub>t arg s" by (cases s) (auto intro: gt_sub_arg) theorem gt_trans: "wary u \ wary t \ wary s \ u >\<^sub>t t \ t >\<^sub>t s \ u >\<^sub>t s" proof (simp only: atomize_imp, rule measure_induct_rule[of "\(u, t, s). {#size u, size t, size s#}" "\(u, t, s). wary u \ wary t \ wary s \ u >\<^sub>t t \ t >\<^sub>t s \ u >\<^sub>t s" "(u, t, s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric]) fix u t s assume ih: "\ua ta sa. {#size ua, size ta, size sa#} < {#size u, size t, size s#} \ wary ua \ wary ta \ wary sa \ ua >\<^sub>t ta \ ta >\<^sub>t sa \ ua >\<^sub>t sa" and wary_u: "wary u" and wary_t: "wary t" and wary_s: "wary s" and u_gt_t: "u >\<^sub>t t" and t_gt_s: "t >\<^sub>t s" have wt_u_ge_t: "wt u \\<^sub>p wt t" and wt_t_ge_s: "wt t \\<^sub>p wt s" using gt_imp_wt u_gt_t t_gt_s by auto hence wt_u_ge_s: "wt u \\<^sub>p wt s" by (rule ge_ge_tpoly_trans) have wary_arg_u: "wary (arg u)" by (rule wary_arg[OF wary_u]) have wary_arg_t: "wary (arg t)" by (rule wary_arg[OF wary_t]) have wary_arg_s: "wary (arg s)" by (rule wary_arg[OF wary_s]) show "u >\<^sub>t s" using t_gt_s proof cases case gt_wt_t_s: gt_wt hence "wt u >\<^sub>p wt s" using wt_u_ge_t ge_gt_tpoly_trans by blast thus ?thesis by (rule gt_wt) next case gt_unary_t_s: gt_unary have t_app: "is_App t" by (metis args_Nil_iff_is_Hd gt_unary_t_s(3) length_greater_0_conv less_numeral_extra(1)) hence nargs_fun_t: "num_args (fun t) < arity_hd (head (fun t))" by (metis tm.collapse(2) wary_AppE wary_t) have \_eq_\: "\\<^sub>h = \\<^sub>h" using gt_unary_t_s(4) unary_wt_sym_0_imp_\\<^sub>h_eq_\\<^sub>h by blast show ?thesis using u_gt_t proof cases case gt_wt_u_t: gt_wt hence "wt u >\<^sub>p wt s" using wt_t_ge_s gt_ge_tpoly_trans by blast thus ?thesis by (rule gt_wt) next case gt_unary_u_t: gt_unary have u_app: "is_App u" by (metis args_Nil_iff_is_Hd gt_unary_u_t(3) length_greater_0_conv less_numeral_extra(1)) hence nargs_fun_u: "num_args (fun u) = 0" by (metis args.simps(1) gt_unary_u_t(3) list.size(3) one_arg_imp_Hd tm.collapse(2)) have arg_u_gt_s: "arg u >\<^sub>t s" using ih[of "arg u" t s] u_app gt_unary_u_t(5) t_gt_s size_arg_lt wary_arg_u wary_s wary_t by force hence arg_u_ge_s: "arg u \\<^sub>t s" by sat { assume "size (arg u) < size t" hence "{#size u, size (arg u), size s#} < {#size u, size t, size s#}" by simp hence ?thesis using ih[of u "arg u" s] arg_u_gt_s gt_arg u_app wary_s wary_u by blast } moreover { assume "size (arg t) < size s" hence "u >\<^sub>t arg t" using ih[of u t "arg t"] args_Nil_iff_is_Hd gt_arg gt_unary_t_s(3) u_gt_t wary_t wary_u by force hence ?thesis using ih[of u "arg t" s] args_Nil_iff_is_Hd gt_unary_t_s(3,5) size_arg_lt wary_arg_t wary_s wary_u by force } moreover { assume sz_u_gt_t: "size u > size t" and sz_t_gt_s: "size t > size s" { assume hd_u_eq_s: "head u = head s" hence ary_hd_s: "arity_hd (head s) = 1" using ground_heads_arity gt_unary_u_t(3,4) hd_u_eq_s one_enat_def wary_num_args_le_arity_head wary_u by fastforce have extf: "extf f (>\<^sub>t) (args u) (args s)" for f proof (cases "args s") case Nil thus ?thesis by (metis \_eq_\ args.elims args_Nil_iff_is_Hd extf_snoc_if_\\<^sub>h_eq_\\<^sub>h length_0_conv nargs_fun_u tm.sel(4) u_app) next case args_s: (Cons sa ss) hence ss: "ss = []" by (cases s, simp, metis One_nat_def antisym_conv ary_hd_s diff_Suc_1 enat_ord_simps(1) le_add2 length_0_conv length_Cons list.size(4) one_enat_def wary_num_args_le_arity_head wary_s) have sa: "sa = arg s" by (metis apps.simps(1) apps.simps(2) args_s tm.sel(6) tm_collapse_apps ss) have s_app: "is_App s" using args_Nil_iff_is_Hd args_s by force have args_u: "args u = [arg u]" by (metis append_Nil args.simps(2) args_Nil_iff_is_Hd gt_unary_u_t(3) length_0_conv nargs_fun_u tm.collapse(2) zero_neq_one) have max_sz_arg_u_t_arg_t: "Max {size (arg t), size t, size (arg u)} < size u" using size_arg_lt sz_u_gt_t t_app u_app by fastforce have "{#size (arg u), size t, size (arg t)#} < {#size u, size t, size s#}" using max_sz_arg_u_t_arg_t by (auto intro!: Max_lt_imp_lt_mset) hence arg_u_gt_arg_t: "arg u >\<^sub>t arg t" using ih[OF _ wary_arg_u wary_t wary_arg_t] args_Nil_iff_is_Hd gt_arg gt_unary_t_s(3) gt_unary_u_t(5) wary_t by force have max_sz_arg_s_s_arg_t: "Max {size (arg s), size s, size (arg t)} < size u" using s_app t_app size_arg_lt sz_t_gt_s sz_u_gt_t by force have "{#size (arg t), size s, size (arg s)#} < {#size u, size t, size s#}" using max_sz_arg_s_s_arg_t by (auto intro!: Max_lt_imp_lt_mset) hence arg_t_gt_arg_s: "arg t >\<^sub>t arg s" using ih[OF _ wary_arg_t wary_s wary_arg_s] gt_unary_t_s(5) gt_arg args_Nil_iff_is_Hd args_s wary_s by force have "{#size (arg u), size (arg t), size (arg s)#} < {#size u, size t, size s#}" by (auto intro!: add_mset_lt_lt_lt simp: size_arg_lt u_app t_app s_app) hence "arg u >\<^sub>t arg s" using ih[of "arg u" "arg t" "arg s"] arg_u_gt_arg_t arg_t_gt_arg_s wary_arg_s wary_arg_t wary_arg_u by blast thus ?thesis unfolding args_u args_s ss sa by (metis extf_singleton gt_irrefl wary_arg_u) qed have ?thesis by (rule gt_same[OF wt_u_ge_s hd_u_eq_s]) (simp add: extf) } moreover { assume "head u >\<^sub>h\<^sub>d head s" hence ?thesis by (rule gt_diff[OF wt_u_ge_s]) } moreover { assume "head s >\<^sub>h\<^sub>d head u" hence False using gt_hd_def gt_hd_irrefl gt_sym_antisym gt_unary_u_t(4) unary_wt_sym_0_gt by blast } moreover { assume "\ head u \\\<^sub>h\<^sub>d head s" hence ?thesis by (rule gt_unary[OF wt_u_ge_s _ gt_unary_u_t(3,4) arg_u_ge_s]) } ultimately have ?thesis unfolding comp_hd_def by sat } ultimately show ?thesis by (meson less_le_trans linorder_not_le size_arg_lt t_app u_app) next case gt_diff_u_t: gt_diff have False using gt_diff_u_t(2) gt_hd_def gt_hd_irrefl gt_sym_antisym gt_unary_t_s(4) unary_wt_sym_0_gt by blast thus ?thesis by sat next case gt_same_u_t: gt_same have hd_u_ncomp_s: "\ head u \\\<^sub>h\<^sub>d head s" by (rule gt_unary_t_s(2)[folded gt_same_u_t(2)]) have "\f \ ground_heads (head u). arity_sym f = 1 \ wt_sym f = 0" by (rule gt_unary_t_s(4)[folded gt_same_u_t(2)]) hence "arity_hd (head u) = 1" by (metis dual_order.order_iff_strict gr_implies_not_zero_hmset ground_heads_arity gt_same_u_t(2) head_fun hmset_of_enat_1 hmset_of_enat_less lt_1_iff_eq_0_hmset nargs_fun_t) hence "num_args u \ 1" using of_nat_le_hmset wary_num_args_le_arity_head\<^sub>h wary_u by fastforce hence nargs_u: "num_args u = 1" by (cases "args u", metis Hd_head_id \_eq_\ append_Nil args.simps(2) ex_in_conv[THEN iffD2, OF ground_heads_nonempty] gt_same_u_t(2,3) gt_unary_t_s(3) head_fun list.size(3) not_extf_gt_nil_singleton_if_\\<^sub>h_eq_\\<^sub>h one_arg_imp_Hd tm.collapse(2)[OF t_app] wary_arg_t, simp) hence u_app: "is_App u" by (cases u) auto have "arg u >\<^sub>t arg t" by (metis extf_singleton[THEN iffD1] append_Nil args.simps args_Nil_iff_is_Hd comp_hd_def gt_hd_def gt_irrefl gt_same_u_t(2,3) gt_unary_t_s(2,3) head_fun length_0_conv nargs_u one_arg_imp_Hd t_app tm.collapse(2) u_gt_t wary_u) moreover have "{#size (arg u), size (arg t), size s#} < {#size u, size t, size s#}" by (auto intro!: add_mset_lt_lt_lt simp: size_arg_lt u_app t_app) ultimately have "arg u >\<^sub>t s" using ih[OF _ wary_arg_u wary_arg_t wary_s] gt_unary_t_s(5) by blast hence arg_u_ge_s: "arg u \\<^sub>t s" by sat show ?thesis by (rule gt_unary[OF wt_u_ge_s hd_u_ncomp_s nargs_u _ arg_u_ge_s]) (simp add: gt_same_u_t(2) gt_unary_t_s(4)) qed next case gt_diff_t_s: gt_diff show ?thesis using u_gt_t proof cases case gt_wt_u_t: gt_wt hence "wt u >\<^sub>p wt s" using wt_t_ge_s gt_ge_tpoly_trans by blast thus ?thesis by (rule gt_wt) next case gt_unary_u_t: gt_unary have u_app: "is_App u" by (metis args_Nil_iff_is_Hd gt_unary_u_t(3) length_greater_0_conv less_numeral_extra(1)) hence "arg u >\<^sub>t s" using ih[of "arg u" t s] gt_unary_u_t(5) t_gt_s size_arg_lt wary_arg_u wary_s wary_t by force hence arg_u_ge_s: "arg u \\<^sub>t s" by sat { assume "head u = head s" hence False using gt_diff_t_s(2) gt_unary_u_t(2) unfolding comp_hd_def by force } moreover { assume "head s >\<^sub>h\<^sub>d head u" hence False using gt_hd_def gt_hd_irrefl gt_sym_antisym gt_unary_u_t(4) unary_wt_sym_0_gt by blast } moreover { assume "head u >\<^sub>h\<^sub>d head s" hence ?thesis by (rule gt_diff[OF wt_u_ge_s]) } moreover { assume "\ head u \\\<^sub>h\<^sub>d head s" hence ?thesis by (rule gt_unary[OF wt_u_ge_s _ gt_unary_u_t(3,4) arg_u_ge_s]) } ultimately show ?thesis unfolding comp_hd_def by sat next case gt_diff_u_t: gt_diff have "head u >\<^sub>h\<^sub>d head s" using gt_diff_u_t(2) gt_diff_t_s(2) gt_hd_trans by blast thus ?thesis by (rule gt_diff[OF wt_u_ge_s]) next case gt_same_u_t: gt_same have "head u >\<^sub>h\<^sub>d head s" using gt_diff_t_s(2) gt_same_u_t(2) by simp thus ?thesis by (rule gt_diff[OF wt_u_ge_s]) qed next case gt_same_t_s: gt_same show ?thesis using u_gt_t proof cases case gt_wt_u_t: gt_wt hence "wt u >\<^sub>p wt s" using wt_t_ge_s gt_ge_tpoly_trans by blast thus ?thesis by (rule gt_wt) next case gt_unary_u_t: gt_unary have "is_App u" by (metis args_Nil_iff_is_Hd gt_unary_u_t(3) length_greater_0_conv less_numeral_extra(1)) hence "arg u >\<^sub>t s" using ih[of "arg u" t s] gt_unary_u_t(5) t_gt_s size_arg_lt wary_arg_u wary_s wary_t by force hence arg_u_ge_s: "arg u \\<^sub>t s" by sat have "\ head u \\\<^sub>h\<^sub>d head s" using gt_same_t_s(2) gt_unary_u_t(2) by simp thus ?thesis by (rule gt_unary[OF wt_u_ge_s _ gt_unary_u_t(3,4) arg_u_ge_s]) next case gt_diff_u_t: gt_diff have "head u >\<^sub>h\<^sub>d head s" using gt_diff_u_t(2) gt_same_t_s(2) by simp thus ?thesis by (rule gt_diff[OF wt_u_ge_s]) next case gt_same_u_t: gt_same have hd_u_s: "head u = head s" by (simp only: gt_same_t_s(2) gt_same_u_t(2)) let ?S = "set (args u) \ set (args t) \ set (args s)" have gt_trans_args: "\ua \ ?S. \ta \ ?S. \sa \ ?S. ua >\<^sub>t ta \ ta >\<^sub>t sa \ ua >\<^sub>t sa" proof clarify fix sa ta ua assume ua_in: "ua \ ?S" and ta_in: "ta \ ?S" and sa_in: "sa \ ?S" and ua_gt_ta: "ua >\<^sub>t ta" and ta_gt_sa: "ta >\<^sub>t sa" have wary_sa: "wary sa" and wary_ta: "wary ta" and wary_ua: "wary ua" using wary_args ua_in ta_in sa_in wary_u wary_t wary_s by blast+ show "ua >\<^sub>t sa" by (auto intro!: ih[OF Max_lt_imp_lt_mset wary_ua wary_ta wary_sa ua_gt_ta ta_gt_sa]) (meson ua_in ta_in sa_in Un_iff max.strict_coboundedI1 max.strict_coboundedI2 size_in_args)+ qed have "\f \ ground_heads (head u). extf f (>\<^sub>t) (args u) (args s)" by (clarify, rule extf_trans_from_irrefl[of ?S _ "args t", OF _ _ _ _ _ gt_trans_args]) (auto simp: gt_same_u_t(2,3) gt_same_t_s(3) wary_args wary_u wary_t wary_s gt_irrefl) thus ?thesis by (rule gt_same[OF wt_u_ge_s hd_u_s]) qed qed qed lemma gt_antisym: "wary s \ wary t \ t >\<^sub>t s \ \ s >\<^sub>t t" using gt_irrefl gt_trans by blast subsection \Subterm Property\ lemma gt_sub_fun: "App s t >\<^sub>t s" proof (cases "wt (App s t) >\<^sub>p wt s") case True thus ?thesis using gt_wt by simp next case False hence \_eq_\: "\\<^sub>h = \\<^sub>h" using wt_App_ge_fun dual_order.order_iff_strict wt_App_arg_\\<^sub>h wt_\\<^sub>h_imp_\\<^sub>h_eq_\\<^sub>h unfolding gt_tpoly_def ge_tpoly_def by fast have hd_st: "head (App s t) = head s" by auto have extf: "\f \ ground_heads (head (App s t)). extf f (>\<^sub>t) (args (App s t)) (args s)" by (simp add: \_eq_\ extf_snoc_if_\\<^sub>h_eq_\\<^sub>h) show ?thesis by (rule gt_same[OF wt_App_ge_fun hd_st extf]) qed theorem gt_proper_sub: "wary t \ proper_sub s t \ t >\<^sub>t s" by (induct t) (auto intro: gt_sub_fun gt_sub_arg gt_trans sub.intros wary_sub) subsection \Compatibility with Functions\ lemma gt_compat_fun: assumes wary_t: "wary t" and t'_gt_t: "t' >\<^sub>t t" shows "App s t' >\<^sub>t App s t" proof (rule gt_same; clarify?) show "wt (App s t') \\<^sub>p wt (App s t)" using gt_imp_wt[OF t'_gt_t, unfolded ge_tpoly_def] by (cases s rule: tm_exhaust_apps, auto simp del: apps_append simp: ge_tpoly_def App_apps eval_ztpoly_nonneg intro: ordered_comm_semiring_class.comm_mult_left_mono) next fix f have "extf f (>\<^sub>t) (args s @ [t']) (args s @ [t])" using t'_gt_t by (metis extf_compat_list gt_irrefl[OF wary_t]) thus "extf f (>\<^sub>t) (args (App s t')) (args (App s t))" by simp qed simp theorem gt_compat_fun_strong: assumes wary_t: "wary t" and t'_gt_t: "t' >\<^sub>t t" shows "apps s (t' # us) >\<^sub>t apps s (t # us)" proof (induct us rule: rev_induct) case Nil show ?case using t'_gt_t by (auto intro!: gt_compat_fun[OF wary_t]) next case (snoc u us) note ih = snoc let ?v' = "apps s (t' # us @ [u])" let ?v = "apps s (t # us @ [u])" have "wt ?v' \\<^sub>p wt ?v" using gt_imp_wt[OF ih] by (cases s rule: tm_exhaust_apps, simp del: apps_append add: App_apps apps_append[symmetric] ge_tpoly_def, subst (1 2) zip_eq_butlast_last, simp+) moreover have "head ?v' = head ?v" by simp moreover have "\f \ ground_heads (head ?v'). extf f (>\<^sub>t) (args ?v') (args ?v)" by (metis args_apps extf_compat_list gt_irrefl[OF wary_t] t'_gt_t) ultimately show ?case by (rule gt_same) qed subsection \Compatibility with Arguments\ theorem gt_compat_arg_weak: assumes wary_st: "wary (App s t)" and wary_s't: "wary (App s' t)" and coef_s'_0_ge_s: "coef s' 0 \\<^sub>p coef s 0" and s'_gt_s: "s' >\<^sub>t s" shows "App s' t >\<^sub>t App s t" proof - obtain \ ss where s: "s = apps (Hd \) ss" by (metis tm_exhaust_apps) obtain \' ss' where s': "s' = apps (Hd \') ss'" by (metis tm_exhaust_apps) have len_ss_lt: "of_nat (length ss) < arity_sym\<^sub>h (min_ground_head \)" using wary_st[unfolded s] ground_heads_arity\<^sub>h less_le_trans min_ground_head_in_ground_heads by (metis (no_types) tm_collapse_apps tm_inject_apps wary_AppE\<^sub>h) have \_etc: "\\<^sub>h + \\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss) - 1) = \\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))" if wary: "wary (App (apps (Hd \) ss) t)" for \ ss proof (cases "\\<^sub>h > 0") case True then obtain n where n: "of_nat n = arity_sym\<^sub>h (min_ground_head \)" by (metis arity_sym\<^sub>h_if_\\<^sub>h_gt_0_E) have "of_nat (length ss) < arity_sym\<^sub>h (min_ground_head \)" using wary by (metis (no_types) wary_AppE\<^sub>h ground_heads_arity\<^sub>h le_less_trans min_ground_head_in_ground_heads not_le tm_collapse_apps tm_inject_apps) thus ?thesis by (fold n, subst of_nat_1[symmetric], fold of_nat_minus_hmset, simp, metis Suc_diff_Suc mult_Suc_right of_nat_add of_nat_mult) qed simp have coef_\'_ge_\: "coef_hd \' (length ss') \\<^sub>p coef_hd \ (length ss)" by (rule coef_s'_0_ge_s[unfolded s s', simplified]) have wt_s'_ge_s: "wt s' \\<^sub>p wt s" by (rule gt_imp_wt[OF s'_gt_s]) have \_tms_len_ss_tms_wt_t_le: "eval_ztpoly A (coef_hd \ (length ss)) * eval_ztpoly A (wt t) \ eval_ztpoly A (coef_hd \' (length ss')) * eval_ztpoly A (wt t)" if legal: "legal_zpassign A" for A using legal coef_\'_ge_\[unfolded ge_tpoly_def] by (simp add: eval_ztpoly_nonneg mult_right_mono) have wt_s't_ge_st: "wt (App s' t) \\<^sub>p wt (App s t)" unfolding s s' by (clarsimp simp del: apps_append simp: App_apps ge_tpoly_def add_ac(1)[symmetric] intro!: add_mono[OF _ \_tms_len_ss_tms_wt_t_le], rule add_le_imp_le_left[of "zhmset_of \\<^sub>h"], unfold add_ac(1)[symmetric] add.commute[of 1] diff_diff_add[symmetric], subst (1 3) ac_simps(3)[unfolded add_ac(1)[symmetric]], subst (1 3) add_ac(1), simp only: zhmset_of_plus[symmetric] \_etc[OF wary_st[unfolded s]] \_etc[OF wary_s't[unfolded s']] add_ac(1) wt_s'_ge_s[unfolded s s', unfolded ge_tpoly_def add_ac(1)[symmetric], simplified]) show ?thesis using s'_gt_s proof cases case gt_wt_s'_s: gt_wt have "wt (App s' t) >\<^sub>p wt (App s t)" unfolding s s' by (clarsimp simp del: apps_append simp: App_apps gt_tpoly_def add_ac(1)[symmetric] intro!: add_less_le_mono[OF _ \_tms_len_ss_tms_wt_t_le], rule add_less_imp_less_left[of "zhmset_of \\<^sub>h"], unfold add_ac(1)[symmetric] add.commute[of 1] diff_diff_add[symmetric], subst (1 3) ac_simps(3)[unfolded add_ac(1)[symmetric]], subst (1 3) add_ac(1), simp only: zhmset_of_plus[symmetric] \_etc[OF wary_st[unfolded s]] \_etc[OF wary_s't[unfolded s']] add_ac(1) gt_wt_s'_s[unfolded s s', unfolded gt_tpoly_def add_ac(1)[symmetric], simplified]) thus ?thesis by (rule gt_wt) next case gt_unary_s'_s: gt_unary have False by (metis ground_heads_arity\<^sub>h gt_unary_s'_s(3) gt_unary_s'_s(4) hmset_of_enat_1 leD of_nat_1 wary_AppE\<^sub>h wary_s't) thus ?thesis by sat next case gt_diff_s'_s: gt_diff show ?thesis by (rule gt_diff[OF wt_s't_ge_st]) (simp add: gt_diff_s'_s(2)) next case gt_same_s'_s: gt_same have hd_s't: "head (App s' t) = head (App s t)" by (simp add: gt_same_s'_s(2)) have "\f \ ground_heads (head (App s' t)). extf f (>\<^sub>t) (args (App s' t)) (args (App s t))" using gt_same_s'_s(3) by (auto intro: extf_compat_append_right) thus ?thesis by (rule gt_same[OF wt_s't_ge_st hd_s't]) qed qed subsection \Stability under Substitution\ primrec subst_zpassign :: "('v \ ('s, 'v) tm) \ ('v pvar \ zhmultiset) \ 'v pvar \ zhmultiset" where "subst_zpassign \ A (PWt x) = eval_ztpoly A (wt (\ x)) - zhmset_of (\\<^sub>h * arity_sym\<^sub>h (min_ground_head (Var x)))" | "subst_zpassign \ A (PCoef x i) = eval_ztpoly A (coef (\ x) i)" lemma legal_subst_zpassign: assumes legal: "legal_zpassign A" and wary_\: "wary_subst \" shows "legal_zpassign (subst_zpassign \ A)" unfolding legal_zpassign_def proof fix v show "subst_zpassign \ A v \ min_zpassign v" proof (cases v) case v: (PWt x) obtain \ ss where \x: "\ x = apps (Hd \) ss" by (rule tm_exhaust_apps) have ghd_\: "ground_heads \ \ ground_heads_var x" using wary_\[unfolded wary_subst_def, rule_format, of x, unfolded \x] by simp have "zhmset_of (wt_sym (min_ground_head (Var x)) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head (Var x))) \ eval_ztpoly A (wt0 \) + zhmset_of (\\<^sub>h * arity_sym\<^sub>h (min_ground_head \))" proof - have mgh_x_min: "zhmset_of (wt_sym (min_ground_head (Var x)) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head (Var x))) \ zhmset_of (wt_sym (min_ground_head \) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head \))" by (simp add: zmset_of_le zhmset_of_le ghd_\ min_ground_head_antimono) have wt_mgh_le_wt0: "zhmset_of (wt_sym (min_ground_head \)) \ eval_ztpoly A (wt0 \)" using wt0_ge_min_ground_head[OF legal] by blast show ?thesis by (rule order_trans[OF mgh_x_min]) (simp add: zhmset_of_plus wt_mgh_le_wt0) qed also have "\ \ eval_ztpoly A (wt0 \) + zhmset_of ((\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))) + of_nat (length ss) * \\<^sub>h)" proof - have "zhmset_of (\\<^sub>h * arity_sym\<^sub>h (min_ground_head \)) \ zhmset_of (\\<^sub>h * (of_nat (length ss) + (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))))" by (metis add.commute le_minus_plus_same_hmset mult_le_mono2_hmset zhmset_of_le) thus ?thesis by (simp add: add.commute add.left_commute distrib_left mult.commute) qed also have "\ \ eval_ztpoly A (wt0 \) + zhmset_of ((\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))) + of_nat (length ss) * \\<^sub>h)" using \\<^sub>h_le_\\<^sub>h zhmset_of_le by auto also have "\ \ eval_ztpoly A (wt0 \) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))) + wt_args 0 A \ ss" using wt_args_ge_length_times_\\<^sub>h[OF legal] by (simp add: algebra_simps zhmset_of_plus zhmset_of_times of_nat_zhmset) finally have wt_x_le_\ssts: "zhmset_of (wt_sym (min_ground_head (Var x)) + \\<^sub>h * arity_sym\<^sub>h (min_ground_head (Var x))) \ eval_ztpoly A (wt0 \) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - of_nat (length ss))) + wt_args 0 A \ ss" by assumption show ?thesis using wt_x_le_\ssts[unfolded wt_args_def] by (simp add: v \x comp_def le_diff_eq add.assoc[symmetric] ZHMSet_plus[symmetric] zmset_of_plus[symmetric] hmsetmset_plus[symmetric] zmset_of_le) next case (PCoef x i) thus ?thesis using coef_gt_0[OF legal, unfolded zero_less_iff_1_le_hmset] by (simp add: zhmset_of_1 zero_less_iff_1_le_zhmset) qed qed lemma wt_subst: assumes legal: "legal_zpassign A" and wary_\: "wary_subst \" shows "wary s \ eval_ztpoly A (wt (subst \ s)) = eval_ztpoly (subst_zpassign \ A) (wt s)" proof (induct s rule: tm_induct_apps) case (apps \ ss) note ih = this(1) and wary_\ss = this(2) have wary_nth_ss: "\i. i < length ss \ wary (ss ! i)" using wary_args[OF _ wary_\ss] by force show ?case proof (cases \) case \: (Var x) show ?thesis proof (cases "\ x" rule: tm_exhaust_apps) case \x: (apps \ ts) have wary_\x: "wary (\ x)" using wary_\ wary_subst_def by blast have coef_subst: "\i. eval_tpoly A (zhmset_of_tpoly (coef_hd \ (i + length ts))) = eval_tpoly (subst_zpassign \ A) (zhmset_of_tpoly (coef_hd (Var x) i))" by (simp add: \x) have tedious_ary_arith: "arity_sym\<^sub>h (min_ground_head (Var x)) + (arity_sym\<^sub>h (min_ground_head \) - (of_nat (length ss) + of_nat (length ts))) = arity_sym\<^sub>h (min_ground_head \) - of_nat (length ts) + (arity_sym\<^sub>h (min_ground_head (Var x)) - of_nat (length ss))" if \_gt_0: "\\<^sub>h > 0" proof - obtain m where m: "of_nat m = arity_sym\<^sub>h (min_ground_head (Var x))" by (metis arity_sym\<^sub>h_if_\\<^sub>h_gt_0_E[OF \_gt_0]) obtain n where n: "of_nat n = arity_sym\<^sub>h (min_ground_head \)" by (metis arity_sym\<^sub>h_if_\\<^sub>h_gt_0_E[OF \_gt_0]) have "m \ length ss" unfolding of_nat_le_hmset[symmetric] m using wary_\ss[unfolded \] by (cases rule: wary_cases_apps\<^sub>h, clarsimp, metis arity_hd.simps(1) enat_ile enat_ord_simps(1) ground_heads_arity hmset_of_enat_inject hmset_of_enat_of_nat le_trans m min_ground_head_in_ground_heads of_nat_eq_enat of_nat_le_hmset_of_enat_iff) moreover have n_ge_len_ss_ts: "n \ length ss + length ts" proof - have "of_nat (length ss) + of_nat (length ts) \ arity_hd\<^sub>h \ + of_nat (length ts)" using wary_\ss wary_cases_apps\<^sub>h by fastforce also have "\ = arity_var\<^sub>h x + of_nat (length ts)" by (simp add: \) also have "\ \ arity\<^sub>h (\ x) + of_nat (length ts)" using wary_\ wary_subst_def by auto also have "\ = arity\<^sub>h (apps (Hd \) ts) + of_nat (length ts)" by (simp add: \x) also have "\ = arity_hd\<^sub>h \" using wary_\x[unfolded \x] by (cases rule: wary_cases_apps\<^sub>h, cases "arity_hd \", simp add: of_nat_add[symmetric] of_nat_minus_hmset[symmetric], metis \_gt_0 arity_hd_ne_infinity_if_\_gt_0 of_nat_0 of_nat_less_hmset) also have "\ \ arity_sym\<^sub>h (min_ground_head \)" using ground_heads_arity\<^sub>h min_ground_head_in_ground_heads by blast finally show ?thesis unfolding of_nat_le_hmset[symmetric] n by simp qed moreover have "n \ length ts" using n_ge_len_ss_ts by simp ultimately show ?thesis by (fold m n of_nat_add of_nat_minus_hmset, unfold of_nat_inject_hmset, fastforce) qed have "eval_tpoly A (zhmset_of_tpoly (wt (subst \ (apps (Hd (Var x)) ss)))) = eval_tpoly A (zhmset_of_tpoly (wt0 \)) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - (of_nat (length ts) + of_nat (length ss)))) + wt_args 0 A \ (ts @ map (subst \) ss)" by (simp del: apps_append add: apps_append[symmetric] \x wt_args_def comp_def) also have "\ = eval_tpoly A (zhmset_of_tpoly (wt0 \)) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - (of_nat (length ts) + of_nat (length ss)))) + wt_args 0 A \ ts + wt_args (length ts) A \ (map (subst \) ss)" by (simp add: wt_args_def zip_append_0_upt[of ts "map (subst \) ss", simplified]) also have "\ = eval_tpoly A (zhmset_of_tpoly (wt0 \)) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head \) - (of_nat (length ts) + of_nat (length ss)))) + wt_args 0 A \ ts + wt_args 0 (subst_zpassign \ A) (Var x) ss" by (auto intro!: arg_cong[of _ _ sum_list] nth_map_conv simp: wt_args_def coef_subst add.commute zhmset_of_times ih[OF nth_mem wary_nth_ss]) also have "\ = eval_tpoly (subst_zpassign \ A) (zhmset_of_tpoly (wt0 (Var x))) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h (min_ground_head (Var x)) - of_nat (length ss))) + wt_args 0 (subst_zpassign \ A) (Var x) ss" by (simp add: \x wt_args_def comp_def algebra_simps ring_distribs(1)[symmetric] zhmset_of_times zhmset_of_plus[symmetric] zhmset_of_0[symmetric]) (use tedious_ary_arith in fastforce) also have "\ = eval_tpoly (subst_zpassign \ A) (zhmset_of_tpoly (wt (apps (Hd (Var x)) ss)))" by (simp add: wt_args_def comp_def) finally show ?thesis unfolding \ by assumption qed next case \: (Sym f) have "eval_tpoly A (zhmset_of_tpoly (wt (subst \ (apps (Hd (Sym f)) ss)))) = zhmset_of (wt_sym f) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h f - of_nat (length ss))) + wt_args 0 A (Sym f) (map (subst \) ss)" by (simp add: wt_args_def comp_def) also have "\ = zhmset_of (wt_sym f) + zhmset_of (\\<^sub>h * (arity_sym\<^sub>h f - of_nat (length ss))) + wt_args 0 (subst_zpassign \ A) (Sym f) ss" by (auto simp: wt_args_def ih[OF _ wary_nth_ss] intro!: arg_cong[of _ _ sum_list] nth_map_conv) also have "\ = eval_tpoly (subst_zpassign \ A) (zhmset_of_tpoly (wt (apps (Hd (Sym f)) ss)))" by (simp add: wt_args_def comp_def) finally show ?thesis unfolding \ by assumption qed qed theorem gt_subst: assumes wary_\: "wary_subst \" shows "wary t \ wary s \ t >\<^sub>t s \ subst \ t >\<^sub>t subst \ s" proof (simp only: atomize_imp, rule measure_induct_rule[of "\(t, s). {#size t, size s#}" "\(t, s). wary t \ wary s \ t >\<^sub>t s \ subst \ t >\<^sub>t subst \ s" "(t, s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric]) fix t s assume ih: "\ta sa. {#size ta, size sa#} < {#size t, size s#} \ wary ta \ wary sa \ ta >\<^sub>t sa \ subst \ ta >\<^sub>t subst \ sa" and wary_t: "wary t" and wary_s: "wary s" and t_gt_s: "t >\<^sub>t s" show "subst \ t >\<^sub>t subst \ s" using t_gt_s proof cases case gt_wt_t_s: gt_wt have "wt (subst \ t) >\<^sub>p wt (subst \ s)" by (auto simp: gt_tpoly_def wary_s wary_t wt_subst[OF _ wary_\] intro: gt_wt_t_s[unfolded gt_tpoly_def, rule_format] elim: legal_subst_zpassign[OF _ wary_\]) thus ?thesis by (rule gt_wt) next assume wt_t_ge_s: "wt t \\<^sub>p wt s" have wt_\t_ge_\s: "wt (subst \ t) \\<^sub>p wt (subst \ s)" by (auto simp: ge_tpoly_def wary_s wary_t wt_subst[OF _ wary_\] intro: wt_t_ge_s[unfolded ge_tpoly_def, rule_format] elim: legal_subst_zpassign[OF _ wary_\]) { case gt_unary have wary_\t: "wary (subst \ t)" by (simp add: wary_subst_wary wary_t wary_\) show ?thesis proof (cases t) case Hd hence False using gt_unary(3) by simp thus ?thesis by sat next case t: (App t1 t2) hence t2: "t2 = arg t" by simp hence wary_t2: "wary t2" using wary_t by blast show ?thesis proof (cases "t2 = s") case True moreover have "subst \ t >\<^sub>t subst \ t2" using gt_sub_arg wary_\t unfolding t by simp ultimately show ?thesis by simp next case t2_ne_s: False hence t2_gt_s: "t2 >\<^sub>t s" using gt_unary(5) t2 by blast have "subst \ t2 >\<^sub>t subst \ s" by (rule ih[OF _ wary_t2 wary_s t2_gt_s]) (simp add: t) thus ?thesis by (metis gt_sub_arg gt_trans subst.simps(2) t wary_\ wary_\t wary_s wary_subst_wary wary_t2) qed qed } { case _: gt_diff note hd_t_gt_hd_s = this(2) have "head (subst \ t) >\<^sub>h\<^sub>d head (subst \ s)" by (meson hd_t_gt_hd_s wary_subst_ground_heads gt_hd_def rev_subsetD wary_\) thus ?thesis by (rule gt_diff[OF wt_\t_ge_\s]) } { case _: gt_same note hd_s_eq_hd_t = this(2) and extf = this(3) have hd_\t: "head (subst \ t) = head (subst \ s)" by (simp add: hd_s_eq_hd_t) { fix f assume f_in_grs: "f \ ground_heads (head (subst \ t))" let ?S = "set (args t) \ set (args s)" have extf_args_s_t: "extf f (>\<^sub>t) (args t) (args s)" using extf f_in_grs wary_subst_ground_heads wary_\ by blast have "extf f (>\<^sub>t) (map (subst \) (args t)) (map (subst \) (args s))" proof (rule extf_map[of ?S, OF _ _ _ _ _ _ extf_args_s_t]) show "\x \ ?S. \ subst \ x >\<^sub>t subst \ x" using gt_irrefl wary_t wary_s wary_args wary_\ wary_subst_wary by fastforce next show "\z \ ?S. \y \ ?S. \x \ ?S. subst \ z >\<^sub>t subst \ y \ subst \ y >\<^sub>t subst \ x \ subst \ z >\<^sub>t subst \ x" using gt_trans wary_t wary_s wary_args wary_\ wary_subst_wary by (metis Un_iff) next have sz_a: "\ta \ ?S. \sa \ ?S. {#size ta, size sa#} < {#size t, size s#}" by (fastforce intro: Max_lt_imp_lt_mset dest: size_in_args) show "\y \ ?S. \x \ ?S. y >\<^sub>t x \ subst \ y >\<^sub>t subst \ x" using ih sz_a size_in_args wary_t wary_s wary_args wary_\ wary_subst_wary by fastforce qed auto hence "extf f (>\<^sub>t) (args (subst \ t)) (args (subst \ s))" by (auto simp: hd_s_eq_hd_t intro: extf_compat_append_left) } hence "\f \ ground_heads (head (subst \ t)). extf f (>\<^sub>t) (args (subst \ t)) (args (subst \ s))" by blast thus ?thesis by (rule gt_same[OF wt_\t_ge_\s hd_\t]) } qed qed subsection \Totality on Ground Terms\ lemma wt_total_ground: assumes gr_t: "ground t" and gr_s: "ground s" shows "wt t >\<^sub>p wt s \ wt s >\<^sub>p wt t \ wt t =\<^sub>p wt s" unfolding gt_tpoly_def eq_tpoly_def by (subst (1 2 3) ground_eval_ztpoly_wt_eq[OF gr_t, of _ undefined], subst (1 2 3) ground_eval_ztpoly_wt_eq[OF gr_s, of _ undefined], auto) + theorem gt_total_ground: - assumes - extf_total: "\f. ext_total (extf f)" and - gr_t: "ground t" and - gr_s: "ground s" - shows "t >\<^sub>t s \ s >\<^sub>t t \ t = s" - using gr_t gr_s -proof (induct t arbitrary: s rule: tm_induct_apps) - case t: (apps \ ts) - note ih = this(1) and gr_t = this(2) and gr_s = this(3) + assumes extf_total: "\f. ext_total (extf f)" + shows "ground t \ ground s \ t >\<^sub>t s \ s >\<^sub>t t \ t = s" +proof (simp only: atomize_imp, + rule measure_induct_rule[of "\(t, s). {# size t, size s #}" + "\(t, s). ground t \ ground s \ t >\<^sub>t s \ s >\<^sub>t t \ t = s" "(t, s)", simplified prod.case], + simp only: split_paired_all prod.case atomize_imp[symmetric]) + fix t s :: "('s, 'v) tm" + assume + ih: "\ta sa. {# size ta, size sa #} < {# size t, size s #} \ ground ta \ ground sa \ + ta >\<^sub>t sa \ sa >\<^sub>t ta \ ta = sa" and + gr_t: "ground t" and gr_s: "ground s" - let ?t = "apps (Hd \) ts" + let ?case = "t >\<^sub>t s \ s >\<^sub>t t \ t = s" { - assume "wt ?t >\<^sub>p wt s" - hence "?t >\<^sub>t s" - by (rule gt_wt) - } - moreover - { - assume "wt s >\<^sub>p wt ?t" - hence "s >\<^sub>t ?t" + assume "wt t >\<^sub>p wt s" + hence "t >\<^sub>t s" by (rule gt_wt) } moreover { - assume "wt ?t =\<^sub>p wt s" - hence wt_t_ge_s: "wt ?t \\<^sub>p wt s" and wt_s_ge_t: "wt s \\<^sub>p wt ?t" + assume "wt s >\<^sub>p wt t" + hence "s >\<^sub>t t" + by (rule gt_wt) + } + moreover + { + assume "wt t =\<^sub>p wt s" + hence wt_t_ge_s: "wt t \\<^sub>p wt s" and wt_s_ge_t: "wt s \\<^sub>p wt t" by (simp add: eq_tpoly_def ge_tpoly_def)+ - have ?case - proof (cases s rule: tm_exhaust_apps) - case s: (apps \ ss) - obtain g where \: "\ = Sym g" - by (metis ground_head[OF gr_t] hd.collapse(2) head_apps tm.sel(1)) - obtain f where \: "\ = Sym f" - using s by (metis ground_head[OF gr_s] hd.collapse(2) head_apps tm.sel(1)) - - { - assume g_gt_f: "g >\<^sub>s f" - have "?t >\<^sub>t s" - by (rule gt_diff[OF wt_t_ge_s]) (simp add: \ \ s g_gt_f gt_hd_def) - } - moreover - { - assume f_gt_g: "f >\<^sub>s g" - have "s >\<^sub>t ?t" - by (rule gt_diff[OF wt_s_ge_t]) (simp add: \ \ s f_gt_g gt_hd_def) - } - moreover - { - assume g_eq_f: "g = f" - hence hd_t: "head ?t = head s" - using \ \ t s by force - note hd_s = hd_t[symmetric] + obtain g where \: "head t = Sym g" + by (metis ground_head[OF gr_t] hd.collapse(2)) + obtain f where \: "head s = Sym f" + by (metis ground_head[OF gr_s] hd.collapse(2)) - have gr_ts: "\t \ set ts. ground t" - using gr_t by auto - have gr_ss: "\s \ set ss. ground s" - using gr_s s by auto + { + assume g_gt_f: "g >\<^sub>s f" + have "t >\<^sub>t s" + by (rule gt_diff[OF wt_t_ge_s]) (simp add: \ \ g_gt_f gt_hd_def) + } + moreover + { + assume f_gt_g: "f >\<^sub>s g" + have "s >\<^sub>t t" + by (rule gt_diff[OF wt_s_ge_t]) (simp add: \ \ f_gt_g gt_hd_def) + } + moreover + { + assume g_eq_f: "g = f" + hence hd_t: "head t = head s" + using \ \ by force + note hd_s = hd_t[symmetric] - have ?thesis - proof (cases "ts = ss") - case ts_eq_ss: True - show ?thesis - using s \ \ g_eq_f ts_eq_ss by blast - next - case False - hence "extf g (>\<^sub>t) ts ss \ extf g (>\<^sub>t) ss ts" - using ih gr_ss gr_ts - ext_total.total[OF extf_total, rule_format, of "set ts" "set ss" "(>\<^sub>t)" ts ss g] - by blast - moreover - { - assume extf: "extf g (>\<^sub>t) ts ss" - have "?t >\<^sub>t s" - by (rule gt_same[OF wt_t_ge_s hd_t]) (simp add: extf \ s) - } - moreover - { - assume extf: "extf g (>\<^sub>t) ss ts" - have "s >\<^sub>t ?t" - by (rule gt_same[OF wt_s_ge_t hd_s]) (simp add: extf[unfolded g_eq_f] \ s) - } - ultimately show ?thesis - by sat - qed - } - ultimately show ?thesis - using gt_sym_total by blast - qed + let ?ts = "args t" + let ?ss = "args s" + + have gr_ts: "\t \ set ?ts. ground t" + using gr_t ground_args by auto + have gr_ss: "\s \ set ?ss. ground s" + using gr_s ground_args by auto + + have ?case + proof (cases "?ts = ?ss") + case ts_eq_ss: True + show ?thesis + using \ \ g_eq_f ts_eq_ss by (simp add: tm_expand_apps) + next + case False + hence "extf g (>\<^sub>t) ?ts ?ss \ extf g (>\<^sub>t) ?ss ?ts" + using ih gr_ss gr_ts less_multiset_doubletons + ext_total.total[OF extf_total, rule_format, of "set ?ts \ set ?ss" "(>\<^sub>t)" ?ts ?ss g] + by (metis Un_commute Un_iff in_lists_iff_set size_in_args sup_ge2) + moreover + { + assume extf: "extf g (>\<^sub>t) ?ts ?ss" + have "t >\<^sub>t s" + by (rule gt_same[OF wt_t_ge_s hd_t]) (simp add: extf \) + } + moreover + { + assume extf: "extf g (>\<^sub>t) ?ss ?ts" + have "s >\<^sub>t t" + by (rule gt_same[OF wt_s_ge_t hd_s]) (simp add: extf[unfolded g_eq_f] \) + } + ultimately show ?thesis + by sat + qed + } + ultimately have ?case + using gt_sym_total by blast } ultimately show ?case using wt_total_ground[OF gr_t gr_s] by fast qed subsection \Well-foundedness\ abbreviation gtw :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix ">\<^sub>t\<^sub>w" 50) where "(>\<^sub>t\<^sub>w) \ \t s. wary t \ wary s \ t >\<^sub>t s" abbreviation gtwg :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix ">\<^sub>t\<^sub>w\<^sub>g" 50) where "(>\<^sub>t\<^sub>w\<^sub>g) \ \t s. ground t \ t >\<^sub>t\<^sub>w s" lemma ground_gt_unary: assumes gr_t: "ground t" shows "\ gt_unary t s" proof assume gt_unary_t_s: "gt_unary t s" hence "t >\<^sub>t s" using gt_iff_wt_unary_diff_same by blast hence gr_s: "ground s" using gr_t gt_imp_vars by blast have ngr_t_or_s: "\ ground t \ \ ground s" using gt_unary_t_s by cases (blast dest: ground_head not_comp_hd_imp_Var) show False using gr_t gr_s ngr_t_or_s by sat qed theorem gt_wf: "wfP (\s t. t >\<^sub>t\<^sub>w s)" proof - have ground_wfP: "wfP (\s t. t >\<^sub>t\<^sub>w\<^sub>g s)" unfolding wfP_iff_no_inf_chain proof assume "\f. inf_chain (>\<^sub>t\<^sub>w\<^sub>g) f" then obtain t where t_bad: "bad (>\<^sub>t\<^sub>w\<^sub>g) t" unfolding inf_chain_def bad_def by blast let ?ff = "worst_chain (>\<^sub>t\<^sub>w\<^sub>g) (\t s. size t > size s)" let ?A = min_passign note wf_sz = wf_app[OF wellorder_class.wf, of size, simplified] have ffi_ground: "\i. ground (?ff i)" and ffi_wary: "\i. wary (?ff i)" using worst_chain_bad[OF wf_sz t_bad, unfolded inf_chain_def] by fast+ have "inf_chain (>\<^sub>t\<^sub>w\<^sub>g) ?ff" by (rule worst_chain_bad[OF wf_sz t_bad]) hence bad_wt_diff_same: "inf_chain (\t s. ground t \ (gt_wt t s \ gt_diff t s \ gt_same t s)) ?ff" unfolding inf_chain_def using gt_iff_wt_unary_diff_same ground_gt_unary by blast have wf_wt: "wf {(s, t). ground t \ gt_wt t s}" by (rule wf_subset[OF wf_app[of _ "eval_tpoly ?A \ wt", OF wf_less_hmultiset]], simp add: gt_wt.simps gt_tpoly_def, fold zhmset_of_less, auto simp: legal_min_zpassign gt_wt.simps gt_tpoly_def) have wt_O_diff_same: "{(s, t). ground t \ gt_wt t s} O {(s, t). ground t \ wt t =\<^sub>p wt s \ (gt_diff t s \ gt_same t s)} \ {(s, t). ground t \ gt_wt t s}" unfolding gt_wt.simps gt_diff.simps gt_same.simps by (auto intro: ge_gt_tpoly_trans) have wt_diff_same_as_union: "{(s, t). ground t \ (gt_wt t s \ gt_diff t s \ gt_same t s)} = {(s, t). ground t \ gt_wt t s} \ {(s, t). ground t \ wt t =\<^sub>p wt s \ (gt_diff t s \ gt_same t s)}" using gt_ge_tpoly_trans gt_tpoly_irrefl wt_ge_vars wt_total_ground by (fastforce simp: gt_wt.simps gt_diff.simps gt_same.simps) obtain k1 where bad_diff_same: "inf_chain (\t s. ground t \ wt t =\<^sub>p wt s \ (gt_diff t s \ gt_same t s)) (\i. ?ff (i + k1))" using wf_infinite_down_chain_compatible[OF wf_wt _ wt_O_diff_same, of ?ff] bad_wt_diff_same unfolding inf_chain_def wt_diff_same_as_union[symmetric] by auto have "wf {(s, t). ground s \ ground t \ wt t =\<^sub>p wt s \ sym (head t) >\<^sub>s sym (head s)}" using gt_sym_wf unfolding wfP_def wf_iff_no_infinite_down_chain by fast moreover have "{(s, t). ground t \ wt t =\<^sub>p wt s \ gt_diff t s} \ {(s, t). ground s \ ground t \ wt t =\<^sub>p wt s \ sym (head t) >\<^sub>s sym (head s)}" proof (clarsimp, intro conjI) fix s t assume gr_t: "ground t" and gt_diff_t_s: "gt_diff t s" thus gr_s: "ground s" using gt_iff_wt_unary_diff_same gt_imp_vars by fastforce show "sym (head t) >\<^sub>s sym (head s)" using gt_diff_t_s by cases (simp add: gt_hd_def gr_s gr_t ground_hd_in_ground_heads) qed ultimately have wf_diff: "wf {(s, t). ground t \ wt t =\<^sub>p wt s \ gt_diff t s}" by (rule wf_subset) have diff_O_same: "{(s, t). ground t \ wt t =\<^sub>p wt s \ gt_diff t s} O {(s, t). ground t \ wt t =\<^sub>p wt s \ gt_same t s} \ {(s, t). ground t \ wt t =\<^sub>p wt s \ gt_diff t s}" unfolding gt_diff.simps gt_same.simps by (auto intro: ge_ge_tpoly_trans simp: eq_tpoly_def) have diff_same_as_union: "{(s, t). ground t \ wt t =\<^sub>p wt s \ (gt_diff t s \ gt_same t s)} = {(s, t). ground t \ wt t =\<^sub>p wt s \ gt_diff t s} \ {(s, t). ground t \ wt t =\<^sub>p wt s \ gt_same t s}" by auto obtain k2 where bad_same: "inf_chain (\t s. ground t \ wt t =\<^sub>p wt s \ gt_same t s) (\i. ?ff (i + k2))" using wf_infinite_down_chain_compatible[OF wf_diff _ diff_O_same, of "\i. ?ff (i + k1)"] bad_diff_same unfolding inf_chain_def diff_same_as_union[symmetric] by (auto simp: add.assoc) hence hd_sym: "\i. is_Sym (head (?ff (i + k2)))" unfolding inf_chain_def by (simp add: ground_head) define f where "f = sym (head (?ff k2))" define w where "w = eval_tpoly ?A (wt (?ff k2))" have "head (?ff (i + k2)) = Sym f \ eval_tpoly ?A (wt (?ff (i + k2))) = w" for i proof (induct i) case 0 thus ?case by (auto simp: f_def w_def hd.collapse(2)[OF hd_sym, of 0, simplified]) next case (Suc ia) thus ?case using bad_same unfolding inf_chain_def gt_same.simps zhmset_of_inject[symmetric] by (simp add: eq_tpoly_def legal_min_zpassign) qed note hd_eq_f = this[THEN conjunct1] and wt_eq_w = this[THEN conjunct2] define max_args where "max_args = (if \\<^sub>h = 0 then sum_coefs w else the_enat (arity_sym f))" have nargs_le_max_args: "num_args (?ff (i + k2)) \ max_args" for i proof (cases "\\<^sub>h = 0") case \_ne_0: False hence ary_f_ne_inf: "arity_sym f \ \" using arity_sym_ne_infinity_if_\_gt_0 of_nat_0 by blast have "enat (num_args (worst_chain (\t s. ground t \ t >\<^sub>t\<^sub>w s) (\t s. size s < size t) (i + k2))) \ arity_sym f" using wary_num_args_le_arity_head[OF ffi_wary[of "i + k2"]] by (simp add: hd_eq_f) with \_ne_0 show ?thesis by (simp del: enat_ord_simps add: max_args_def enat_ord_simps(1)[symmetric] enat_the_enat_iden[OF ary_f_ne_inf]) next case \_eq_0: True show ?thesis using sum_coefs_ge_num_args_if_\\<^sub>h_eq_0[OF legal_min_passign \_eq_0 ffi_wary[of "i + k2"]] by (simp add: max_args_def \_eq_0 wt_eq_w) qed let ?U_of = "\i. set (args (?ff (i + k2)))" define U where "U = (\i. ?U_of i)" have gr_u: "\u. u \ U \ ground u" unfolding U_def by (blast dest: ground_args[OF _ ffi_ground]) have wary_u: "\u. u \ U \ wary u" unfolding U_def by (blast dest: wary_args[OF _ ffi_wary]) have "\ bad (>\<^sub>t\<^sub>w\<^sub>g) u" if u_in: "u \ ?U_of i" for u i proof assume u_bad: "bad (>\<^sub>t\<^sub>w\<^sub>g) u" have sz_u: "size u < size (?ff (i + k2))" by (rule size_in_args[OF u_in]) show False proof (cases "i + k2") case 0 thus False using sz_u min_worst_chain_0[OF wf_sz u_bad] by simp next case Suc hence gt: "?ff (i + k2 - 1) >\<^sub>t\<^sub>w ?ff (i + k2)" using worst_chain_pred[OF wf_sz t_bad] by auto moreover have "?ff (i + k2) >\<^sub>t\<^sub>w u" using gt gt_proper_sub sub_args sz_u u_in wary_args by auto ultimately have "?ff (i + k2 - 1) >\<^sub>t\<^sub>w u" using gt_trans by blast thus False using Suc sz_u min_worst_chain_Suc[OF wf_sz u_bad] ffi_ground by fastforce qed qed hence u_good: "\u. u \ U \ \ bad (>\<^sub>t\<^sub>w\<^sub>g) u" unfolding U_def by blast let ?gtwu = "\t s. t \ U \ t >\<^sub>t\<^sub>w s" have gtwu_irrefl: "\x. \ ?gtwu x x" using gt_irrefl by auto have "\i j. \t \ set (args (?ff (i + k2))). \s \ set (args (?ff (j + k2))). t >\<^sub>t s \ t \ U \ t >\<^sub>t\<^sub>w s" using wary_u unfolding U_def by blast moreover have "\i. extf f (>\<^sub>t) (args (?ff (i + k2))) (args (?ff (Suc i + k2)))" using bad_same hd_eq_f unfolding inf_chain_def gt_same.simps by auto ultimately have "\i. extf f ?gtwu (args (?ff (i + k2))) (args (?ff (Suc i + k2)))" by (rule extf_mono_strong) hence "inf_chain (extf f ?gtwu) (\i. args (?ff (i + k2)))" unfolding inf_chain_def by blast hence nwf_ext: "\ wfP (\xs ys. length ys \ max_args \ length xs \ max_args \ extf f ?gtwu ys xs)" unfolding inf_chain_def wfP_def wf_iff_no_infinite_down_chain using nargs_le_max_args by fast have gtwu_le_gtwg: "?gtwu \ (>\<^sub>t\<^sub>w\<^sub>g)" by (auto intro!: gr_u) have "wfP (\s t. ?gtwu t s)" unfolding wfP_iff_no_inf_chain proof (intro notI, elim exE) fix f assume bad_f: "inf_chain ?gtwu f" hence bad_f0: "bad ?gtwu (f 0)" by (rule inf_chain_bad) hence "f 0 \ U" using bad_f unfolding inf_chain_def by blast hence "\ bad (>\<^sub>t\<^sub>w\<^sub>g) (f 0)" using u_good by blast hence "\ bad ?gtwu (f 0)" using bad_f inf_chain_bad inf_chain_subset[OF _ gtwu_le_gtwg] by blast thus False using bad_f0 by sat qed hence wf_ext: "wfP (\xs ys. length ys \ max_args \ length xs \ max_args \ extf f ?gtwu ys xs)" using extf_wf_bounded[of ?gtwu] gtwu_irrefl by blast show False using nwf_ext wf_ext by blast qed let ?subst = "subst grounding_\" have "wfP (\s t. ?subst t >\<^sub>t\<^sub>w\<^sub>g ?subst s)" by (rule wfP_app[OF ground_wfP]) hence "wfP (\s t. ?subst t >\<^sub>t\<^sub>w ?subst s)" by (simp add: ground_grounding_\) thus ?thesis by (auto intro: wfP_subset wary_subst_wary[OF wary_grounding_\] gt_subst[OF wary_grounding_\]) qed end end diff --git a/thys/Lambda_Free_RPOs/Extension_Orders.thy b/thys/Lambda_Free_RPOs/Extension_Orders.thy --- a/thys/Lambda_Free_RPOs/Extension_Orders.thy +++ b/thys/Lambda_Free_RPOs/Extension_Orders.thy @@ -1,2056 +1,2071 @@ (* Title: Extension Orders Author: Heiko Becker , 2016 Author: Jasmin Blanchette , 2016 Author: Dmitriy Traytel , 2014 Maintainer: Jasmin Blanchette *) section \Extension Orders\ theory Extension_Orders imports Lambda_Free_Util Infinite_Chain "HOL-Cardinals.Wellorder_Extension" begin text \ This theory defines locales for categorizing extension orders used for orders on \\\-free higher-order terms and defines variants of the lexicographic and multiset orders. \ subsection \Locales\ locale ext = fixes ext :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" assumes mono_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt' y x) \ ext gt ys xs \ ext gt' ys xs" and map: "finite A \ ys \ lists A \ xs \ lists A \ (\x \ A. \ gt (f x) (f x)) \ (\z \ A. \y \ A. \x \ A. gt (f z) (f y) \ gt (f y) (f x) \ gt (f z) (f x)) \ (\y \ A. \x \ A. gt y x \ gt (f y) (f x)) \ ext gt ys xs \ ext gt (map f ys) (map f xs)" begin lemma mono[mono]: "gt \ gt' \ ext gt \ ext gt'" using mono_strong by blast end locale ext_irrefl = ext + assumes irrefl: "(\x \ set xs. \ gt x x) \ \ ext gt xs xs" locale ext_trans = ext + assumes trans: "zs \ lists A \ ys \ lists A \ xs \ lists A \ (\z \ A. \y \ A. \x \ A. gt z y \ gt y x \ gt z x) \ ext gt zs ys \ ext gt ys xs \ ext gt zs xs" locale ext_irrefl_before_trans = ext_irrefl + assumes trans_from_irrefl: "finite A \ zs \ lists A \ ys \ lists A \ xs \ lists A \ (\x \ A. \ gt x x) \ (\z \ A. \y \ A. \x \ A. gt z y \ gt y x \ gt z x) \ ext gt zs ys \ ext gt ys xs \ ext gt zs xs" locale ext_trans_before_irrefl = ext_trans + assumes irrefl_from_trans: "(\z \ set xs. \y \ set xs. \x \ set xs. gt z y \ gt y x \ gt z x) \ (\x \ set xs. \ gt x x) \ \ ext gt xs xs" locale ext_irrefl_trans_strong = ext_irrefl + assumes trans_strong: "(\z \ set zs. \y \ set ys. \x \ set xs. gt z y \ gt y x \ gt z x) \ ext gt zs ys \ ext gt ys xs \ ext gt zs xs" sublocale ext_irrefl_trans_strong < ext_irrefl_before_trans by standard (erule irrefl, metis in_listsD trans_strong) sublocale ext_irrefl_trans_strong < ext_trans by standard (metis in_listsD trans_strong) sublocale ext_irrefl_trans_strong < ext_trans_before_irrefl by standard (rule irrefl) locale ext_snoc = ext + assumes snoc: "ext gt (xs @ [x]) xs" locale ext_compat_cons = ext + assumes compat_cons: "ext gt ys xs \ ext gt (x # ys) (x # xs)" begin lemma compat_append_left: "ext gt ys xs \ ext gt (zs @ ys) (zs @ xs)" by (induct zs) (auto intro: compat_cons) end locale ext_compat_snoc = ext + assumes compat_snoc: "ext gt ys xs \ ext gt (ys @ [x]) (xs @ [x])" begin lemma compat_append_right: "ext gt ys xs \ ext gt (ys @ zs) (xs @ zs)" by (induct zs arbitrary: xs ys rule: rev_induct) (auto intro: compat_snoc simp del: append_assoc simp: append_assoc[symmetric]) end locale ext_compat_list = ext + assumes compat_list: "y \ x \ gt y x \ ext gt (xs @ y # xs') (xs @ x # xs')" locale ext_singleton = ext + assumes singleton: "y \ x \ ext gt [y] [x] \ gt y x" locale ext_compat_list_strong = ext_compat_cons + ext_compat_snoc + ext_singleton begin lemma compat_list: "y \ x \ gt y x \ ext gt (xs @ y # xs') (xs @ x # xs')" using compat_append_left[of gt "y # xs'" "x # xs'" xs] compat_append_right[of gt, of "[y]" "[x]" xs'] singleton[of y x gt] by fastforce end sublocale ext_compat_list_strong < ext_compat_list by standard (fact compat_list) locale ext_total = ext + - assumes total: "(\y \ B. \x \ A. gt y x \ gt x y \ y = x) \ ys \ lists B \ xs \ lists A \ + assumes total: "(\y \ A. \x \ A. gt y x \ gt x y \ y = x) \ ys \ lists A \ xs \ lists A \ ext gt ys xs \ ext gt xs ys \ ys = xs" locale ext_wf = ext + assumes wf: "wfP (\x y. gt y x) \ wfP (\xs ys. ext gt ys xs)" locale ext_hd_or_tl = ext + assumes hd_or_tl: "(\z y x. gt z y \ gt y x \ gt z x) \ (\y x. gt y x \ gt x y \ y = x) \ length ys = length xs \ ext gt (y # ys) (x # xs) \ gt y x \ ext gt ys xs" locale ext_wf_bounded = ext_irrefl_before_trans + ext_hd_or_tl begin context fixes gt :: "'a \ 'a \ bool" assumes gt_irrefl: "\z. \ gt z z" and gt_trans: "\z y x. gt z y \ gt y x \ gt z x" and gt_total: "\y x. gt y x \ gt x y \ y = x" and gt_wf: "wfP (\x y. gt y x)" begin lemma irrefl_gt: "\ ext gt xs xs" using irrefl gt_irrefl by simp lemma trans_gt: "ext gt zs ys \ ext gt ys xs \ ext gt zs xs" by (rule trans_from_irrefl[of "set zs \ set ys \ set xs" zs ys xs gt]) (auto intro: gt_trans simp: gt_irrefl) lemma hd_or_tl_gt: "length ys = length xs \ ext gt (y # ys) (x # xs) \ gt y x \ ext gt ys xs" by (rule hd_or_tl) (auto intro: gt_trans simp: gt_total) lemma wf_same_length_if_total: "wfP (\xs ys. length ys = n \ length xs = n \ ext gt ys xs)" proof (induct n) case 0 thus ?case unfolding wfP_def wf_def using irrefl by auto next case (Suc n) note ih = this(1) define gt_hd where "\ys xs. gt_hd ys xs \ gt (hd ys) (hd xs)" define gt_tl where "\ys xs. gt_tl ys xs \ ext gt (tl ys) (tl xs)" have hd_tl: "gt_hd ys xs \ gt_tl ys xs" if len_ys: "length ys = Suc n" and len_xs: "length xs = Suc n" and ys_gt_xs: "ext gt ys xs" for n ys xs using len_ys len_xs ys_gt_xs unfolding gt_hd_def gt_tl_def by (cases xs; cases ys) (auto simp: hd_or_tl_gt) show ?case unfolding wfP_iff_no_inf_chain proof (intro notI) let ?gtsn = "\ys xs. length ys = n \ length xs = n \ ext gt ys xs" let ?gtsSn = "\ys xs. length ys = Suc n \ length xs = Suc n \ ext gt ys xs" let ?gttlSn = "\ys xs. length ys = Suc n \ length xs = Suc n \ gt_tl ys xs" assume "\f. inf_chain ?gtsSn f" then obtain xs where xs_bad: "bad ?gtsSn xs" unfolding inf_chain_def bad_def by blast let ?ff = "worst_chain ?gtsSn gt_hd" have wf_hd: "wf {(xs, ys). gt_hd ys xs}" unfolding gt_hd_def by (rule wfP_app[OF gt_wf, of hd, unfolded wfP_def]) have "inf_chain ?gtsSn ?ff" by (rule worst_chain_bad[OF wf_hd xs_bad]) moreover have "\ gt_hd (?ff i) (?ff (Suc i))" for i by (rule worst_chain_not_gt[OF wf_hd xs_bad]) (blast intro: trans_gt) ultimately have tl_bad: "inf_chain ?gttlSn ?ff" unfolding inf_chain_def using hd_tl by blast have "\ inf_chain ?gtsn (tl \ ?ff)" using wfP_iff_no_inf_chain[THEN iffD1, OF ih] by blast hence tl_good: "\ inf_chain ?gttlSn ?ff" unfolding inf_chain_def gt_tl_def by force show False using tl_bad tl_good by sat qed qed lemma wf_bounded_if_total: "wfP (\xs ys. length ys \ n \ length xs \ n \ ext gt ys xs)" unfolding wfP_iff_no_inf_chain proof (intro notI, induct n rule: less_induct) case (less n) note ih = this(1) and ex_bad = this(2) let ?gtsle = "\ys xs. length ys \ n \ length xs \ n \ ext gt ys xs" obtain xs where xs_bad: "bad ?gtsle xs" using ex_bad unfolding inf_chain_def bad_def by blast let ?ff = "worst_chain ?gtsle (\ys xs. length ys > length xs)" note wf_len = wf_app[OF wellorder_class.wf, of length, simplified] have ff_bad: "inf_chain ?gtsle ?ff" by (rule worst_chain_bad[OF wf_len xs_bad]) have ffi_bad: "\i. bad ?gtsle (?ff i)" by (rule inf_chain_bad[OF ff_bad]) have len_le_n: "\i. length (?ff i) \ n" using worst_chain_pred[OF wf_len xs_bad] by simp have len_le_Suc: "\i. length (?ff i) \ length (?ff (Suc i))" using worst_chain_not_gt[OF wf_len xs_bad] not_le_imp_less by (blast intro: trans_gt) show False proof (cases "\k. length (?ff k) = n") case False hence len_lt_n: "\i. length (?ff i) < n" using len_le_n by (blast intro: le_neq_implies_less) hence nm1_le: "n - 1 < n" by fastforce let ?gtslt = "\ys xs. length ys \ n - 1 \ length xs \ n - 1 \ ext gt ys xs" have "inf_chain ?gtslt ?ff" using ff_bad len_lt_n unfolding inf_chain_def by (metis (no_types, lifting) Suc_diff_1 le_antisym nat_neq_iff not_less0 not_less_eq_eq) thus False using ih[OF nm1_le] by blast next case True then obtain k where len_eq_n: "length (?ff k) = n" by blast let ?gtssl = "\ys xs. length ys = n \ length xs = n \ ext gt ys xs" have len_eq_n: "length (?ff (i + k)) = n" for i by (induct i) (simp add: len_eq_n, metis (lifting) len_le_n len_le_Suc add_Suc dual_order.antisym) have "inf_chain ?gtsle (\i. ?ff (i + k))" by (rule inf_chain_offset[OF ff_bad]) hence "inf_chain ?gtssl (\i. ?ff (i + k))" unfolding inf_chain_def using len_eq_n by presburger hence "\ wfP (\xs ys. ?gtssl ys xs)" using wfP_iff_no_inf_chain by blast thus False using wf_same_length_if_total[of n] by sat qed qed end context fixes gt :: "'a \ 'a \ bool" assumes gt_irrefl: "\z. \ gt z z" and gt_wf: "wfP (\x y. gt y x)" begin lemma wf_bounded: "wfP (\xs ys. length ys \ n \ length xs \ n \ ext gt ys xs)" proof - obtain Ge' where gt_sub_Ge': "{(x, y). gt y x} \ Ge'" and Ge'_wo: "Well_order Ge'" and Ge'_fld: "Field Ge' = UNIV" using total_well_order_extension[OF gt_wf[unfolded wfP_def]] by blast define gt' where "\y x. gt' y x \ y \ x \ (x, y) \ Ge'" have gt_imp_gt': "gt \ gt'" by (auto simp: gt'_def gt_irrefl intro: gt_sub_Ge'[THEN subsetD]) have gt'_irrefl: "\z. \ gt' z z" unfolding gt'_def by simp have gt'_trans: "\z y x. gt' z y \ gt' y x \ gt' z x" using Ge'_wo unfolding gt'_def well_order_on_def linear_order_on_def partial_order_on_def preorder_on_def trans_def antisym_def by blast have "wf {(x, y). (x, y) \ Ge' \ x \ y}" by (rule Ge'_wo[unfolded well_order_on_def set_diff_eq case_prod_eta[symmetric, of "\xy. xy \ Ge' \ xy \ Id"] pair_in_Id_conv, THEN conjunct2]) moreover have "\y x. (x, y) \ Ge' \ x \ y \ y \ x \ (x, y) \ Ge'" by auto ultimately have gt'_wf: "wfP (\x y. gt' y x)" unfolding wfP_def gt'_def by simp have gt'_total: "\x y. gt' y x \ gt' x y \ y = x" using Ge'_wo unfolding gt'_def well_order_on_def linear_order_on_def total_on_def Ge'_fld by blast have "wfP (\xs ys. length ys \ n \ length xs \ n \ ext gt' ys xs)" using wf_bounded_if_total gt'_total gt'_irrefl gt'_trans gt'_wf by blast thus ?thesis by (rule wfP_subset) (auto intro: mono[OF gt_imp_gt', THEN predicate2D]) qed end end subsection \Lexicographic Extension\ inductive lexext :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" for gt where lexext_Nil: "lexext gt (y # ys) []" | lexext_Cons: "gt y x \ lexext gt (y # ys) (x # xs)" | lexext_Cons_eq: "lexext gt ys xs \ lexext gt (x # ys) (x # xs)" lemma lexext_simps[simp]: "lexext gt ys [] \ ys \ []" "\ lexext gt [] xs" "lexext gt (y # ys) (x # xs) \ gt y x \ x = y \ lexext gt ys xs" proof show "lexext gt ys [] \ (ys \ [])" by (metis lexext.cases list.distinct(1)) next show "ys \ [] \ lexext gt ys []" by (metis lexext_Nil list.exhaust) next show "\ lexext gt [] xs" using lexext.cases by auto next show "lexext gt (y # ys) (x # xs) = (gt y x \ x = y \ lexext gt ys xs)" proof - have fwdd: "lexext gt (y # ys) (x # xs) \ gt y x \ x = y \ lexext gt ys xs" proof assume "lexext gt (y # ys) (x # xs)" thus "gt y x \ x = y \ lexext gt ys xs" using lexext.cases by blast qed have backd: "gt y x \ x = y \ lexext gt ys xs \ lexext gt (y # ys) (x # xs)" by (simp add: lexext_Cons lexext_Cons_eq) show "lexext gt (y # ys) (x # xs) = (gt y x \ x = y \ lexext gt ys xs)" using fwdd backd by blast qed qed lemma lexext_mono_strong: assumes "\y \ set ys. \x \ set xs. gt y x \ gt' y x" and "lexext gt ys xs" shows "lexext gt' ys xs" using assms by (induct ys xs rule: list_induct2') auto lemma lexext_map_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt (f y) (f x)) \ lexext gt ys xs \ lexext gt (map f ys) (map f xs)" by (induct ys xs rule: list_induct2') auto lemma lexext_irrefl: assumes "\x \ set xs. \ gt x x" shows "\ lexext gt xs xs" using assms by (induct xs) auto lemma lexext_trans_strong: assumes "\z \ set zs. \y \ set ys. \x \ set xs. gt z y \ gt y x \ gt z x" and "lexext gt zs ys" and "lexext gt ys xs" shows "lexext gt zs xs" using assms proof (induct zs arbitrary: ys xs) case (Cons z zs) note zs_trans = this(1) show ?case using Cons(2-4) proof (induct ys arbitrary: xs rule: list.induct) case (Cons y ys) note ys_trans = this(1) and gt_trans = this(2) and zzs_gt_yys = this(3) and yys_gt_xs = this(4) show ?case proof (cases xs) case xs: (Cons x xs) thus ?thesis proof (unfold xs) note yys_gt_xxs = yys_gt_xs[unfolded xs] note gt_trans = gt_trans[unfolded xs] let ?case = "lexext gt (z # zs) (x # xs)" { assume "gt z y" and "gt y x" hence ?case using gt_trans by simp } moreover { assume "gt z y" and "x = y" hence ?case by simp } moreover { assume "y = z" and "gt y x" hence ?case by simp } moreover { assume y_eq_z: "y = z" and zs_gt_ys: "lexext gt zs ys" and x_eq_y: "x = y" and ys_gt_xs: "lexext gt ys xs" have "lexext gt zs xs" by (rule zs_trans[OF _ zs_gt_ys ys_gt_xs]) (meson gt_trans[simplified]) hence ?case by (simp add: x_eq_y y_eq_z) } ultimately show ?case using zzs_gt_yys yys_gt_xxs by force qed qed auto qed auto qed auto lemma lexext_snoc: "lexext gt (xs @ [x]) xs" by (induct xs) auto lemmas lexext_compat_cons = lexext_Cons_eq lemma lexext_compat_snoc_if_same_length: assumes "length ys = length xs" and "lexext gt ys xs" shows "lexext gt (ys @ [x]) (xs @ [x])" using assms(2,1) by (induct rule: lexext.induct) auto lemma lexext_compat_list: "gt y x \ lexext gt (xs @ y # xs') (xs @ x # xs')" by (induct xs) auto lemma lexext_singleton: "lexext gt [y] [x] \ gt y x" by simp lemma lexext_total: "(\y \ B. \x \ A. gt y x \ gt x y \ y = x) \ ys \ lists B \ xs \ lists A \ lexext gt ys xs \ lexext gt xs ys \ ys = xs" by (induct ys xs rule: list_induct2') auto lemma lexext_hd_or_tl: "lexext gt (y # ys) (x # xs) \ gt y x \ lexext gt ys xs" by auto interpretation lexext: ext lexext by standard (fact lexext_mono_strong, rule lexext_map_strong, metis in_listsD) interpretation lexext: ext_irrefl_trans_strong lexext by standard (fact lexext_irrefl, fact lexext_trans_strong) interpretation lexext: ext_snoc lexext by standard (fact lexext_snoc) interpretation lexext: ext_compat_cons lexext by standard (fact lexext_compat_cons) interpretation lexext: ext_compat_list lexext by standard (rule lexext_compat_list) interpretation lexext: ext_singleton lexext by standard (rule lexext_singleton) interpretation lexext: ext_total lexext by standard (fact lexext_total) interpretation lexext: ext_hd_or_tl lexext by standard (rule lexext_hd_or_tl) interpretation lexext: ext_wf_bounded lexext by standard subsection \Reverse (Right-to-Left) Lexicographic Extension\ abbreviation lexext_rev :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" where "lexext_rev gt ys xs \ lexext gt (rev ys) (rev xs)" lemma lexext_rev_simps[simp]: "lexext_rev gt ys [] \ ys \ []" "\ lexext_rev gt [] xs" "lexext_rev gt (ys @ [y]) (xs @ [x]) \ gt y x \ x = y \ lexext_rev gt ys xs" by simp+ lemma lexext_rev_cons_cons: assumes "length ys = length xs" shows "lexext_rev gt (y # ys) (x # xs) \ lexext_rev gt ys xs \ ys = xs \ gt y x" using assms proof (induct arbitrary: y x rule: rev_induct2) case Nil thus ?case by simp next case (snoc y' ys x' xs) show ?case using snoc(2) by auto qed lemma lexext_rev_mono_strong: assumes "\y \ set ys. \x \ set xs. gt y x \ gt' y x" and "lexext_rev gt ys xs" shows "lexext_rev gt' ys xs" using assms by (simp add: lexext_mono_strong) lemma lexext_rev_map_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt (f y) (f x)) \ lexext_rev gt ys xs \ lexext_rev gt (map f ys) (map f xs)" by (simp add: lexext_map_strong rev_map) lemma lexext_rev_irrefl: assumes "\x \ set xs. \ gt x x" shows "\ lexext_rev gt xs xs" using assms by (simp add: lexext_irrefl) lemma lexext_rev_trans_strong: assumes "\z \ set zs. \y \ set ys. \x \ set xs. gt z y \ gt y x \ gt z x" and "lexext_rev gt zs ys" and "lexext_rev gt ys xs" shows "lexext_rev gt zs xs" using assms(1) lexext_trans_strong[OF _ assms(2,3), unfolded set_rev] by sat lemma lexext_rev_compat_cons_if_same_length: assumes "length ys = length xs" and "lexext_rev gt ys xs" shows "lexext_rev gt (x # ys) (x # xs)" using assms by (simp add: lexext_compat_snoc_if_same_length) lemma lexext_rev_compat_snoc: "lexext_rev gt ys xs \ lexext_rev gt (ys @ [x]) (xs @ [x])" by (simp add: lexext_compat_cons) lemma lexext_rev_compat_list: "gt y x \ lexext_rev gt (xs @ y # xs') (xs @ x # xs')" by (induct xs' rule: rev_induct) auto lemma lexext_rev_singleton: "lexext_rev gt [y] [x] \ gt y x" by simp lemma lexext_rev_total: "(\y \ B. \x \ A. gt y x \ gt x y \ y = x) \ ys \ lists B \ xs \ lists A \ lexext_rev gt ys xs \ lexext_rev gt xs ys \ ys = xs" by (rule lexext_total[of _ _ _ "rev ys" "rev xs", simplified]) lemma lexext_rev_hd_or_tl: assumes "length ys = length xs" and "lexext_rev gt (y # ys) (x # xs)" shows "gt y x \ lexext_rev gt ys xs" using assms lexext_rev_cons_cons by fastforce interpretation lexext_rev: ext lexext_rev by standard (fact lexext_rev_mono_strong, rule lexext_rev_map_strong, metis in_listsD) interpretation lexext_rev: ext_irrefl_trans_strong lexext_rev by standard (fact lexext_rev_irrefl, fact lexext_rev_trans_strong) interpretation lexext_rev: ext_compat_snoc lexext_rev by standard (fact lexext_rev_compat_snoc) interpretation lexext_rev: ext_compat_list lexext_rev by standard (rule lexext_rev_compat_list) interpretation lexext_rev: ext_singleton lexext_rev by standard (rule lexext_rev_singleton) interpretation lexext_rev: ext_total lexext_rev by standard (fact lexext_rev_total) interpretation lexext_rev: ext_hd_or_tl lexext_rev by standard (rule lexext_rev_hd_or_tl) interpretation lexext_rev: ext_wf_bounded lexext_rev by standard subsection \Generic Length Extension\ definition lenext :: "('a list \ 'a list \ bool) \ 'a list \ 'a list \ bool" where "lenext gts ys xs \ length ys > length xs \ length ys = length xs \ gts ys xs" lemma lenext_mono_strong: "(gts ys xs \ gts' ys xs) \ lenext gts ys xs \ lenext gts' ys xs" and lenext_map_strong: "(length ys = length xs \ gts ys xs \ gts (map f ys) (map f xs)) \ lenext gts ys xs \ lenext gts (map f ys) (map f xs)" and lenext_irrefl: "\ gts xs xs \ \ lenext gts xs xs" and lenext_trans: "(gts zs ys \ gts ys xs \ gts zs xs) \ lenext gts zs ys \ lenext gts ys xs \ lenext gts zs xs" and lenext_snoc: "lenext gts (xs @ [x]) xs" and lenext_compat_cons: "(length ys = length xs \ gts ys xs \ gts (x # ys) (x # xs)) \ lenext gts ys xs \ lenext gts (x # ys) (x # xs)" and lenext_compat_snoc: "(length ys = length xs \ gts ys xs \ gts (ys @ [x]) (xs @ [x])) \ lenext gts ys xs \ lenext gts (ys @ [x]) (xs @ [x])" and lenext_compat_list: "gts (xs @ y # xs') (xs @ x # xs') \ lenext gts (xs @ y # xs') (xs @ x # xs')" and lenext_singleton: "lenext gts [y] [x] \ gts [y] [x]" and lenext_total: "(gts ys xs \ gts xs ys \ ys = xs) \ lenext gts ys xs \ lenext gts xs ys \ ys = xs" and lenext_hd_or_tl: "(length ys = length xs \ gts (y # ys) (x # xs) \ gt y x \ gts ys xs) \ lenext gts (y # ys) (x # xs) \ gt y x \ lenext gts ys xs" unfolding lenext_def by auto subsection \Length-Lexicographic Extension\ abbreviation len_lexext :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" where "len_lexext gt \ lenext (lexext gt)" lemma len_lexext_mono_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt' y x) \ len_lexext gt ys xs \ len_lexext gt' ys xs" by (rule lenext_mono_strong[OF lexext_mono_strong]) lemma len_lexext_map_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt (f y) (f x)) \ len_lexext gt ys xs \ len_lexext gt (map f ys) (map f xs)" by (rule lenext_map_strong) (metis lexext_map_strong) lemma len_lexext_irrefl: "(\x \ set xs. \ gt x x) \ \ len_lexext gt xs xs" by (rule lenext_irrefl[OF lexext_irrefl]) lemma len_lexext_trans_strong: "(\z \ set zs. \y \ set ys. \x \ set xs. gt z y \ gt y x \ gt z x) \ len_lexext gt zs ys \ len_lexext gt ys xs \ len_lexext gt zs xs" by (rule lenext_trans[OF lexext_trans_strong]) lemma len_lexext_snoc: "len_lexext gt (xs @ [x]) xs" by (rule lenext_snoc) lemma len_lexext_compat_cons: "len_lexext gt ys xs \ len_lexext gt (x # ys) (x # xs)" by (intro lenext_compat_cons lexext_compat_cons) lemma len_lexext_compat_snoc: "len_lexext gt ys xs \ len_lexext gt (ys @ [x]) (xs @ [x])" by (intro lenext_compat_snoc lexext_compat_snoc_if_same_length) lemma len_lexext_compat_list: "gt y x \ len_lexext gt (xs @ y # xs') (xs @ x # xs')" by (intro lenext_compat_list lexext_compat_list) lemma len_lexext_singleton[simp]: "len_lexext gt [y] [x] \ gt y x" by (simp only: lenext_singleton lexext_singleton) lemma len_lexext_total: "(\y \ B. \x \ A. gt y x \ gt x y \ y = x) \ ys \ lists B \ xs \ lists A \ len_lexext gt ys xs \ len_lexext gt xs ys \ ys = xs" by (rule lenext_total[OF lexext_total]) lemma len_lexext_iff_lenlex: "len_lexext gt ys xs \ (xs, ys) \ lenlex {(x, y). gt y x}" proof - { assume "length xs = length ys" hence "lexext gt ys xs \ (xs, ys) \ lex {(x, y). gt y x}" by (induct xs ys rule: list_induct2) auto } thus ?thesis unfolding lenext_def lenlex_conv by auto qed lemma len_lexext_wf: "wfP (\x y. gt y x) \ wfP (\xs ys. len_lexext gt ys xs)" unfolding wfP_def len_lexext_iff_lenlex by (simp add: wf_lenlex) lemma len_lexext_hd_or_tl: "len_lexext gt (y # ys) (x # xs) \ gt y x \ len_lexext gt ys xs" using lenext_hd_or_tl lexext_hd_or_tl by metis interpretation len_lexext: ext len_lexext by standard (fact len_lexext_mono_strong, rule len_lexext_map_strong, metis in_listsD) interpretation len_lexext: ext_irrefl_trans_strong len_lexext by standard (fact len_lexext_irrefl, fact len_lexext_trans_strong) interpretation len_lexext: ext_snoc len_lexext by standard (fact len_lexext_snoc) interpretation len_lexext: ext_compat_cons len_lexext by standard (fact len_lexext_compat_cons) interpretation len_lexext: ext_compat_snoc len_lexext by standard (fact len_lexext_compat_snoc) interpretation len_lexext: ext_compat_list len_lexext by standard (rule len_lexext_compat_list) interpretation len_lexext: ext_singleton len_lexext by standard (rule len_lexext_singleton) interpretation len_lexext: ext_total len_lexext by standard (fact len_lexext_total) interpretation len_lexext: ext_wf len_lexext by standard (fact len_lexext_wf) interpretation len_lexext: ext_hd_or_tl len_lexext by standard (rule len_lexext_hd_or_tl) interpretation len_lexext: ext_wf_bounded len_lexext by standard subsection \Reverse (Right-to-Left) Length-Lexicographic Extension\ abbreviation len_lexext_rev :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" where "len_lexext_rev gt \ lenext (lexext_rev gt)" lemma len_lexext_rev_mono_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt' y x) \ len_lexext_rev gt ys xs \ len_lexext_rev gt' ys xs" by (rule lenext_mono_strong) (rule lexext_rev_mono_strong) lemma len_lexext_rev_map_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt (f y) (f x)) \ len_lexext_rev gt ys xs \ len_lexext_rev gt (map f ys) (map f xs)" by (rule lenext_map_strong) (rule lexext_rev_map_strong) lemma len_lexext_rev_irrefl: "(\x \ set xs. \ gt x x) \ \ len_lexext_rev gt xs xs" by (rule lenext_irrefl) (rule lexext_rev_irrefl) lemma len_lexext_rev_trans_strong: "(\z \ set zs. \y \ set ys. \x \ set xs. gt z y \ gt y x \ gt z x) \ len_lexext_rev gt zs ys \ len_lexext_rev gt ys xs \ len_lexext_rev gt zs xs" by (rule lenext_trans) (rule lexext_rev_trans_strong) lemma len_lexext_rev_snoc: "len_lexext_rev gt (xs @ [x]) xs" by (rule lenext_snoc) lemma len_lexext_rev_compat_cons: "len_lexext_rev gt ys xs \ len_lexext_rev gt (x # ys) (x # xs)" by (intro lenext_compat_cons lexext_rev_compat_cons_if_same_length) lemma len_lexext_rev_compat_snoc: "len_lexext_rev gt ys xs \ len_lexext_rev gt (ys @ [x]) (xs @ [x])" by (intro lenext_compat_snoc lexext_rev_compat_snoc) lemma len_lexext_rev_compat_list: "gt y x \ len_lexext_rev gt (xs @ y # xs') (xs @ x # xs')" by (intro lenext_compat_list lexext_rev_compat_list) lemma len_lexext_rev_singleton[simp]: "len_lexext_rev gt [y] [x] \ gt y x" by (simp only: lenext_singleton lexext_rev_singleton) lemma len_lexext_rev_total: "(\y \ B. \x \ A. gt y x \ gt x y \ y = x) \ ys \ lists B \ xs \ lists A \ len_lexext_rev gt ys xs \ len_lexext_rev gt xs ys \ ys = xs" by (rule lenext_total[OF lexext_rev_total]) lemma len_lexext_rev_iff_len_lexext: "len_lexext_rev gt ys xs \ len_lexext gt (rev ys) (rev xs)" unfolding lenext_def by simp lemma len_lexext_rev_wf: "wfP (\x y. gt y x) \ wfP (\xs ys. len_lexext_rev gt ys xs)" unfolding len_lexext_rev_iff_len_lexext by (rule wfP_app[of "\xs ys. len_lexext gt ys xs" rev, simplified]) (rule len_lexext_wf) lemma len_lexext_rev_hd_or_tl: "len_lexext_rev gt (y # ys) (x # xs) \ gt y x \ len_lexext_rev gt ys xs" using lenext_hd_or_tl lexext_rev_hd_or_tl by metis interpretation len_lexext_rev: ext len_lexext_rev by standard (fact len_lexext_rev_mono_strong, rule len_lexext_rev_map_strong, metis in_listsD) interpretation len_lexext_rev: ext_irrefl_trans_strong len_lexext_rev by standard (fact len_lexext_rev_irrefl, fact len_lexext_rev_trans_strong) interpretation len_lexext_rev: ext_snoc len_lexext_rev by standard (fact len_lexext_rev_snoc) interpretation len_lexext_rev: ext_compat_cons len_lexext_rev by standard (fact len_lexext_rev_compat_cons) interpretation len_lexext_rev: ext_compat_snoc len_lexext_rev by standard (fact len_lexext_rev_compat_snoc) interpretation len_lexext_rev: ext_compat_list len_lexext_rev by standard (rule len_lexext_rev_compat_list) interpretation len_lexext_rev: ext_singleton len_lexext_rev by standard (rule len_lexext_rev_singleton) interpretation len_lexext_rev: ext_total len_lexext_rev by standard (fact len_lexext_rev_total) interpretation len_lexext_rev: ext_wf len_lexext_rev by standard (fact len_lexext_rev_wf) interpretation len_lexext_rev: ext_hd_or_tl len_lexext_rev by standard (rule len_lexext_rev_hd_or_tl) interpretation len_lexext_rev: ext_wf_bounded len_lexext_rev by standard subsection \Dershowitz--Manna Multiset Extension\ definition msetext_dersh where "msetext_dersh gt ys xs = (let N = mset ys; M = mset xs in (\Y X. Y \ {#} \ Y \# N \ M = (N - Y) + X \ (\x. x \# X \ (\y. y \# Y \ gt y x))))" text \ The following proof is based on that of @{thm[source] less_multiset\<^sub>D\<^sub>M_imp_mult}. \ lemma msetext_dersh_imp_mult_rel: assumes ys_a: "ys \ lists A" and xs_a: "xs \ lists A" and ys_gt_xs: "msetext_dersh gt ys xs" shows "(mset xs, mset ys) \ mult {(x, y). x \ A \ y \ A \ gt y x}" proof - obtain Y X where y_nemp: "Y \ {#}" and y_sub_ys: "Y \# mset ys" and xs_eq: "mset xs = mset ys - Y + X" and ex_y: "\x. x \# X \ (\y. y \# Y \ gt y x)" using ys_gt_xs[unfolded msetext_dersh_def Let_def] by blast have ex_y': "\x. x \# X \ (\y. y \# Y \ x \ A \ y \ A \ gt y x)" using ex_y y_sub_ys xs_eq ys_a xs_a by (metis in_listsD mset_subset_eqD set_mset_mset union_iff) hence "(mset ys - Y + X, mset ys - Y + Y) \ mult {(x, y). x \ A \ y \ A \ gt y x}" using y_nemp y_sub_ys by (intro one_step_implies_mult) (auto simp: Bex_def trans_def) thus ?thesis using xs_eq y_sub_ys by (simp add: subset_mset.diff_add) qed lemma msetext_dersh_imp_mult: "msetext_dersh gt ys xs \ (mset xs, mset ys) \ mult {(x, y). gt y x}" using msetext_dersh_imp_mult_rel[of _ UNIV] by auto lemma mult_imp_msetext_dersh_rel: assumes ys_a: "set_mset (mset ys) \ A" and xs_a: "set_mset (mset xs) \ A" and in_mult: "(mset xs, mset ys) \ mult {(x, y). x \ A \ y \ A \ gt y x}" and trans: "\z \ A. \y \ A. \x \ A. gt z y \ gt y x \ gt z x" shows "msetext_dersh gt ys xs" using in_mult ys_a xs_a unfolding mult_def msetext_dersh_def Let_def proof induct case (base Ys) then obtain y M0 X where "Ys = M0 + {#y#}" and "mset xs = M0 + X" and "\a. a \# X \ gt y a" unfolding mult1_def by auto thus ?case by (auto intro: exI[of _ "{#y#}"] exI[of _ X]) next case (step Ys Zs) note ys_zs_in_mult1 = this(2) and ih = this(3) and zs_a = this(4) and xs_a = this(5) have Ys_a: "set_mset Ys \ A" using ys_zs_in_mult1 zs_a unfolding mult1_def by auto obtain Y X where y_nemp: "Y \ {#}" and y_sub_ys: "Y \# Ys" and xs_eq: "mset xs = Ys - Y + X" and ex_y: "\x. x \# X \ (\y. y \# Y \ gt y x)" using ih[OF Ys_a xs_a] by blast obtain z M0 Ya where zs_eq: "Zs = M0 + {#z#}" and ys_eq: "Ys = M0 + Ya" and z_gt: "\y. y \# Ya \ y \ A \ z \ A \ gt z y" using ys_zs_in_mult1[unfolded mult1_def] by auto let ?Za = "Y - Ya + {#z#}" let ?Xa = "X + Ya + (Y - Ya) - Y" have xa_sub_x_ya: "set_mset ?Xa \ set_mset (X + Ya)" by (metis diff_subset_eq_self in_diffD subsetI subset_mset.diff_diff_right) have x_a: "set_mset X \ A" using xs_a xs_eq by auto have ya_a: "set_mset Ya \ A" by (simp add: subsetI z_gt) have ex_y': "\y. y \# Y - Ya + {#z#} \ gt y x" if x_in: "x \# X + Ya" for x proof (cases "x \# X") case True then obtain y where y_in: "y \# Y" and y_gt_x: "gt y x" using ex_y by blast show ?thesis proof (cases "y \# Ya") case False hence "y \# Y - Ya + {#z#}" using y_in by fastforce thus ?thesis using y_gt_x by blast next case True hence "y \ A" and "z \ A" and "gt z y" using z_gt by blast+ hence "gt z x" using trans y_gt_x x_a ya_a x_in by (meson subsetCE union_iff) thus ?thesis by auto qed next case False hence "x \# Ya" using x_in by auto hence "x \ A" and "z \ A" and "gt z x" using z_gt by blast+ thus ?thesis by auto qed show ?case proof (rule exI[of _ ?Za], rule exI[of _ ?Xa], intro conjI) show "Y - Ya + {#z#} \# Zs" using mset_subset_eq_mono_add subset_eq_diff_conv y_sub_ys ys_eq zs_eq by fastforce next show "mset xs = Zs - (Y - Ya + {#z#}) + (X + Ya + (Y - Ya) - Y)" unfolding xs_eq ys_eq zs_eq by (auto simp: multiset_eq_iff) next show "\x. x \# X + Ya + (Y - Ya) - Y \ (\y. y \# Y - Ya + {#z#} \ gt y x)" using ex_y' xa_sub_x_ya by blast qed auto qed lemma msetext_dersh_mono_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt' y x) \ msetext_dersh gt ys xs \ msetext_dersh gt' ys xs" unfolding msetext_dersh_def Let_def by (metis mset_subset_eqD mset_subset_eq_add_right set_mset_mset) lemma msetext_dersh_map_strong: assumes compat_f: "\y \ set ys. \x \ set xs. gt y x \ gt (f y) (f x)" and ys_gt_xs: "msetext_dersh gt ys xs" shows "msetext_dersh gt (map f ys) (map f xs)" proof - obtain Y X where y_nemp: "Y \ {#}" and y_sub_ys: "Y \# mset ys" and xs_eq: "mset xs = mset ys - Y + X" and ex_y: "\x. x \# X \ (\y. y \# Y \ gt y x)" using ys_gt_xs[unfolded msetext_dersh_def Let_def mset_map] by blast have x_sub_xs: "X \# mset xs" using xs_eq by simp let ?fY = "image_mset f Y" let ?fX = "image_mset f X" show ?thesis unfolding msetext_dersh_def Let_def mset_map proof (intro exI conjI) show "image_mset f (mset xs) = image_mset f (mset ys) - ?fY + ?fX" using xs_eq[THEN arg_cong, of "image_mset f"] y_sub_ys by (metis image_mset_Diff image_mset_union) next obtain y where y: "\x. x \# X \ y x \# Y \ gt (y x) x" using ex_y by moura show "\fx. fx \# ?fX \ (\fy. fy \# ?fY \ gt fy fx)" proof (intro allI impI) fix fx assume "fx \# ?fX" then obtain x where fx: "fx = f x" and x_in: "x \# X" by auto hence y_in: "y x \# Y" and y_gt: "gt (y x) x" using y[rule_format, OF x_in] by blast+ hence "f (y x) \# ?fY \ gt (f (y x)) (f x)" using compat_f y_sub_ys x_sub_xs x_in by (metis image_eqI in_image_mset mset_subset_eqD set_mset_mset) thus "\fy. fy \# ?fY \ gt fy fx" unfolding fx by auto qed qed (auto simp: y_nemp y_sub_ys image_mset_subseteq_mono) qed lemma msetext_dersh_trans: assumes zs_a: "zs \ lists A" and ys_a: "ys \ lists A" and xs_a: "xs \ lists A" and trans: "\z \ A. \y \ A. \x \ A. gt z y \ gt y x \ gt z x" and zs_gt_ys: "msetext_dersh gt zs ys" and ys_gt_xs: "msetext_dersh gt ys xs" shows "msetext_dersh gt zs xs" proof (rule mult_imp_msetext_dersh_rel[OF _ _ _ trans]) show "set_mset (mset zs) \ A" using zs_a by auto next show "set_mset (mset xs) \ A" using xs_a by auto next let ?Gt = "{(x, y). x \ A \ y \ A \ gt y x}" have "(mset xs, mset ys) \ mult ?Gt" by (rule msetext_dersh_imp_mult_rel[OF ys_a xs_a ys_gt_xs]) moreover have "(mset ys, mset zs) \ mult ?Gt" by (rule msetext_dersh_imp_mult_rel[OF zs_a ys_a zs_gt_ys]) ultimately show "(mset xs, mset zs) \ mult ?Gt" unfolding mult_def by simp qed lemma msetext_dersh_irrefl_from_trans: assumes trans: "\z \ set xs. \y \ set xs. \x \ set xs. gt z y \ gt y x \ gt z x" and irrefl: "\x \ set xs. \ gt x x" shows "\ msetext_dersh gt xs xs" unfolding msetext_dersh_def Let_def proof clarify fix Y X assume y_nemp: "Y \ {#}" and y_sub_xs: "Y \# mset xs" and xs_eq: "mset xs = mset xs - Y + X" and ex_y: "\x. x \# X \ (\y. y \# Y \ gt y x)" have x_eq_y: "X = Y" using y_sub_xs xs_eq by (metis diff_union_cancelL subset_mset.diff_add) let ?Gt = "{(y, x). y \# Y \ x \# Y \ gt y x}" have "?Gt \ set_mset Y \ set_mset Y" by auto hence fin: "finite ?Gt" by (auto dest!: infinite_super) moreover have "irrefl ?Gt" unfolding irrefl_def using irrefl y_sub_xs by (fastforce dest!: set_mset_mono) moreover have "trans ?Gt" unfolding trans_def using trans y_sub_xs by (fastforce dest!: set_mset_mono) ultimately have acyc: "acyclic ?Gt" by (rule finite_irrefl_trans_imp_wf[THEN wf_acyclic]) have fin_y: "finite (set_mset Y)" using y_sub_xs by simp hence cyc: "\ acyclic ?Gt" proof (rule finite_nonempty_ex_succ_imp_cyclic) show "\x \# Y. \y \# Y. (y, x) \ ?Gt" using ex_y[unfolded x_eq_y] by auto qed (auto simp: y_nemp) show False using acyc cyc by sat qed lemma msetext_dersh_snoc: "msetext_dersh gt (xs @ [x]) xs" unfolding msetext_dersh_def Let_def proof (intro exI conjI) show "mset xs = mset (xs @ [x]) - {#x#} + {#}" by simp qed auto lemma msetext_dersh_compat_cons: assumes ys_gt_xs: "msetext_dersh gt ys xs" shows "msetext_dersh gt (x # ys) (x # xs)" proof - obtain Y X where y_nemp: "Y \ {#}" and y_sub_ys: "Y \# mset ys" and xs_eq: "mset xs = mset ys - Y + X" and ex_y: "\x. x \# X \ (\y. y \# Y \ gt y x)" using ys_gt_xs[unfolded msetext_dersh_def Let_def mset_map] by blast show ?thesis unfolding msetext_dersh_def Let_def proof (intro exI conjI) show "Y \# mset (x # ys)" using y_sub_ys by (metis add_mset_add_single mset.simps(2) mset_subset_eq_add_left subset_mset.add_increasing2) next show "mset (x # xs) = mset (x # ys) - Y + X" proof - have "X + (mset ys - Y) = mset xs" by (simp add: union_commute xs_eq) hence "mset (x # xs) = X + (mset (x # ys) - Y)" by (metis add_mset_add_single mset.simps(2) mset_subset_eq_multiset_union_diff_commute union_mset_add_mset_right y_sub_ys) thus ?thesis by (simp add: union_commute) qed qed (auto simp: y_nemp ex_y) qed lemma msetext_dersh_compat_snoc: "msetext_dersh gt ys xs \ msetext_dersh gt (ys @ [x]) (xs @ [x])" using msetext_dersh_compat_cons[of gt ys xs x] unfolding msetext_dersh_def by simp lemma msetext_dersh_compat_list: assumes y_gt_x: "gt y x" shows "msetext_dersh gt (xs @ y # xs') (xs @ x # xs')" unfolding msetext_dersh_def Let_def proof (intro exI conjI) show "mset (xs @ x # xs') = mset (xs @ y # xs') - {#y#} + {#x#}" by auto qed (auto intro: y_gt_x) lemma msetext_dersh_singleton: "msetext_dersh gt [y] [x] \ gt y x" unfolding msetext_dersh_def Let_def by (auto dest: nonempty_subseteq_mset_eq_singleton simp: nonempty_subseteq_mset_iff_singleton) lemma msetext_dersh_wf: assumes wf_gt: "wfP (\x y. gt y x)" shows "wfP (\xs ys. msetext_dersh gt ys xs)" proof (rule wfP_subset, rule wfP_app[of "\xs ys. (xs, ys) \ mult {(x, y). gt y x}" mset]) show "wfP (\xs ys. (xs, ys) \ mult {(x, y). gt y x})" using wf_gt unfolding wfP_def by (auto intro: wf_mult) next show "(\xs ys. msetext_dersh gt ys xs) \ (\x y. (mset x, mset y) \ mult {(x, y). gt y x})" using msetext_dersh_imp_mult by blast qed interpretation msetext_dersh: ext msetext_dersh by standard (fact msetext_dersh_mono_strong, rule msetext_dersh_map_strong, metis in_listsD) interpretation msetext_dersh: ext_trans_before_irrefl msetext_dersh by standard (fact msetext_dersh_trans, fact msetext_dersh_irrefl_from_trans) interpretation msetext_dersh: ext_snoc msetext_dersh by standard (fact msetext_dersh_snoc) interpretation msetext_dersh: ext_compat_cons msetext_dersh by standard (fact msetext_dersh_compat_cons) interpretation msetext_dersh: ext_compat_snoc msetext_dersh by standard (fact msetext_dersh_compat_snoc) interpretation msetext_dersh: ext_compat_list msetext_dersh by standard (rule msetext_dersh_compat_list) interpretation msetext_dersh: ext_singleton msetext_dersh by standard (rule msetext_dersh_singleton) interpretation msetext_dersh: ext_wf msetext_dersh by standard (fact msetext_dersh_wf) subsection \Huet--Oppen Multiset Extension\ definition msetext_huet where "msetext_huet gt ys xs = (let N = mset ys; M = mset xs in M \ N \ (\x. count M x > count N x \ (\y. gt y x \ count N y > count M y)))" lemma msetext_huet_imp_count_gt: assumes ys_gt_xs: "msetext_huet gt ys xs" shows "\x. count (mset ys) x > count (mset xs) x" proof - obtain x where "count (mset ys) x \ count (mset xs) x" using ys_gt_xs[unfolded msetext_huet_def Let_def] by (fastforce intro: multiset_eqI) moreover { assume "count (mset ys) x < count (mset xs) x" hence ?thesis using ys_gt_xs[unfolded msetext_huet_def Let_def] by blast } moreover { assume "count (mset ys) x > count (mset xs) x" hence ?thesis by fast } ultimately show ?thesis by fastforce qed lemma msetext_huet_imp_dersh: assumes huet: "msetext_huet gt ys xs" shows "msetext_dersh gt ys xs" proof (unfold msetext_dersh_def Let_def, intro exI conjI) let ?X = "mset xs - mset ys" let ?Y = "mset ys - mset xs" show "?Y \ {#}" by (metis msetext_huet_imp_count_gt[OF huet] empty_iff in_diff_count set_mset_empty) show "?Y \# mset ys" by auto show "mset xs = mset ys - ?Y + ?X" by (metis add.commute diff_intersect_right_idem multiset_inter_def subset_mset.inf.cobounded2 subset_mset.le_imp_diff_is_add) show "\x. x \# ?X \ (\y. y \# ?Y \ gt y x)" using huet[unfolded msetext_huet_def Let_def, THEN conjunct2] by (meson in_diff_count) qed text \ The following proof is based on that of @{thm[source] mult_imp_less_multiset\<^sub>H\<^sub>O}. \ lemma mult_imp_msetext_huet: assumes irrefl: "irreflp gt" and trans: "transp gt" and in_mult: "(mset xs, mset ys) \ mult {(x, y). gt y x}" shows "msetext_huet gt ys xs" using in_mult unfolding mult_def msetext_huet_def Let_def proof (induct rule: trancl_induct) case (base Ys) thus ?case using irrefl unfolding irreflp_def msetext_huet_def Let_def mult1_def by (auto 0 3 split: if_splits) next case (step Ys Zs) have asym[unfolded antisym_def, simplified]: "antisymp gt" by (rule irreflp_transp_imp_antisymP[OF irrefl trans]) from step(3) have "mset xs \ Ys" and **: "\x. count Ys x < count (mset xs) x \ (\y. gt y x \ count (mset xs) y < count Ys y)" by blast+ from step(2) obtain M0 a K where *: "Zs = M0 + {#a#}" "Ys = M0 + K" "a \# K" "\b. b \# K \ gt a b" using irrefl unfolding mult1_def irreflp_def by force have "mset xs \ Zs" proof (cases "K = {#}") case True thus ?thesis using \mset xs \ Ys\ ** *(1,2) irrefl[unfolded irreflp_def] by (metis One_nat_def add.comm_neutral count_single diff_union_cancelL lessI minus_multiset.rep_eq not_add_less2 plus_multiset.rep_eq union_commute zero_less_diff) next case False thus ?thesis proof - obtain aa :: "'a \ 'a" where f1: "\a. \ count Ys a < count (mset xs) a \ gt (aa a) a \ count (mset xs) (aa a) < count Ys (aa a)" using "**" by moura have f2: "K + M0 = Ys" using "*"(2) union_ac(2) by blast have f3: "\aa. count Zs aa = count M0 aa + count {#a#} aa" by (simp add: "*"(1)) have f4: "\a. count Ys a = count K a + count M0 a" using f2 by auto have f5: "count K a = 0" by (meson "*"(3) count_inI) have "Zs - M0 = {#a#}" using "*"(1) add_diff_cancel_left' by blast then have f6: "count M0 a < count Zs a" by (metis in_diff_count union_single_eq_member) have "\m. count m a = 0 + count m a" by simp moreover { assume "aa a \ a" then have "mset xs = Zs \ count Zs (aa a) < count K (aa a) + count M0 (aa a) \ count K (aa a) + count M0 (aa a) < count Zs (aa a)" using f5 f3 f2 f1 "*"(4) asym by (auto dest!: antisympD) } ultimately show ?thesis using f6 f5 f4 f1 by (metis less_imp_not_less) qed qed moreover { assume "count Zs a \ count (mset xs) a" with \a \# K\ have "count Ys a < count (mset xs) a" unfolding *(1,2) by (auto simp add: not_in_iff) with ** obtain z where z: "gt z a" "count (mset xs) z < count Ys z" by blast with * have "count Ys z \ count Zs z" using asym by (auto simp: intro: count_inI dest: antisympD) with z have "\z. gt z a \ count (mset xs) z < count Zs z" by auto } note count_a = this { fix y assume count_y: "count Zs y < count (mset xs) y" have "\x. gt x y \ count (mset xs) x < count Zs x" proof (cases "y = a") case True with count_y count_a show ?thesis by auto next case False show ?thesis proof (cases "y \# K") case True with *(4) have "gt a y" by simp then show ?thesis by (cases "count Zs a \ count (mset xs) a", blast dest: count_a trans[unfolded transp_def, rule_format], auto dest: count_a) next case False with \y \ a\ have "count Zs y = count Ys y" unfolding *(1,2) by (simp add: not_in_iff) with count_y ** obtain z where z: "gt z y" "count (mset xs) z < count Ys z" by auto show ?thesis proof (cases "z \# K") case True with *(4) have "gt a z" by simp with z(1) show ?thesis by (cases "count Zs a \ count (mset xs) a") (blast dest: count_a not_le_imp_less trans[unfolded transp_def, rule_format])+ next case False with \a \# K\ have "count Ys z \ count Zs z" unfolding * by (auto simp add: not_in_iff) with z show ?thesis by auto qed qed qed } ultimately show ?case unfolding msetext_huet_def Let_def by blast qed theorem msetext_huet_eq_dersh: "irreflp gt \ transp gt \ msetext_dersh gt = msetext_huet gt" using msetext_huet_imp_dersh msetext_dersh_imp_mult mult_imp_msetext_huet by fast lemma msetext_huet_mono_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt' y x) \ msetext_huet gt ys xs \ msetext_huet gt' ys xs" unfolding msetext_huet_def by (metis less_le_trans mem_Collect_eq not_le not_less0 set_mset_mset[unfolded set_mset_def]) lemma msetext_huet_map: assumes fin: "finite A" and ys_a: "ys \ lists A" and xs_a: "xs \ lists A" and irrefl_f: "\x \ A. \ gt (f x) (f x)" and trans_f: "\z \ A. \y \ A. \x \ A. gt (f z) (f y) \ gt (f y) (f x) \ gt (f z) (f x)" and compat_f: "\y \ A. \x \ A. gt y x \ gt (f y) (f x)" and ys_gt_xs: "msetext_huet gt ys xs" shows "msetext_huet gt (map f ys) (map f xs)" (is "msetext_huet _ ?fys ?fxs") proof - have irrefl: "\x \ A. \ gt x x" using irrefl_f compat_f by blast have ms_xs_ne_ys: "mset xs \ mset ys" and ex_gt: "\x. count (mset ys) x < count (mset xs) x \ (\y. gt y x \ count (mset xs) y < count (mset ys) y)" using ys_gt_xs[unfolded msetext_huet_def Let_def] by blast+ have ex_y: "\y. gt (f y) (f x) \ count (mset ?fxs) (f y) < count (mset (map f ys)) (f y)" if cnt_x: "count (mset xs) x > count (mset ys) x" for x proof - have x_in_a: "x \ A" using cnt_x xs_a dual_order.strict_trans2 by fastforce obtain y where y_gt_x: "gt y x" and cnt_y: "count (mset ys) y > count (mset xs) y" using cnt_x ex_gt by blast have y_in_a: "y \ A" using cnt_y ys_a dual_order.strict_trans2 by fastforce have wf_gt_f: "wfP (\y x. y \ A \ x \ A \ gt (f y) (f x))" by (rule finite_irreflp_transp_imp_wfp) (auto elim: trans_f[rule_format] simp: fin irrefl_f Collect_case_prod_Sigma irreflp_def transp_def) obtain yy where fyy_gt_fx: "gt (f yy) (f x)" and cnt_yy: "count (mset ys) yy > count (mset xs) yy" and max_yy: "\y \ A. yy \ A \ gt (f y) (f yy) \ gt (f y) (f x) \ count (mset xs) y \ count (mset ys) y" using wfP_eq_minimal[THEN iffD1, OF wf_gt_f, rule_format, of y "{y. gt (f y) (f x) \ count (mset xs) y < count (mset ys) y}", simplified] y_gt_x cnt_y by (metis compat_f not_less x_in_a y_in_a) have yy_in_a: "yy \ A" using cnt_yy ys_a dual_order.strict_trans2 by fastforce { assume "count (mset ?fxs) (f yy) \ count (mset ?fys) (f yy)" then obtain u where fu_eq_fyy: "f u = f yy" and cnt_u: "count (mset xs) u > count (mset ys) u" using count_image_mset_le_imp_lt cnt_yy mset_map by (metis (mono_tags)) have u_in_a: "u \ A" using cnt_u xs_a dual_order.strict_trans2 by fastforce obtain v where v_gt_u: "gt v u" and cnt_v: "count (mset ys) v > count (mset xs) v" using cnt_u ex_gt by blast have v_in_a: "v \ A" using cnt_v ys_a dual_order.strict_trans2 by fastforce have fv_gt_fu: "gt (f v) (f u)" using v_gt_u compat_f v_in_a u_in_a by blast hence fv_gt_fyy: "gt (f v) (f yy)" by (simp only: fu_eq_fyy) have "gt (f v) (f x)" using fv_gt_fyy fyy_gt_fx v_in_a yy_in_a x_in_a trans_f by blast hence False using max_yy[rule_format, of v] fv_gt_fyy v_in_a yy_in_a cnt_v by linarith } thus ?thesis using fyy_gt_fx leI by blast qed show ?thesis unfolding msetext_huet_def Let_def proof (intro conjI allI impI) { assume len_eq: "length xs = length ys" obtain x where cnt_x: "count (mset xs) x > count (mset ys) x" using len_eq ms_xs_ne_ys by (metis size_eq_ex_count_lt size_mset) hence "mset ?fxs \ mset ?fys" using ex_y by fastforce } thus "mset ?fxs \ mset (map f ys)" by (metis length_map size_mset) next fix fx assume cnt_fx: "count (mset ?fxs) fx > count (mset ?fys) fx" then obtain x where fx: "fx = f x" and cnt_x: "count (mset xs) x > count (mset ys) x" using count_image_mset_lt_imp_lt mset_map by (metis (mono_tags)) thus "\fy. gt fy fx \ count (mset ?fxs) fy < count (mset (map f ys)) fy" using ex_y[OF cnt_x] by blast qed qed lemma msetext_huet_irrefl: "(\x \ set xs. \ gt x x) \ \ msetext_huet gt xs xs" unfolding msetext_huet_def by simp lemma msetext_huet_trans_from_irrefl: assumes fin: "finite A" and zs_a: "zs \ lists A" and ys_a: "ys \ lists A" and xs_a: "xs \ lists A" and irrefl: "\x \ A. \ gt x x" and trans: "\z \ A. \y \ A. \x \ A. gt z y \ gt y x \ gt z x" and zs_gt_ys: "msetext_huet gt zs ys" and ys_gt_xs: "msetext_huet gt ys xs" shows "msetext_huet gt zs xs" proof - have wf_gt: "wfP (\y x. y \ A \ x \ A \ gt y x)" by (rule finite_irreflp_transp_imp_wfp) (auto elim: trans[rule_format] simp: fin irrefl Collect_case_prod_Sigma irreflp_def transp_def) show ?thesis unfolding msetext_huet_def Let_def proof (intro conjI allI impI) obtain x where cnt_x: "count (mset zs) x > count (mset ys) x" using msetext_huet_imp_count_gt[OF zs_gt_ys] by blast have x_in_a: "x \ A" using cnt_x zs_a dual_order.strict_trans2 by fastforce obtain xx where cnt_xx: "count (mset zs) xx > count (mset ys) xx" and max_xx: "\y \ A. xx \ A \ gt y xx \ count (mset ys) y \ count (mset zs) y" using wfP_eq_minimal[THEN iffD1, OF wf_gt, rule_format, of x "{y. count (mset ys) y < count (mset zs) y}", simplified] cnt_x by force have xx_in_a: "xx \ A" using cnt_xx zs_a dual_order.strict_trans2 by fastforce show "mset xs \ mset zs" proof (cases "count (mset ys) xx \ count (mset xs) xx") case True thus ?thesis using cnt_xx by fastforce next case False hence "count (mset ys) xx < count (mset xs) xx" by fastforce then obtain z where z_gt_xx: "gt z xx" and cnt_z: "count (mset ys) z > count (mset xs) z" using ys_gt_xs[unfolded msetext_huet_def Let_def] by blast have z_in_a: "z \ A" using cnt_z ys_a dual_order.strict_trans2 by fastforce have "count (mset zs) z \ count (mset ys) z" using max_xx[rule_format, of z] z_in_a xx_in_a z_gt_xx by blast moreover { assume "count (mset zs) z < count (mset ys) z" then obtain u where u_gt_z: "gt u z" and cnt_u: "count (mset ys) u < count (mset zs) u" using zs_gt_ys[unfolded msetext_huet_def Let_def] by blast have u_in_a: "u \ A" using cnt_u zs_a dual_order.strict_trans2 by fastforce have u_gt_xx: "gt u xx" using trans u_in_a z_in_a xx_in_a u_gt_z z_gt_xx by blast have False using max_xx[rule_format, of u] u_in_a xx_in_a u_gt_xx cnt_u by fastforce } ultimately have "count (mset zs) z = count (mset ys) z" by fastforce thus ?thesis using cnt_z by fastforce qed next fix x assume cnt_x_xz: "count (mset zs) x < count (mset xs) x" have x_in_a: "x \ A" using cnt_x_xz xs_a dual_order.strict_trans2 by fastforce let ?case = "\y. gt y x \ count (mset zs) y > count (mset xs) y" { assume cnt_x: "count (mset zs) x < count (mset ys) x" then obtain y where y_gt_x: "gt y x" and cnt_y: "count (mset zs) y > count (mset ys) y" using zs_gt_ys[unfolded msetext_huet_def Let_def] by blast have y_in_a: "y \ A" using cnt_y zs_a dual_order.strict_trans2 by fastforce obtain yy where yy_gt_x: "gt yy x" and cnt_yy: "count (mset zs) yy > count (mset ys) yy" and max_yy: "\y \ A. yy \ A \ gt y yy \ gt y x \ count (mset ys) y \ count (mset zs) y" using wfP_eq_minimal[THEN iffD1, OF wf_gt, rule_format, of y "{y. gt y x \ count (mset ys) y < count (mset zs) y}", simplified] y_gt_x cnt_y by force have yy_in_a: "yy \ A" using cnt_yy zs_a dual_order.strict_trans2 by fastforce have ?case proof (cases "count (mset ys) yy \ count (mset xs) yy") case True thus ?thesis using yy_gt_x cnt_yy by fastforce next case False hence "count (mset ys) yy < count (mset xs) yy" by fastforce then obtain z where z_gt_yy: "gt z yy" and cnt_z: "count (mset ys) z > count (mset xs) z" using ys_gt_xs[unfolded msetext_huet_def Let_def] by blast have z_in_a: "z \ A" using cnt_z ys_a dual_order.strict_trans2 by fastforce have z_gt_x: "gt z x" using trans z_in_a yy_in_a x_in_a z_gt_yy yy_gt_x by blast have "count (mset zs) z \ count (mset ys) z" using max_yy[rule_format, of z] z_in_a yy_in_a z_gt_yy z_gt_x by blast moreover { assume "count (mset zs) z < count (mset ys) z" then obtain u where u_gt_z: "gt u z" and cnt_u: "count (mset ys) u < count (mset zs) u" using zs_gt_ys[unfolded msetext_huet_def Let_def] by blast have u_in_a: "u \ A" using cnt_u zs_a dual_order.strict_trans2 by fastforce have u_gt_yy: "gt u yy" using trans u_in_a z_in_a yy_in_a u_gt_z z_gt_yy by blast have u_gt_x: "gt u x" using trans u_in_a z_in_a x_in_a u_gt_z z_gt_x by blast have False using max_yy[rule_format, of u] u_in_a yy_in_a u_gt_yy u_gt_x cnt_u by fastforce } ultimately have "count (mset zs) z = count (mset ys) z" by fastforce thus ?thesis using z_gt_x cnt_z by fastforce qed } moreover { assume "count (mset zs) x \ count (mset ys) x" hence "count (mset ys) x < count (mset xs) x" using cnt_x_xz by fastforce then obtain y where y_gt_x: "gt y x" and cnt_y: "count (mset ys) y > count (mset xs) y" using ys_gt_xs[unfolded msetext_huet_def Let_def] by blast have y_in_a: "y \ A" using cnt_y ys_a dual_order.strict_trans2 by fastforce obtain yy where yy_gt_x: "gt yy x" and cnt_yy: "count (mset ys) yy > count (mset xs) yy" and max_yy: "\y \ A. yy \ A \ gt y yy \ gt y x \ count (mset xs) y \ count (mset ys) y" using wfP_eq_minimal[THEN iffD1, OF wf_gt, rule_format, of y "{y. gt y x \ count (mset xs) y < count (mset ys) y}", simplified] y_gt_x cnt_y by force have yy_in_a: "yy \ A" using cnt_yy ys_a dual_order.strict_trans2 by fastforce have ?case proof (cases "count (mset zs) yy \ count (mset ys) yy") case True thus ?thesis using yy_gt_x cnt_yy by fastforce next case False hence "count (mset zs) yy < count (mset ys) yy" by fastforce then obtain z where z_gt_yy: "gt z yy" and cnt_z: "count (mset zs) z > count (mset ys) z" using zs_gt_ys[unfolded msetext_huet_def Let_def] by blast have z_in_a: "z \ A" using cnt_z zs_a dual_order.strict_trans2 by fastforce have z_gt_x: "gt z x" using trans z_in_a yy_in_a x_in_a z_gt_yy yy_gt_x by blast have "count (mset ys) z \ count (mset xs) z" using max_yy[rule_format, of z] z_in_a yy_in_a z_gt_yy z_gt_x by blast moreover { assume "count (mset ys) z < count (mset xs) z" then obtain u where u_gt_z: "gt u z" and cnt_u: "count (mset xs) u < count (mset ys) u" using ys_gt_xs[unfolded msetext_huet_def Let_def] by blast have u_in_a: "u \ A" using cnt_u ys_a dual_order.strict_trans2 by fastforce have u_gt_yy: "gt u yy" using trans u_in_a z_in_a yy_in_a u_gt_z z_gt_yy by blast have u_gt_x: "gt u x" using trans u_in_a z_in_a x_in_a u_gt_z z_gt_x by blast have False using max_yy[rule_format, of u] u_in_a yy_in_a u_gt_yy u_gt_x cnt_u by fastforce } ultimately have "count (mset ys) z = count (mset xs) z" by fastforce thus ?thesis using z_gt_x cnt_z by fastforce qed } ultimately show "\y. gt y x \ count (mset xs) y < count (mset zs) y" by fastforce qed qed lemma msetext_huet_snoc: "msetext_huet gt (xs @ [x]) xs" unfolding msetext_huet_def Let_def by simp lemma msetext_huet_compat_cons: "msetext_huet gt ys xs \ msetext_huet gt (x # ys) (x # xs)" unfolding msetext_huet_def Let_def by auto lemma msetext_huet_compat_snoc: "msetext_huet gt ys xs \ msetext_huet gt (ys @ [x]) (xs @ [x])" unfolding msetext_huet_def Let_def by auto lemma msetext_huet_compat_list: "y \ x \ gt y x \ msetext_huet gt (xs @ y # xs') (xs @ x # xs')" unfolding msetext_huet_def Let_def by auto lemma msetext_huet_singleton: "y \ x \ msetext_huet gt [y] [x] \ gt y x" unfolding msetext_huet_def by simp lemma msetext_huet_wf: "wfP (\x y. gt y x) \ wfP (\xs ys. msetext_huet gt ys xs)" by (erule wfP_subset[OF msetext_dersh_wf]) (auto intro: msetext_huet_imp_dersh) lemma msetext_huet_hd_or_tl: assumes trans: "\z y x. gt z y \ gt y x \ gt z x" and total: "\y x. gt y x \ gt x y \ y = x" and len_eq: "length ys = length xs" and yys_gt_xxs: "msetext_huet gt (y # ys) (x # xs)" shows "gt y x \ msetext_huet gt ys xs" proof - let ?Y = "mset (y # ys)" let ?X = "mset (x # xs)" let ?Ya = "mset ys" let ?Xa = "mset xs" have Y_ne_X: "?Y \ ?X" and ex_gt_Y: "\xa. count ?X xa > count ?Y xa \ \ya. gt ya xa \ count ?Y ya > count ?X ya" using yys_gt_xxs[unfolded msetext_huet_def Let_def] by auto obtain yy where yy: "\xa. count ?X xa > count ?Y xa \ gt (yy xa) xa \ count ?Y (yy xa) > count ?X (yy xa)" using ex_gt_Y by metis have cnt_Y_pres: "count ?Ya xa > count ?Xa xa" if "count ?Y xa > count ?X xa" and "xa \ y" for xa using that by (auto split: if_splits) have cnt_X_pres: "count ?Xa xa > count ?Ya xa" if "count ?X xa > count ?Y xa" and "xa \ x" for xa using that by (auto split: if_splits) { assume y_eq_x: "y = x" have "?Xa \ ?Ya" using y_eq_x Y_ne_X by simp moreover have "\xa. count ?Xa xa > count ?Ya xa \ \ya. gt ya xa \ count ?Ya ya > count ?Xa ya" proof - fix xa :: 'a assume a1: "count (mset ys) xa < count (mset xs) xa" from ex_gt_Y obtain aa :: "'a \ 'a" where f3: "\a. \ count (mset (y # ys)) a < count (mset (x # xs)) a \ gt (aa a) a \ count (mset (x # xs)) (aa a) < count (mset (y # ys)) (aa a)" by (metis (full_types)) then have f4: "\a. count (mset (x # xs)) (aa a) < count (mset (x # ys)) (aa a) \ \ count (mset (x # ys)) a < count (mset (x # xs)) a" using y_eq_x by meson have "\a as aa. count (mset ((a::'a) # as)) aa = count (mset as) aa \ aa = a" by fastforce then have "xa = x \ count (mset (x # xs)) (aa xa) < count (mset (x # ys)) (aa xa)" using f4 a1 by (metis (no_types)) then show "\a. gt a xa \ count (mset xs) a < count (mset ys) a" using f3 y_eq_x a1 by (metis (no_types) Suc_less_eq count_add_mset mset.simps(2)) qed ultimately have "msetext_huet gt ys xs" unfolding msetext_huet_def Let_def by simp } moreover { assume x_gt_y: "gt x y" and y_ngt_x: "\ gt y x" hence y_ne_x: "y \ x" by fast obtain z where z_cnt: "count ?X z > count ?Y z" using size_eq_ex_count_lt[of ?Y ?X] size_mset size_mset len_eq Y_ne_X by auto have Xa_ne_Ya: "?Xa \ ?Ya" proof (cases "z = x") case True hence "yy z \ y" using y_ngt_x yy z_cnt by blast hence "count ?Ya (yy z) > count ?Xa (yy z)" using cnt_Y_pres yy z_cnt by blast thus ?thesis by auto next case False hence "count ?Xa z > count ?Ya z" using z_cnt cnt_X_pres by blast thus ?thesis by auto qed have "\ya. gt ya xa \ count ?Ya ya > count ?Xa ya" if xa_cnta: "count ?Xa xa > count ?Ya xa" for xa proof (cases "xa = y") case xa_eq_y: True { assume "count ?Ya x > count ?Xa x" moreover have "gt x xa" unfolding xa_eq_y by (rule x_gt_y) ultimately have ?thesis by fast } moreover { assume "count ?Xa x \ count ?Ya x" hence x_cnt: "count ?X x > count ?Y x" by (simp add: y_ne_x) hence yyx_gt_x: "gt (yy x) x" and yyx_cnt: "count ?Y (yy x) > count ?X (yy x)" using yy by blast+ have yyx_ne_y: "yy x \ y" using y_ngt_x yyx_gt_x by auto have "gt (yy x) xa" unfolding xa_eq_y using trans yyx_gt_x x_gt_y by blast moreover have "count ?Ya (yy x) > count ?Xa (yy x)" using cnt_Y_pres yyx_cnt yyx_ne_y by blast ultimately have ?thesis by blast } ultimately show ?thesis by fastforce next case False hence xa_cnt: "count ?X xa > count ?Y xa" using xa_cnta by fastforce show ?thesis proof (cases "yy xa = y \ count ?Ya y \ count ?Xa y") case yyxa_ne_y_or: False have yyxa_gt_xa: "gt (yy xa) xa" and yyxa_cnt: "count ?Y (yy xa) > count ?X (yy xa)" using yy[OF xa_cnt] by blast+ have "count ?Ya (yy xa) > count ?Xa (yy xa)" using cnt_Y_pres yyxa_cnt yyxa_ne_y_or by fastforce thus ?thesis using yyxa_gt_xa by blast next case True note yyxa_eq_y = this[THEN conjunct1] and y_cnt = this[THEN conjunct2] { assume "count ?Ya x > count ?Xa x" moreover have "gt x xa" using trans x_gt_y xa_cnt yy yyxa_eq_y by blast ultimately have ?thesis by fast } moreover { assume "count ?Xa x \ count ?Ya x" hence x_cnt: "count ?X x > count ?Y x" by (simp add: y_ne_x) hence yyx_gt_x: "gt (yy x) x" and yyx_cnt: "count ?Y (yy x) > count ?X (yy x)" using yy by blast+ have yyx_ne_y: "yy x \ y" using y_ngt_x yyx_gt_x by auto have "gt (yy x) xa" using trans x_gt_y xa_cnt yy yyx_gt_x yyxa_eq_y by blast moreover have "count ?Ya (yy x) > count ?Xa (yy x)" using cnt_Y_pres yyx_cnt yyx_ne_y by blast ultimately have ?thesis by blast } ultimately show ?thesis by fastforce qed qed hence "msetext_huet gt ys xs" unfolding msetext_huet_def Let_def using Xa_ne_Ya by fast } ultimately show ?thesis using total by blast qed interpretation msetext_huet: ext msetext_huet by standard (fact msetext_huet_mono_strong, fact msetext_huet_map) interpretation msetext_huet: ext_irrefl_before_trans msetext_huet by standard (fact msetext_huet_irrefl, fact msetext_huet_trans_from_irrefl) interpretation msetext_huet: ext_snoc msetext_huet by standard (fact msetext_huet_snoc) interpretation msetext_huet: ext_compat_cons msetext_huet by standard (fact msetext_huet_compat_cons) interpretation msetext_huet: ext_compat_snoc msetext_huet by standard (fact msetext_huet_compat_snoc) interpretation msetext_huet: ext_compat_list msetext_huet by standard (fact msetext_huet_compat_list) interpretation msetext_huet: ext_singleton msetext_huet by standard (fact msetext_huet_singleton) interpretation msetext_huet: ext_wf msetext_huet by standard (fact msetext_huet_wf) interpretation msetext_huet: ext_hd_or_tl msetext_huet by standard (rule msetext_huet_hd_or_tl) interpretation msetext_huet: ext_wf_bounded msetext_huet by standard subsection \Componentwise Extension\ definition cwiseext :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" where "cwiseext gt ys xs \ length ys = length xs \ (\i < length ys. gt (ys ! i) (xs ! i) \ ys ! i = xs ! i) \ (\i < length ys. gt (ys ! i) (xs ! i))" lemma cwiseext_imp_len_lexext: assumes cw: "cwiseext gt ys xs" shows "len_lexext gt ys xs" proof - have len_eq: "length ys = length xs" using cw[unfolded cwiseext_def] by sat moreover have "lexext gt ys xs" proof - obtain j where j_len: "j < length ys" and j_gt: "gt (ys ! j) (xs ! j)" using cw[unfolded cwiseext_def] by blast then obtain j0 where j0_len: "j0 < length ys" and j0_gt: "gt (ys ! j0) (xs ! j0)" and j0_min: "\i. i < j0 \ \ gt (ys ! i) (xs ! i)" using wf_eq_minimal[THEN iffD1, OF wf_less, rule_format, of _ "{i. gt (ys ! i) (xs ! i)}", simplified, OF j_gt] by (metis less_trans nat_neq_iff) have j0_eq: "\i. i < j0 \ ys ! i = xs ! i" using cw[unfolded cwiseext_def] by (metis j0_len j0_min less_trans) have "lexext gt (drop j0 ys) (drop j0 xs)" using lexext_Cons[of gt _ _ "drop (Suc j0) ys" "drop (Suc j0) xs", OF j0_gt] by (metis Cons_nth_drop_Suc j0_len len_eq) thus ?thesis using cw len_eq j0_len j0_min proof (induct j0 arbitrary: ys xs) case (Suc k) note ih0 = this(1) and gts_dropSk = this(2) and cw = this(3) and len_eq = this(4) and Sk_len = this(5) and Sk_min = this(6) have Sk_eq: "\i. i < Suc k \ ys ! i = xs ! i" using cw[unfolded cwiseext_def] by (metis Sk_len Sk_min less_trans) have k_len: "k < length ys" using Sk_len by simp have k_min: "\i. i < k \ \ gt (ys ! i) (xs ! i)" using Sk_min by simp have k_eq: "\i. i < k \ ys ! i = xs ! i" using Sk_eq by simp note ih = ih0[OF _ cw len_eq k_len k_min] show ?case proof (cases "k < length ys") case k_lt_ys: True note k_lt_xs = k_lt_ys[unfolded len_eq] obtain x where x: "x = xs ! k" by simp hence y: "x = ys ! k" using Sk_eq[of k] by simp have dropk_xs: "drop k xs = x # drop (Suc k) xs" using k_lt_xs x by (simp add: Cons_nth_drop_Suc) have dropk_ys: "drop k ys = x # drop (Suc k) ys" using k_lt_ys y by (simp add: Cons_nth_drop_Suc) show ?thesis by (rule ih, unfold dropk_xs dropk_ys, rule lexext_Cons_eq[OF gts_dropSk]) next case False hence "drop k xs = []" and "drop k ys = []" using len_eq by simp_all hence "lexext gt [] []" using gts_dropSk by simp hence "lexext gt (drop k ys) (drop k xs)" by simp thus ?thesis by (rule ih) qed qed simp qed ultimately show ?thesis unfolding lenext_def by sat qed lemma cwiseext_mono_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt' y x) \ cwiseext gt ys xs \ cwiseext gt' ys xs" unfolding cwiseext_def by (induct, force, fast) lemma cwiseext_map_strong: "(\y \ set ys. \x \ set xs. gt y x \ gt (f y) (f x)) \ cwiseext gt ys xs \ cwiseext gt (map f ys) (map f xs)" unfolding cwiseext_def by auto lemma cwiseext_irrefl: "(\x \ set xs. \ gt x x) \ \ cwiseext gt xs xs" unfolding cwiseext_def by (blast intro: nth_mem) lemma cwiseext_trans_strong: assumes "\z \ set zs. \y \ set ys. \x \ set xs. gt z y \ gt y x \ gt z x" and "cwiseext gt zs ys" and "cwiseext gt ys xs" shows "cwiseext gt zs xs" using assms unfolding cwiseext_def by (metis (mono_tags) nth_mem) lemma cwiseext_compat_cons: "cwiseext gt ys xs \ cwiseext gt (x # ys) (x # xs)" unfolding cwiseext_def proof (elim conjE, intro conjI) assume "length ys = length xs" and "\i < length ys. gt (ys ! i) (xs ! i) \ ys ! i = xs ! i" thus "\i < length (x # ys). gt ((x # ys) ! i) ((x # xs) ! i) \ (x # ys) ! i = (x # xs) ! i" by (simp add: nth_Cons') next assume "\i < length ys. gt (ys ! i) (xs ! i)" thus "\i < length (x # ys). gt ((x # ys) ! i) ((x # xs) ! i)" by fastforce qed auto lemma cwiseext_compat_snoc: "cwiseext gt ys xs \ cwiseext gt (ys @ [x]) (xs @ [x])" unfolding cwiseext_def proof (elim conjE, intro conjI) assume "length ys = length xs" and "\i < length ys. gt (ys ! i) (xs ! i) \ ys ! i = xs ! i" thus "\i < length (ys @ [x]). gt ((ys @ [x]) ! i) ((xs @ [x]) ! i) \ (ys @ [x]) ! i = (xs @ [x]) ! i" by (simp add: nth_append) next assume "length ys = length xs" and "\i < length ys. gt (ys ! i) (xs ! i)" thus "\i < length (ys @ [x]). gt ((ys @ [x]) ! i) ((xs @ [x]) ! i)" by (metis length_append_singleton less_Suc_eq nth_append) qed auto lemma cwiseext_compat_list: assumes y_gt_x: "gt y x" shows "cwiseext gt (xs @ y # xs') (xs @ x # xs')" unfolding cwiseext_def proof (intro conjI) show "\i < length (xs @ y # xs'). gt ((xs @ y # xs') ! i) ((xs @ x # xs') ! i) \ (xs @ y # xs') ! i = (xs @ x # xs') ! i" using y_gt_x by (simp add: nth_Cons' nth_append) next show "\i < length (xs @ y # xs'). gt ((xs @ y # xs') ! i) ((xs @ x # xs') ! i)" using y_gt_x by (metis add_diff_cancel_right' append_is_Nil_conv diff_less length_append length_greater_0_conv list.simps(3) nth_append_length) qed auto lemma cwiseext_singleton: "cwiseext gt [y] [x] \ gt y x" unfolding cwiseext_def by auto lemma cwiseext_wf: "wfP (\x y. gt y x) \ wfP (\xs ys. cwiseext gt ys xs)" by (auto intro: cwiseext_imp_len_lexext wfP_subset[OF len_lexext_wf]) lemma cwiseext_hd_or_tl: "cwiseext gt (y # ys) (x # xs) \ gt y x \ cwiseext gt ys xs" unfolding cwiseext_def proof (elim conjE, intro disj_imp[THEN iffD2, rule_format] conjI) assume "\i < length (y # ys). gt ((y # ys) ! i) ((x # xs) ! i)" and "\ gt y x" thus "\i < length ys. gt (ys ! i) (xs ! i)" by (metis (no_types) One_nat_def diff_le_self diff_less dual_order.strict_trans2 length_Cons less_Suc_eq linorder_neqE_nat not_less0 nth_Cons') qed auto locale ext_cwiseext = ext_compat_list + ext_compat_cons begin context fixes gt :: "'a \ 'a \ bool" assumes gt_irrefl: "\ gt x x" and trans_gt: "ext gt zs ys \ ext gt ys xs \ ext gt zs xs" begin lemma assumes ys_gtcw_xs: "cwiseext gt ys xs" shows "ext gt ys xs" proof - have "length ys = length xs" by (rule ys_gtcw_xs[unfolded cwiseext_def, THEN conjunct1]) thus ?thesis using ys_gtcw_xs proof (induct rule: list_induct2) case Nil thus ?case unfolding cwiseext_def by simp next case (Cons y ys x xs) note len_ys_eq_xs = this(1) and ih = this(2) and yys_gtcw_xxs = this(3) have xys_gts_xxs: "ext gt (x # ys) (x # xs)" if ys_ne_xs: "ys \ xs" proof - have ys_gtcw_xs: "cwiseext gt ys xs" using yys_gtcw_xxs unfolding cwiseext_def proof (elim conjE, intro conjI) assume "\i < length (y # ys). gt ((y # ys) ! i) ((x # xs) ! i) \ (y # ys) ! i = (x # xs) ! i" hence ge: "\i < length ys. gt (ys ! i) (xs ! i) \ ys ! i = xs ! i" by auto thus "\i < length ys. gt (ys ! i) (xs ! i)" using ys_ne_xs len_ys_eq_xs nth_equalityI by blast qed auto hence "ext gt ys xs" by (rule ih) thus "ext gt (x # ys) (x # xs)" by (rule compat_cons) qed have "gt y x \ y = x" using yys_gtcw_xxs unfolding cwiseext_def by fastforce moreover { assume y_eq_x: "y = x" have ?case proof (cases "ys = xs") case True hence False using y_eq_x gt_irrefl yys_gtcw_xxs unfolding cwiseext_def by presburger thus ?thesis by sat next case False thus ?thesis using y_eq_x xys_gts_xxs by simp qed } moreover { assume "y \ x" and "gt y x" hence yys_gts_xys: "ext gt (y # ys) (x # ys)" using compat_list[of _ _ gt "[]"] by simp have ?case proof (cases "ys = xs") case ys_eq_xs: True thus ?thesis using yys_gts_xys by simp next case False thus ?thesis using yys_gts_xys xys_gts_xxs trans_gt by blast qed } ultimately show ?case by sat qed qed end end interpretation cwiseext: ext cwiseext by standard (fact cwiseext_mono_strong, rule cwiseext_map_strong, metis in_listsD) interpretation cwiseext: ext_irrefl_trans_strong cwiseext by standard (fact cwiseext_irrefl, fact cwiseext_trans_strong) interpretation cwiseext: ext_compat_cons cwiseext by standard (fact cwiseext_compat_cons) interpretation cwiseext: ext_compat_snoc cwiseext by standard (fact cwiseext_compat_snoc) interpretation cwiseext: ext_compat_list cwiseext by standard (rule cwiseext_compat_list) interpretation cwiseext: ext_singleton cwiseext by standard (rule cwiseext_singleton) interpretation cwiseext: ext_wf cwiseext by standard (rule cwiseext_wf) interpretation cwiseext: ext_hd_or_tl cwiseext by standard (rule cwiseext_hd_or_tl) interpretation cwiseext: ext_wf_bounded cwiseext by standard +(* TODO: move ? *) +lemma less_multiset_doubletons: + assumes + "y < t \ y < s" + "x < t \ x < s" + shows + "{# y, x#} < {# t, s#}" + unfolding less_multiset\<^sub>D\<^sub>M +proof (intro exI) + let ?X = "{# t, s#}" + let ?Y = "{#y, x#}" + show "?X \ {#} \ ?X \# {#t, s#} \ {#y, x#} = {#t, s#} - ?X + ?Y \ (\k. k \# ?Y \ (\a. a \# ?X \ k < a))" + using add_eq_conv_diff assms(1) assms(2) by auto +qed + end diff --git a/thys/Lambda_Free_RPOs/Lambda_Free_RPO_Std.thy b/thys/Lambda_Free_RPOs/Lambda_Free_RPO_Std.thy --- a/thys/Lambda_Free_RPOs/Lambda_Free_RPO_Std.thy +++ b/thys/Lambda_Free_RPOs/Lambda_Free_RPO_Std.thy @@ -1,826 +1,826 @@ (* Title: The Graceful Recursive Path Order for Lambda-Free Higher-Order Terms Author: Jasmin Blanchette , 2016 Author: Uwe Waldmann , 2016 Author: Daniel Wand , 2016 Maintainer: Jasmin Blanchette *) section \The Graceful Recursive Path Order for Lambda-Free Higher-Order Terms\ theory Lambda_Free_RPO_Std imports Lambda_Free_Term Extension_Orders abbrevs ">t" = ">\<^sub>t" and "\t" = "\\<^sub>t" begin text \ This theory defines the graceful recursive path order (RPO) for \\\-free higher-order terms. \ subsection \Setup\ locale rpo_basis = ground_heads "(>\<^sub>s)" arity_sym arity_var for gt_sym :: "'s \ 's \ bool" (infix ">\<^sub>s" 50) and arity_sym :: "'s \ enat" and arity_var :: "'v \ enat" + fixes extf :: "'s \ (('s, 'v) tm \ ('s, 'v) tm \ bool) \ ('s, 'v) tm list \ ('s, 'v) tm list \ bool" assumes extf_ext_trans_before_irrefl: "ext_trans_before_irrefl (extf f)" and extf_ext_compat_cons: "ext_compat_cons (extf f)" and extf_ext_compat_list: "ext_compat_list (extf f)" begin lemma extf_ext_trans: "ext_trans (extf f)" by (rule ext_trans_before_irrefl.axioms(1)[OF extf_ext_trans_before_irrefl]) lemma extf_ext: "ext (extf f)" by (rule ext_trans.axioms(1)[OF extf_ext_trans]) lemmas extf_mono_strong = ext.mono_strong[OF extf_ext] lemmas extf_mono = ext.mono[OF extf_ext, mono] lemmas extf_map = ext.map[OF extf_ext] lemmas extf_trans = ext_trans.trans[OF extf_ext_trans] lemmas extf_irrefl_from_trans = ext_trans_before_irrefl.irrefl_from_trans[OF extf_ext_trans_before_irrefl] lemmas extf_compat_append_left = ext_compat_cons.compat_append_left[OF extf_ext_compat_cons] lemmas extf_compat_list = ext_compat_list.compat_list[OF extf_ext_compat_list] definition chkvar :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where [simp]: "chkvar t s \ vars_hd (head s) \ vars t" end locale rpo = rpo_basis _ _ arity_sym arity_var for arity_sym :: "'s \ enat" and arity_var :: "'v \ enat" begin subsection \Inductive Definitions\ definition chksubs :: "(('s, 'v) tm \ ('s, 'v) tm \ bool) \ ('s, 'v) tm \ ('s, 'v) tm \ bool" where [simp]: "chksubs gt t s \ (case s of App s1 s2 \ gt t s1 \ gt t s2 | _ \ True)" lemma chksubs_mono[mono]: "gt \ gt' \ chksubs gt \ chksubs gt'" by (auto simp: tm.case_eq_if) force+ inductive gt :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix ">\<^sub>t" 50) where gt_sub: "is_App t \ (fun t >\<^sub>t s \ fun t = s) \ (arg t >\<^sub>t s \ arg t = s) \ t >\<^sub>t s" | gt_diff: "head t >\<^sub>h\<^sub>d head s \ chkvar t s \ chksubs (>\<^sub>t) t s \ t >\<^sub>t s" | gt_same: "head t = head s \ chksubs (>\<^sub>t) t s \ (\f \ ground_heads (head t). extf f (>\<^sub>t) (args t) (args s)) \ t >\<^sub>t s" abbreviation ge :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix "\\<^sub>t" 50) where "t \\<^sub>t s \ t >\<^sub>t s \ t = s" inductive gt_sub :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_subI: "is_App t \ fun t \\<^sub>t s \ arg t \\<^sub>t s \ gt_sub t s" inductive gt_diff :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_diffI: "head t >\<^sub>h\<^sub>d head s \ chkvar t s \ chksubs (>\<^sub>t) t s \ gt_diff t s" inductive gt_same :: "('s, 'v) tm \ ('s, 'v) tm \ bool" where gt_sameI: "head t = head s \ chksubs (>\<^sub>t) t s \ (\f \ ground_heads (head t). extf f (>\<^sub>t) (args t) (args s)) \ gt_same t s" lemma gt_iff_sub_diff_same: "t >\<^sub>t s \ gt_sub t s \ gt_diff t s \ gt_same t s" by (subst gt.simps) (auto simp: gt_sub.simps gt_diff.simps gt_same.simps) subsection \Transitivity\ lemma gt_fun_imp: "fun t >\<^sub>t s \ t >\<^sub>t s" by (cases t) (auto intro: gt_sub) lemma gt_arg_imp: "arg t >\<^sub>t s \ t >\<^sub>t s" by (cases t) (auto intro: gt_sub) lemma gt_imp_vars: "t >\<^sub>t s \ vars t \ vars s" proof (simp only: atomize_imp, rule measure_induct_rule[of "\(t, s). size t + size s" "\(t, s). t >\<^sub>t s \ vars t \ vars s" "(t, s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric]) fix t s assume ih: "\ta sa. size ta + size sa < size t + size s \ ta >\<^sub>t sa \ vars ta \ vars sa" and t_gt_s: "t >\<^sub>t s" show "vars t \ vars s" using t_gt_s proof cases case gt_sub thus ?thesis using ih[of "fun t" s] ih[of "arg t" s] by (meson add_less_cancel_right subsetD size_arg_lt size_fun_lt subsetI tm.set_sel(5,6)) next case gt_diff show ?thesis proof (cases s) case Hd thus ?thesis using gt_diff(2) by (auto elim: hd.set_cases(2)) next case (App s1 s2) thus ?thesis using gt_diff(3) ih[of t s1] ih[of t s2] by simp qed next case gt_same show ?thesis proof (cases s) case Hd thus ?thesis using gt_same(1) vars_head_subseteq by fastforce next case (App s1 s2) thus ?thesis using gt_same(2) ih[of t s1] ih[of t s2] by simp qed qed qed theorem gt_trans: "u >\<^sub>t t \ t >\<^sub>t s \ u >\<^sub>t s" proof (simp only: atomize_imp, rule measure_induct_rule[of "\(u, t, s). {#size u, size t, size s#}" "\(u, t, s). u >\<^sub>t t \ t >\<^sub>t s \ u >\<^sub>t s" "(u, t, s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric]) fix u t s assume ih: "\ua ta sa. {#size ua, size ta, size sa#} < {#size u, size t, size s#} \ ua >\<^sub>t ta \ ta >\<^sub>t sa \ ua >\<^sub>t sa" and u_gt_t: "u >\<^sub>t t" and t_gt_s: "t >\<^sub>t s" have chkvar: "chkvar u s" by clarsimp (meson u_gt_t t_gt_s gt_imp_vars hd.set_sel(2) vars_head_subseteq subsetCE) have chk_u_s_if: "chksubs (>\<^sub>t) u s" if chk_t_s: "chksubs (>\<^sub>t) t s" proof (cases s) case (App s1 s2) thus ?thesis using chk_t_s by (auto intro: ih[of _ _ s1, OF _ u_gt_t] ih[of _ _ s2, OF _ u_gt_t]) qed auto have fun_u_lt_etc: "is_App u \ {#size (fun u), size t, size s#} < {#size u, size t, size s#}" and arg_u_lt_etc: "is_App u \ {#size (arg u), size t, size s#} < {#size u, size t, size s#}" by (simp_all add: size_fun_lt size_arg_lt) have u_gt_s_if_ui: "is_App u \ fun u \\<^sub>t t \ arg u \\<^sub>t t \ u >\<^sub>t s" using ih[of "fun u" t s, OF fun_u_lt_etc] ih[of "arg u" t s, OF arg_u_lt_etc] gt_arg_imp gt_fun_imp t_gt_s by blast show "u >\<^sub>t s" using t_gt_s proof cases case gt_sub_t_s: gt_sub have u_gt_s_if_chk_u_t: ?thesis if chk_u_t: "chksubs (>\<^sub>t) u t" using gt_sub_t_s(1) proof (cases t) case t: (App t1 t2) show ?thesis using ih[of u t1 s] ih[of u t2 s] gt_sub_t_s(2) chk_u_t unfolding t by auto qed auto show ?thesis using u_gt_t by cases (auto intro: u_gt_s_if_ui u_gt_s_if_chk_u_t) next case gt_diff_t_s: gt_diff show ?thesis using u_gt_t proof cases case gt_diff_u_t: gt_diff have "head u >\<^sub>h\<^sub>d head s" using gt_diff_u_t(1) gt_diff_t_s(1) by (auto intro: gt_hd_trans) thus ?thesis by (rule gt_diff[OF _ chkvar chk_u_s_if[OF gt_diff_t_s(3)]]) next case gt_same_u_t: gt_same have "head u >\<^sub>h\<^sub>d head s" using gt_diff_t_s(1) gt_same_u_t(1) by auto thus ?thesis by (rule gt_diff[OF _ chkvar chk_u_s_if[OF gt_diff_t_s(3)]]) qed (auto intro: u_gt_s_if_ui) next case gt_same_t_s: gt_same show ?thesis using u_gt_t proof cases case gt_diff_u_t: gt_diff have "head u >\<^sub>h\<^sub>d head s" using gt_diff_u_t(1) gt_same_t_s(1) by simp thus ?thesis by (rule gt_diff[OF _ chkvar chk_u_s_if[OF gt_same_t_s(2)]]) next case gt_same_u_t: gt_same have hd_u_s: "head u = head s" using gt_same_u_t(1) gt_same_t_s(1) by simp let ?S = "set (args u) \ set (args t) \ set (args s)" have gt_trans_args: "\ua \ ?S. \ta \ ?S. \sa \ ?S. ua >\<^sub>t ta \ ta >\<^sub>t sa \ ua >\<^sub>t sa" proof clarify fix sa ta ua assume ua_in: "ua \ ?S" and ta_in: "ta \ ?S" and sa_in: "sa \ ?S" and ua_gt_ta: "ua >\<^sub>t ta" and ta_gt_sa: "ta >\<^sub>t sa" show "ua >\<^sub>t sa" by (auto intro!: ih[OF Max_lt_imp_lt_mset ua_gt_ta ta_gt_sa]) (meson ua_in ta_in sa_in Un_iff max.strict_coboundedI1 max.strict_coboundedI2 size_in_args)+ qed have "\f \ ground_heads (head u). extf f (>\<^sub>t) (args u) (args s)" proof (clarify, rule extf_trans[OF _ _ _ gt_trans_args]) fix f assume f_in_grounds: "f \ ground_heads (head u)" show "extf f (>\<^sub>t) (args u) (args t)" using f_in_grounds gt_same_u_t(3) by blast show "extf f (>\<^sub>t) (args t) (args s)" using f_in_grounds gt_same_t_s(3) unfolding gt_same_u_t(1) by blast qed auto thus ?thesis by (rule gt_same[OF hd_u_s chk_u_s_if[OF gt_same_t_s(2)]]) qed (auto intro: u_gt_s_if_ui) qed qed subsection \Irreflexivity\ theorem gt_irrefl: "\ s >\<^sub>t s" proof (standard, induct s rule: measure_induct_rule[of size]) case (less s) note ih = this(1) and s_gt_s = this(2) show False using s_gt_s proof cases case _: gt_sub note is_app = this(1) and si_ge_s = this(2) have s_gt_fun_s: "s >\<^sub>t fun s" and s_gt_arg_s: "s >\<^sub>t arg s" using is_app by (simp_all add: gt_sub) have "fun s >\<^sub>t s \ arg s >\<^sub>t s" using si_ge_s is_app s_gt_arg_s s_gt_fun_s by auto moreover { assume fun_s_gt_s: "fun s >\<^sub>t s" have "fun s >\<^sub>t fun s" by (rule gt_trans[OF fun_s_gt_s s_gt_fun_s]) hence False using ih[of "fun s"] is_app size_fun_lt by blast } moreover { assume arg_s_gt_s: "arg s >\<^sub>t s" have "arg s >\<^sub>t arg s" by (rule gt_trans[OF arg_s_gt_s s_gt_arg_s]) hence False using ih[of "arg s"] is_app size_arg_lt by blast } ultimately show False by sat next case gt_diff thus False by (cases "head s") (auto simp: gt_hd_irrefl) next case gt_same note in_grounds = this(3) obtain si where si_in_args: "si \ set (args s)" and si_gt_si: "si >\<^sub>t si" using in_grounds by (metis (full_types) all_not_in_conv extf_irrefl_from_trans ground_heads_nonempty gt_trans) have "size si < size s" by (rule size_in_args[OF si_in_args]) thus False by (rule ih[OF _ si_gt_si]) qed qed lemma gt_antisym: "t >\<^sub>t s \ \ s >\<^sub>t t" using gt_irrefl gt_trans by blast subsection \Subterm Property\ lemma gt_sub_fun: "App s t >\<^sub>t s" and gt_sub_arg: "App s t >\<^sub>t t" by (auto intro: gt_sub) theorem gt_proper_sub: "proper_sub s t \ t >\<^sub>t s" by (induct t) (auto intro: gt_sub_fun gt_sub_arg gt_trans) subsection \Compatibility with Functions\ lemma gt_compat_fun: assumes t'_gt_t: "t' >\<^sub>t t" shows "App s t' >\<^sub>t App s t" proof - have t'_ne_t: "t' \ t" using gt_antisym t'_gt_t by blast have extf_args_single: "\f \ ground_heads (head s). extf f (>\<^sub>t) (args s @ [t']) (args s @ [t])" by (simp add: extf_compat_list t'_gt_t t'_ne_t) show ?thesis by (rule gt_same) (auto simp: gt_sub gt_sub_fun t'_gt_t intro!: extf_args_single) qed theorem gt_compat_fun_strong: assumes t'_gt_t: "t' >\<^sub>t t" shows "apps s (t' # us) >\<^sub>t apps s (t # us)" proof (induct us rule: rev_induct) case Nil show ?case using t'_gt_t by (auto intro!: gt_compat_fun) next case (snoc u us) note ih = snoc let ?v' = "apps s (t' # us @ [u])" let ?v = "apps s (t # us @ [u])" show ?case proof (rule gt_same) show "chksubs (>\<^sub>t) ?v' ?v" using ih by (auto intro: gt_sub gt_sub_arg) next show "\f \ ground_heads (head ?v'). extf f (>\<^sub>t) (args ?v') (args ?v)" by (metis args_apps extf_compat_list gt_irrefl t'_gt_t) qed simp qed subsection \Compatibility with Arguments\ theorem gt_diff_same_compat_arg: assumes extf_compat_snoc: "\f. ext_compat_snoc (extf f)" and diff_same: "gt_diff s' s \ gt_same s' s" shows "App s' t >\<^sub>t App s t" proof - { assume "s' >\<^sub>t s" hence "App s' t >\<^sub>t s" using gt_sub_fun gt_trans by blast moreover have "App s' t >\<^sub>t t" by (simp add: gt_sub_arg) ultimately have "chksubs (>\<^sub>t) (App s' t) (App s t)" by auto } note chk_s't_st = this show ?thesis using diff_same proof assume "gt_diff s' s" hence s'_gt_s: "s' >\<^sub>t s" and hd_s'_gt_s: "head s' >\<^sub>h\<^sub>d head s" and chkvar_s'_s: "chkvar s' s" and chk_s'_s: "chksubs (>\<^sub>t) s' s" using gt_diff.cases by (auto simp: gt_iff_sub_diff_same) have chkvar_s't_st: "chkvar (App s' t) (App s t)" using chkvar_s'_s by auto show ?thesis by (rule gt_diff[OF _ chkvar_s't_st chk_s't_st[OF s'_gt_s]]) (simp add: hd_s'_gt_s[simplified]) next assume "gt_same s' s" hence s'_gt_s: "s' >\<^sub>t s" and hd_s'_eq_s: "head s' = head s" and chk_s'_s: "chksubs (>\<^sub>t) s' s" and gts_args: "\f \ ground_heads (head s'). extf f (>\<^sub>t) (args s') (args s)" using gt_same.cases by (auto simp: gt_iff_sub_diff_same, metis) have gts_args_t: "\f \ ground_heads (head (App s' t)). extf f (>\<^sub>t) (args (App s' t)) (args (App s t))" using gts_args ext_compat_snoc.compat_append_right[OF extf_compat_snoc] by simp show ?thesis by (rule gt_same[OF _ chk_s't_st[OF s'_gt_s] gts_args_t]) (simp add: hd_s'_eq_s) qed qed subsection \Stability under Substitution\ lemma gt_imp_chksubs_gt: assumes t_gt_s: "t >\<^sub>t s" shows "chksubs (>\<^sub>t) t s" proof - have "is_App s \ t >\<^sub>t fun s \ t >\<^sub>t arg s" using t_gt_s by (meson gt_sub gt_trans) thus ?thesis by (simp add: tm.case_eq_if) qed theorem gt_subst: assumes wary_\: "wary_subst \" shows "t >\<^sub>t s \ subst \ t >\<^sub>t subst \ s" proof (simp only: atomize_imp, rule measure_induct_rule[of "\(t, s). {#size t, size s#}" "\(t, s). t >\<^sub>t s \ subst \ t >\<^sub>t subst \ s" "(t, s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric]) fix t s assume ih: "\ta sa. {#size ta, size sa#} < {#size t, size s#} \ ta >\<^sub>t sa \ subst \ ta >\<^sub>t subst \ sa" and t_gt_s: "t >\<^sub>t s" { assume chk_t_s: "chksubs (>\<^sub>t) t s" have "chksubs (>\<^sub>t) (subst \ t) (subst \ s)" proof (cases s) case s: (Hd \) show ?thesis proof (cases \) case \: (Var x) have psub_x_t: "proper_sub (Hd (Var x)) t" using \ s t_gt_s gt_imp_vars gt_irrefl in_vars_imp_sub by fastforce show ?thesis unfolding \ s by (rule gt_imp_chksubs_gt[OF gt_proper_sub[OF proper_sub_subst]]) (rule psub_x_t) qed (auto simp: s) next case s: (App s1 s2) have "t >\<^sub>t s1" and "t >\<^sub>t s2" using s chk_t_s by auto thus ?thesis using s by (auto intro!: ih[of t s1] ih[of t s2]) qed } note chk_\t_\s_if = this show "subst \ t >\<^sub>t subst \ s" using t_gt_s proof cases case gt_sub_t_s: gt_sub obtain t1 t2 where t: "t = App t1 t2" using gt_sub_t_s(1) by (metis tm.collapse(2)) show ?thesis using gt_sub ih[of t1 s] ih[of t2 s] gt_sub_t_s(2) t by auto next case gt_diff_t_s: gt_diff have "head (subst \ t) >\<^sub>h\<^sub>d head (subst \ s)" by (meson wary_subst_ground_heads gt_diff_t_s(1) gt_hd_def subsetCE wary_\) moreover have "chkvar (subst \ t) (subst \ s)" unfolding chkvar_def using vars_subst_subseteq[OF gt_imp_vars[OF t_gt_s]] vars_head_subseteq by force ultimately show ?thesis by (rule gt_diff[OF _ _ chk_\t_\s_if[OF gt_diff_t_s(3)]]) next case gt_same_t_s: gt_same have hd_\t_eq_\s: "head (subst \ t) = head (subst \ s)" using gt_same_t_s(1) by simp { fix f assume f_in_grounds: "f \ ground_heads (head (subst \ t))" let ?S = "set (args t) \ set (args s)" have extf_args_s_t: "extf f (>\<^sub>t) (args t) (args s)" using gt_same_t_s(3) f_in_grounds wary_\ wary_subst_ground_heads by blast have "extf f (>\<^sub>t) (map (subst \) (args t)) (map (subst \) (args s))" proof (rule extf_map[of ?S, OF _ _ _ _ _ _ extf_args_s_t]) have sz_a: "\ta \ ?S. \sa \ ?S. {#size ta, size sa#} < {#size t, size s#}" by (fastforce intro: Max_lt_imp_lt_mset dest: size_in_args) show "\ta \ ?S. \sa \ ?S. ta >\<^sub>t sa \ subst \ ta >\<^sub>t subst \ sa" using ih sz_a size_in_args by fastforce qed (auto intro!: gt_irrefl elim!: gt_trans) hence "extf f (>\<^sub>t) (args (subst \ t)) (args (subst \ s))" by (auto simp: gt_same_t_s(1) intro: extf_compat_append_left) } hence "\f \ ground_heads (head (subst \ t)). extf f (>\<^sub>t) (args (subst \ t)) (args (subst \ s))" by blast thus ?thesis by (rule gt_same[OF hd_\t_eq_\s chk_\t_\s_if[OF gt_same_t_s(2)]]) qed qed subsection \Totality on Ground Terms\ theorem gt_total_ground: assumes extf_total: "\f. ext_total (extf f)" shows "ground t \ ground s \ t >\<^sub>t s \ s >\<^sub>t t \ t = s" proof (simp only: atomize_imp, - rule measure_induct_rule[of "\(t, s). size t + size s" + rule measure_induct_rule[of "\(t, s). {# size t, size s #}" "\(t, s). ground t \ ground s \ t >\<^sub>t s \ s >\<^sub>t t \ t = s" "(t, s)", simplified prod.case], simp only: split_paired_all prod.case atomize_imp[symmetric]) fix t s :: "('s, 'v) tm" assume - ih: "\ta sa. size ta + size sa < size t + size s \ ground ta \ ground sa \ + ih: "\ta sa. {# size ta, size sa #} < {# size t, size s #} \ ground ta \ ground sa \ ta >\<^sub>t sa \ sa >\<^sub>t ta \ ta = sa" and gr_t: "ground t" and gr_s: "ground s" let ?case = "t >\<^sub>t s \ s >\<^sub>t t \ t = s" have "chksubs (>\<^sub>t) t s \ s >\<^sub>t t" - unfolding chksubs_def tm.case_eq_if using ih[of t "fun s"] ih[of t "arg s"] - by (metis gt_sub add_less_cancel_left gr_s gr_t ground_arg ground_fun size_arg_lt size_fun_lt) + unfolding chksubs_def tm.case_eq_if using ih[of t "fun s"] ih[of t "arg s"] mset_lt_single_iff + by (metis add_mset_lt_right_lt gr_s gr_t ground_arg ground_fun gt_sub size_arg_lt size_fun_lt) moreover have "chksubs (>\<^sub>t) s t \ t >\<^sub>t s" unfolding chksubs_def tm.case_eq_if using ih[of "fun t" s] ih[of "arg t" s] - by (metis gt_sub add_less_cancel_right gr_s gr_t ground_arg ground_fun size_arg_lt size_fun_lt) + by (metis add_mset_lt_left_lt gr_s gr_t ground_arg ground_fun gt_sub size_arg_lt size_fun_lt) moreover { assume chksubs_t_s: "chksubs (>\<^sub>t) t s" and chksubs_s_t: "chksubs (>\<^sub>t) s t" obtain g where g: "head t = Sym g" using gr_t by (metis ground_head hd.collapse(2)) obtain f where f: "head s = Sym f" using gr_s by (metis ground_head hd.collapse(2)) have chkvar_t_s: "chkvar t s" and chkvar_s_t: "chkvar s t" using g f by simp_all { assume g_gt_f: "g >\<^sub>s f" have "t >\<^sub>t s" by (rule gt_diff[OF _ chkvar_t_s chksubs_t_s]) (simp add: g f gt_sym_imp_hd[OF g_gt_f]) } moreover { assume f_gt_g: "f >\<^sub>s g" have "s >\<^sub>t t" by (rule gt_diff[OF _ chkvar_s_t chksubs_s_t]) (simp add: g f gt_sym_imp_hd[OF f_gt_g]) } moreover { assume g_eq_f: "g = f" hence hd_t: "head t = head s" using g f by auto let ?ts = "args t" let ?ss = "args s" have gr_ts: "\ta \ set ?ts. ground ta" using ground_args[OF _ gr_t] by blast have gr_ss: "\sa \ set ?ss. ground sa" using ground_args[OF _ gr_s] by blast { assume ts_eq_ss: "?ts = ?ss" have "t = s" using hd_t ts_eq_ss by (rule tm_expand_apps) } moreover { assume ts_gt_ss: "extf g (>\<^sub>t) ?ts ?ss" have "t >\<^sub>t s" by (rule gt_same[OF hd_t chksubs_t_s]) (auto simp: g ts_gt_ss) } moreover { assume ss_gt_ts: "extf g (>\<^sub>t) ?ss ?ts" have "s >\<^sub>t t" by (rule gt_same[OF hd_t[symmetric] chksubs_s_t]) (auto simp: f[folded g_eq_f] ss_gt_ts) } ultimately have ?case using ih gr_ss gr_ts - ext_total.total[OF extf_total, rule_format, of "set ?ts" "set ?ss" "(>\<^sub>t)" ?ts ?ss g] - by (metis add_strict_mono in_listsI size_in_args) + ext_total.total[OF extf_total, rule_format, of "set ?ts \ set ?ss" "(>\<^sub>t)" ?ts ?ss g] + by (metis Un_iff in_listsI less_multiset_doubletons size_in_args) } ultimately have ?case using gt_sym_total by blast } ultimately show ?case by fast qed subsection \Well-foundedness\ abbreviation gtg :: "('s, 'v) tm \ ('s, 'v) tm \ bool" (infix ">\<^sub>t\<^sub>g" 50) where "(>\<^sub>t\<^sub>g) \ \t s. ground t \ t >\<^sub>t s" theorem gt_wf: assumes extf_wf: "\f. ext_wf (extf f)" shows "wfP (\s t. t >\<^sub>t s)" proof - have ground_wfP: "wfP (\s t. t >\<^sub>t\<^sub>g s)" unfolding wfP_iff_no_inf_chain proof assume "\f. inf_chain (>\<^sub>t\<^sub>g) f" then obtain t where t_bad: "bad (>\<^sub>t\<^sub>g) t" unfolding inf_chain_def bad_def by blast let ?ff = "worst_chain (>\<^sub>t\<^sub>g) (\t s. size t > size s)" let ?U_of = "\i. if is_App (?ff i) then {fun (?ff i)} \ set (args (?ff i)) else {}" note wf_sz = wf_app[OF wellorder_class.wf, of size, simplified] define U where "U = (\i. ?U_of i)" have gr: "\i. ground (?ff i)" using worst_chain_bad[OF wf_sz t_bad, unfolded inf_chain_def] by fast have gr_fun: "\i. ground (fun (?ff i))" by (rule ground_fun[OF gr]) have gr_args: "\i s. s \ set (args (?ff i)) \ ground s" by (rule ground_args[OF _ gr]) have gr_u: "\u. u \ U \ ground u" unfolding U_def by (auto dest: gr_args) (metis (lifting) empty_iff gr_fun) have "\ bad (>\<^sub>t\<^sub>g) u" if u_in: "u \ ?U_of i" for u i proof let ?ti = "?ff i" assume u_bad: "bad (>\<^sub>t\<^sub>g) u" have sz_u: "size u < size ?ti" proof (cases "?ff i") case Hd thus ?thesis using u_in size_in_args by fastforce next case App thus ?thesis using u_in size_in_args insert_iff size_fun_lt by fastforce qed show False proof (cases i) case 0 thus False using sz_u min_worst_chain_0[OF wf_sz u_bad] by simp next case Suc hence "?ff (i - 1) >\<^sub>t ?ff i" using worst_chain_pred[OF wf_sz t_bad] by simp moreover have "?ff i >\<^sub>t u" proof - have u_in: "u \ ?U_of i" using u_in by blast have ffi_ne_u: "?ff i \ u" using sz_u by fastforce hence "is_App (?ff i) \ \ sub u (?ff i) \ ?ff i >\<^sub>t u" using u_in gt_sub sub_args by auto thus "?ff i >\<^sub>t u" using ffi_ne_u u_in gt_proper_sub sub_args by fastforce qed ultimately have "?ff (i - 1) >\<^sub>t u" by (rule gt_trans) thus False using Suc sz_u min_worst_chain_Suc[OF wf_sz u_bad] gr by fastforce qed qed hence u_good: "\u. u \ U \ \ bad (>\<^sub>t\<^sub>g) u" unfolding U_def by blast have bad_diff_same: "inf_chain (\t s. ground t \ (gt_diff t s \ gt_same t s)) ?ff" unfolding inf_chain_def proof (intro allI conjI) fix i show "ground (?ff i)" by (rule gr) have gt: "?ff i >\<^sub>t ?ff (Suc i)" using worst_chain_pred[OF wf_sz t_bad] by blast have "\ gt_sub (?ff i) (?ff (Suc i))" proof assume a: "gt_sub (?ff i) (?ff (Suc i))" hence fi_app: "is_App (?ff i)" and fun_or_arg_fi_ge: "fun (?ff i) \\<^sub>t ?ff (Suc i) \ arg (?ff i) \\<^sub>t ?ff (Suc i)" unfolding gt_sub.simps by blast+ have "fun (?ff i) \ ?U_of i" unfolding U_def using fi_app by auto moreover have "arg (?ff i) \ ?U_of i" unfolding U_def using fi_app arg_in_args by force ultimately obtain uij where uij_in: "uij \ U" and uij_cases: "uij \\<^sub>t ?ff (Suc i)" unfolding U_def using fun_or_arg_fi_ge by blast have "\n. ?ff n >\<^sub>t ?ff (Suc n)" by (rule worst_chain_pred[OF wf_sz t_bad, THEN conjunct2]) hence uij_gt_i_plus_3: "uij >\<^sub>t ?ff (Suc (Suc i))" using gt_trans uij_cases by blast have "inf_chain (>\<^sub>t\<^sub>g) (\j. if j = 0 then uij else ?ff (Suc (i + j)))" unfolding inf_chain_def by (auto intro!: gr gr_u[OF uij_in] uij_gt_i_plus_3 worst_chain_pred[OF wf_sz t_bad]) hence "bad (>\<^sub>t\<^sub>g) uij" unfolding bad_def by fastforce thus False using u_good[OF uij_in] by sat qed thus "gt_diff (?ff i) (?ff (Suc i)) \ gt_same (?ff i) (?ff (Suc i))" using gt unfolding gt_iff_sub_diff_same by sat qed have "wf {(s, t). ground s \ ground t \ sym (head t) >\<^sub>s sym (head s)}" using gt_sym_wf unfolding wfP_def wf_iff_no_infinite_down_chain by fast moreover have "{(s, t). ground t \ gt_diff t s} \ {(s, t). ground s \ ground t \ sym (head t) >\<^sub>s sym (head s)}" proof (clarsimp, intro conjI) fix s t assume gr_t: "ground t" and gt_diff_t_s: "gt_diff t s" thus gr_s: "ground s" using gt_iff_sub_diff_same gt_imp_vars by fastforce show "sym (head t) >\<^sub>s sym (head s)" using gt_diff_t_s ground_head[OF gr_s] ground_head[OF gr_t] by (cases; cases "head s"; cases "head t") (auto simp: gt_hd_def) qed ultimately have wf_diff: "wf {(s, t). ground t \ gt_diff t s}" by (rule wf_subset) have diff_O_same: "{(s, t). ground t \ gt_diff t s} O {(s, t). ground t \ gt_same t s} \ {(s, t). ground t \ gt_diff t s}" unfolding gt_diff.simps gt_same.simps by clarsimp (metis chksubs_def empty_subsetI gt_diff[unfolded chkvar_def] gt_imp_chksubs_gt gt_same gt_trans) have diff_same_as_union: "{(s, t). ground t \ (gt_diff t s \ gt_same t s)} = {(s, t). ground t \ gt_diff t s} \ {(s, t). ground t \ gt_same t s}" by auto obtain k where bad_same: "inf_chain (\t s. ground t \ gt_same t s) (\i. ?ff (i + k))" using wf_infinite_down_chain_compatible[OF wf_diff _ diff_O_same, of ?ff] bad_diff_same unfolding inf_chain_def diff_same_as_union[symmetric] by auto hence hd_sym: "\i. is_Sym (head (?ff (i + k)))" unfolding inf_chain_def by (simp add: ground_head) define f where "f = sym (head (?ff k))" have hd_eq_f: "head (?ff (i + k)) = Sym f" for i proof (induct i) case 0 thus ?case by (auto simp: f_def hd.collapse(2)[OF hd_sym, of 0, simplified]) next case (Suc ia) thus ?case using bad_same unfolding inf_chain_def gt_same.simps by simp qed let ?gtu = "\t s. t \ U \ t >\<^sub>t s" have "t \ set (args (?ff i)) \ t \ ?U_of i" for t i unfolding U_def by (cases "is_App (?ff i)", simp_all, metis (lifting) neq_iff size_in_args sub.cases sub_args tm.discI(2)) moreover have "\i. extf f (>\<^sub>t) (args (?ff (i + k))) (args (?ff (Suc i + k)))" using bad_same hd_eq_f unfolding inf_chain_def gt_same.simps by auto ultimately have "\i. extf f ?gtu (args (?ff (i + k))) (args (?ff (Suc i + k)))" using extf_mono_strong[of _ _ "(>\<^sub>t)" "\t s. t \ U \ t >\<^sub>t s"] unfolding U_def by blast hence "inf_chain (extf f ?gtu) (\i. args (?ff (i + k)))" unfolding inf_chain_def by blast hence nwf_ext: "\ wfP (\xs ys. extf f ?gtu ys xs)" unfolding wfP_iff_no_inf_chain by fast have gtu_le_gtg: "?gtu \ (>\<^sub>t\<^sub>g)" by (auto intro!: gr_u) have "wfP (\s t. ?gtu t s)" unfolding wfP_iff_no_inf_chain proof (intro notI, elim exE) fix f assume bad_f: "inf_chain ?gtu f" hence bad_f0: "bad ?gtu (f 0)" by (rule inf_chain_bad) have "f 0 \ U" using bad_f unfolding inf_chain_def by blast hence good_f0: "\ bad ?gtu (f 0)" using u_good bad_f inf_chain_bad inf_chain_subset[OF _ gtu_le_gtg] by blast show False using bad_f0 good_f0 by sat qed hence wf_ext: "wfP (\xs ys. extf f ?gtu ys xs)" by (rule ext_wf.wf[OF extf_wf, rule_format]) show False using nwf_ext wf_ext by blast qed let ?subst = "subst grounding_\" have "wfP (\s t. ?subst t >\<^sub>t\<^sub>g ?subst s)" by (rule wfP_app[OF ground_wfP]) hence "wfP (\s t. ?subst t >\<^sub>t ?subst s)" by (simp add: ground_grounding_\) thus ?thesis by (auto intro: wfP_subset gt_subst[OF wary_grounding_\]) qed end end