diff --git a/thys/Functional_Ordered_Resolution_Prover/Deterministic_FO_Ordered_Resolution_Prover.thy b/thys/Functional_Ordered_Resolution_Prover/Deterministic_FO_Ordered_Resolution_Prover.thy --- a/thys/Functional_Ordered_Resolution_Prover/Deterministic_FO_Ordered_Resolution_Prover.thy +++ b/thys/Functional_Ordered_Resolution_Prover/Deterministic_FO_Ordered_Resolution_Prover.thy @@ -1,1779 +1,1778 @@ (* Title: A Deterministic Ordered Resolution Prover for First-Order Clauses Author: Jasmin Blanchette , 2017 Author: Anders Schlichtkrull , 2018 Maintainer: Anders Schlichtkrull *) section \A Deterministic Ordered Resolution Prover for First-Order Clauses\ text \ The \deterministic_RP\ prover introduced below is a deterministic program that works on finite lists, committing to a strategy for assigning priorities to clauses. However, it is not fully executable: It abstracts over operations on atoms and employs logical specifications instead of executable functions for auxiliary notions. \ theory Deterministic_FO_Ordered_Resolution_Prover imports Polynomial_Factorization.Missing_List Weighted_FO_Ordered_Resolution_Prover begin subsection \Library\ lemma apfst_fst_snd: "apfst f x = (f (fst x), snd x)" by (rule apfst_conv[of _ "fst x" "snd x" for x, unfolded prod.collapse]) lemma apfst_comp_rpair_const: "apfst f \ (\x. (x, y)) = (\x. (x, y)) \ f" by (simp add: comp_def) (* TODO: Move to Isabelle's "List.thy"? *) lemma length_remove1_less[termination_simp]: "x \ set xs \ length (remove1 x xs) < length xs" by (induct xs) auto (* TODO: Move to "Multiset_More.thy". *) lemma subset_mset_imp_subset_add_mset: "A \# B \ A \# add_mset x B" by (metis add_mset_diff_bothsides diff_subset_eq_self multiset_inter_def subset_mset.inf.absorb2) (* TODO: Move to "Multiset_More.thy"? *) lemma subseq_mset_subseteq_mset: "subseq xs ys \ mset xs \# mset ys" proof (induct xs arbitrary: ys) case (Cons x xs) note Outer_Cons = this then show ?case proof (induct ys) case (Cons y ys) have "subseq xs ys" by (metis Cons.prems(2) subseq_Cons' subseq_Cons2_iff) then show ?case using Cons by (metis mset.simps(2) mset_subset_eq_add_mset_cancel subseq_Cons2_iff subset_mset_imp_subset_add_mset) qed simp qed simp lemma map_filter_neq_eq_filter_map: "map f (filter (\y. f x \ f y) xs) = filter (\z. f x \ z) (map f xs)" by (induct xs) auto lemma mset_map_remdups_gen: "mset (map f (remdups_gen f xs)) = mset (remdups_gen (\x. x) (map f xs))" by (induct f xs rule: remdups_gen.induct) (auto simp: map_filter_neq_eq_filter_map) lemma mset_remdups_gen_ident: "mset (remdups_gen (\x. x) xs) = mset_set (set xs)" proof - have "f = (\x. x) \ mset (remdups_gen f xs) = mset_set (set xs)" for f proof (induct f xs rule: remdups_gen.induct) case (2 f x xs) note ih = this(1) and f = this(2) show ?case unfolding f remdups_gen.simps ih[OF f, unfolded f] mset.simps by (metis finite_set list.simps(15) mset_set.insert_remove removeAll_filter_not_eq remove_code(1) remove_def) qed simp then show ?thesis by simp qed (* FIXME: This is a clone of "Lambda_Free_RPOs". *) lemma wf_app: "wf r \ wf {(x, y). (f x, f y) \ r}" unfolding wf_eq_minimal by (intro allI, drule spec[of _ "f ` Q" for Q]) fast (* FIXME: This is a clone of "Lambda_Free_RPOs". *) lemma wfP_app: "wfP p \ wfP (\x y. p (f x) (f y))" unfolding wfP_def by (rule wf_app[of "{(x, y). p x y}" f, simplified]) (* TODO: Move to Isabelle? *) lemma funpow_fixpoint: "f x = x \ (f ^^ n) x = x" by (induct n) auto lemma rtranclp_imp_eq_image: "(\x y. R x y \ f x = f y) \ R\<^sup>*\<^sup>* x y \ f x = f y" by (erule rtranclp.induct) auto lemma tranclp_imp_eq_image: "(\x y. R x y \ f x = f y) \ R\<^sup>+\<^sup>+ x y \ f x = f y" by (erule tranclp.induct) auto subsection \Prover\ type_synonym 'a lclause = "'a literal list" type_synonym 'a dclause = "'a lclause \ nat" type_synonym 'a dstate = "'a dclause list \ 'a dclause list \ 'a dclause list \ nat" locale deterministic_FO_resolution_prover = weighted_FO_resolution_prover_with_size_timestamp_factors S subst_atm id_subst comp_subst renamings_apart atm_of_atms mgu less_atm size_atm timestamp_factor size_factor for S :: "('a :: wellorder) clause \ 'a clause" and subst_atm :: "'a \ 's \ 'a" and id_subst :: "'s" and comp_subst :: "'s \ 's \ 's" and renamings_apart :: "'a literal multiset list \ 's list" and atm_of_atms :: "'a list \ 'a" and mgu :: "'a set set \ 's option" and less_atm :: "'a \ 'a \ bool" and size_atm :: "'a \ nat" and timestamp_factor :: nat and size_factor :: nat + assumes S_empty: "S C = {#}" begin lemma less_atm_irrefl: "\ less_atm A A" using ex_ground_subst less_atm_ground less_atm_stable unfolding is_ground_subst_def by blast fun wstate_of_dstate :: "'a dstate \ 'a wstate" where "wstate_of_dstate (N, P, Q, n) = (mset (map (apfst mset) N), mset (map (apfst mset) P), mset (map (apfst mset) Q), n)" fun state_of_dstate :: "'a dstate \ 'a state" where "state_of_dstate (N, P, Q, _) = (set (map (mset \ fst) N), set (map (mset \ fst) P), set (map (mset \ fst) Q))" abbreviation clss_of_dstate :: "'a dstate \ 'a clause set" where "clss_of_dstate St \ clss_of_state (state_of_dstate St)" fun is_final_dstate :: "'a dstate \ bool" where "is_final_dstate (N, P, Q, n) \ N = [] \ P = []" declare is_final_dstate.simps [simp del] abbreviation rtrancl_weighted_RP (infix "\\<^sub>w\<^sup>*" 50) where "(\\<^sub>w\<^sup>*) \ (\\<^sub>w)\<^sup>*\<^sup>*" abbreviation trancl_weighted_RP (infix "\\<^sub>w\<^sup>+" 50) where "(\\<^sub>w\<^sup>+) \ (\\<^sub>w)\<^sup>+\<^sup>+" definition is_tautology :: "'a lclause \ bool" where "is_tautology C \ (\A \ set (map atm_of C). Pos A \ set C \ Neg A \ set C)" definition subsume :: "'a lclause list \ 'a lclause \ bool" where "subsume Ds C \ (\D \ set Ds. subsumes (mset D) (mset C))" definition strictly_subsume :: "'a lclause list \ 'a lclause \ bool" where "strictly_subsume Ds C \ (\D \ set Ds. strictly_subsumes (mset D) (mset C))" definition is_reducible_on :: "'a literal \ 'a lclause \ 'a literal \ 'a lclause \ bool" where "is_reducible_on M D L C \ subsumes (mset D + {#- M#}) (mset C + {#L#})" definition is_reducible_lit :: "'a lclause list \ 'a lclause \ 'a literal \ bool" where "is_reducible_lit Ds C L \ (\D \ set Ds. \L' \ set D. \\. - L = L' \l \ \ mset (remove1 L' D) \ \ \# mset C)" primrec reduce :: "'a lclause list \ 'a lclause \ 'a lclause \ 'a lclause" where "reduce _ _ [] = []" | "reduce Ds C (L # C') = (if is_reducible_lit Ds (C @ C') L then reduce Ds C C' else L # reduce Ds (L # C) C')" abbreviation is_irreducible :: "'a lclause list \ 'a lclause \ bool" where "is_irreducible Ds C \ reduce Ds [] C = C" abbreviation is_reducible :: "'a lclause list \ 'a lclause \ bool" where "is_reducible Ds C \ reduce Ds [] C \ C" definition reduce_all :: "'a lclause \ 'a dclause list \ 'a dclause list" where "reduce_all D = map (apfst (reduce [D] []))" fun reduce_all2 :: "'a lclause \ 'a dclause list \ 'a dclause list \ 'a dclause list" where "reduce_all2 _ [] = ([], [])" | "reduce_all2 D (Ci # Cs) = (let (C, i) = Ci; C' = reduce [D] [] C in (if C' = C then apsnd else apfst) (Cons (C', i)) (reduce_all2 D Cs))" fun remove_all :: "'b list \ 'b list \ 'b list" where "remove_all xs [] = xs" | "remove_all xs (y # ys) = (if y \ set xs then remove_all (remove1 y xs) ys else remove_all xs ys)" lemma remove_all_mset_minus: "mset ys \# mset xs \ mset (remove_all xs ys) = mset xs - mset ys" proof (induction ys arbitrary: xs) case (Cons y ys) show ?case proof (cases "y \ set xs") case y_in: True then have subs: "mset ys \# mset (remove1 y xs)" using Cons(2) by (simp add: insert_subset_eq_iff) show ?thesis using y_in Cons subs by auto next case False then show ?thesis using Cons by auto qed qed auto definition resolvent :: "'a lclause \ 'a \'a lclause \ 'a lclause \ 'a lclause" where "resolvent D A CA Ls = map (\M. M \l (the (mgu {insert A (atms_of (mset Ls))}))) (remove_all CA Ls @ D)" definition resolvable :: "'a \ 'a lclause \ 'a lclause \ 'a lclause \ bool" where "resolvable A D CA Ls \ (let \ = (mgu {insert A (atms_of (mset Ls))}) in \ \ None \ Ls \ [] \ maximal_wrt (A \a the \) ((add_mset (Neg A) (mset D)) \ the \) \ strictly_maximal_wrt (A \a the \) ((mset CA - mset Ls) \ the \) \ (\L \ set Ls. is_pos L))" definition resolve_on :: "'a \ 'a lclause \ 'a lclause \ 'a lclause list" where "resolve_on A D CA = map (resolvent D A CA) (filter (resolvable A D CA) (subseqs CA))" definition resolve :: "'a lclause \ 'a lclause \ 'a lclause list" where "resolve C D = concat (map (\L. (case L of Pos A \ [] | Neg A \ if maximal_wrt A (mset D) then resolve_on A (remove1 L D) C else [])) D)" definition resolve_rename :: "'a lclause \ 'a lclause \ 'a lclause list" where "resolve_rename C D = (let \s = renamings_apart [mset D, mset C] in resolve (map (\L. L \l last \s) C) (map (\L. L \l hd \s) D))" definition resolve_rename_either_way :: "'a lclause \ 'a lclause \ 'a lclause list" where "resolve_rename_either_way C D = resolve_rename C D @ resolve_rename D C" fun select_min_weight_clause :: "'a dclause \ 'a dclause list \ 'a dclause" where "select_min_weight_clause Ci [] = Ci" | "select_min_weight_clause Ci (Dj # Djs) = select_min_weight_clause (if weight (apfst mset Dj) < weight (apfst mset Ci) then Dj else Ci) Djs" lemma select_min_weight_clause_in: "select_min_weight_clause P0 P \ set (P0 # P)" by (induct P arbitrary: P0) auto function remdups_clss :: "'a dclause list \ 'a dclause list" where "remdups_clss [] = []" | "remdups_clss (Ci # Cis) = (let Ci' = select_min_weight_clause Ci Cis in Ci' # remdups_clss (filter (\(D, _). mset D \ mset (fst Ci')) (Ci # Cis)))" by pat_completeness auto termination apply (relation "measure length") apply (rule wf_measure) by (metis (mono_tags) in_measure length_filter_less prod.case_eq_if select_min_weight_clause_in) declare remdups_clss.simps(2) [simp del] fun deterministic_RP_step :: "'a dstate \ 'a dstate" where "deterministic_RP_step (N, P, Q, n) = (if \Ci \ set (P @ Q). fst Ci = [] then ([], [], remdups_clss P @ Q, n + length (remdups_clss P)) else (case N of [] \ (case P of [] \ (N, P, Q, n) | P0 # P' \ let (C, i) = select_min_weight_clause P0 P'; N = map (\D. (D, n)) (remdups_gen mset (resolve_rename C C @ concat (map (resolve_rename_either_way C \ fst) Q))); P = filter (\(D, j). mset D \ mset C) P; Q = (C, i) # Q; n = Suc n in (N, P, Q, n)) | (C, i) # N \ let C = reduce (map fst (P @ Q)) [] C in if C = [] then ([], [], [([], i)], Suc n) else if is_tautology C \ subsume (map fst (P @ Q)) C then (N, P, Q, n) else let P = reduce_all C P; (back_to_P, Q) = reduce_all2 C Q; P = back_to_P @ P; Q = filter (Not \ strictly_subsume [C] \ fst) Q; P = filter (Not \ strictly_subsume [C] \ fst) P; P = (C, i) # P in (N, P, Q, n)))" declare deterministic_RP_step.simps [simp del] partial_function (option) deterministic_RP :: "'a dstate \ 'a lclause list option" where "deterministic_RP St = (if is_final_dstate St then let (_, _, Q, _) = St in Some (map fst Q) else deterministic_RP (deterministic_RP_step St))" lemma is_final_dstate_imp_not_weighted_RP: "is_final_dstate St \ \ wstate_of_dstate St \\<^sub>w St'" using wrp.final_weighted_RP by (cases St) (auto intro: wrp.final_weighted_RP simp: is_final_dstate.simps) lemma is_final_dstate_funpow_imp_deterministic_RP_neq_None: "is_final_dstate ((deterministic_RP_step ^^ k) St) \ deterministic_RP St \ None" proof (induct k arbitrary: St) case (Suc k) note ih = this(1) and final_Sk = this(2)[simplified, unfolded funpow_swap1] show ?case using ih[OF final_Sk] by (subst deterministic_RP.simps) (simp add: prod.case_eq_if) qed (subst deterministic_RP.simps, simp add: prod.case_eq_if) lemma is_reducible_lit_mono_cls: "mset C \# mset C' \ is_reducible_lit Ds C L \ is_reducible_lit Ds C' L" unfolding is_reducible_lit_def by (blast intro: subset_mset.order.trans) lemma is_reducible_lit_mset_iff: "mset C = mset C' \ is_reducible_lit Ds C' L \ is_reducible_lit Ds C L" by (metis is_reducible_lit_mono_cls subset_mset.order_refl) lemma is_reducible_lit_remove1_Cons_iff: assumes "L \ set C'" shows "is_reducible_lit Ds (C @ remove1 L (M # C')) L \ is_reducible_lit Ds (M # C @ remove1 L C') L" using assms by (subst is_reducible_lit_mset_iff, auto) lemma reduce_mset_eq: "mset C = mset C' \ reduce Ds C E = reduce Ds C' E" proof (induct E arbitrary: C C') case (Cons L E) note ih = this(1) and mset_eq = this(2) have mset_lc_eq: "mset (L # C) = mset (L # C')" and mset_ce_eq: "mset (C @ E) = mset (C' @ E)" using mset_eq by simp+ show ?case using ih[OF mset_eq] ih[OF mset_lc_eq] by (simp add: is_reducible_lit_mset_iff[OF mset_ce_eq]) qed simp lemma reduce_rotate[simp]: "reduce Ds (C @ [L]) E = reduce Ds (L # C) E" by (rule reduce_mset_eq) simp lemma mset_reduce_subset: "mset (reduce Ds C E) \# mset E" by (induct E arbitrary: C) (auto intro: subset_mset_imp_subset_add_mset) lemma reduce_idem: "reduce Ds C (reduce Ds C E) = reduce Ds C E" by (induct E arbitrary: C) (auto intro!: mset_reduce_subset dest!: is_reducible_lit_mono_cls[of "C @ reduce Ds (L # C) E" "C @ E" Ds L for L E C, rotated]) lemma is_reducible_lit_imp_is_reducible: "L \ set C' \ is_reducible_lit Ds (C @ remove1 L C') L \ reduce Ds C C' \ C'" proof (induct C' arbitrary: C) case (Cons M C') note ih = this(1) and l_in = this(2) and l_red = this(3) show ?case proof (cases "is_reducible_lit Ds (C @ C') M") case True then show ?thesis by simp (metis mset.simps(2) mset_reduce_subset multi_self_add_other_not_self subset_mset.eq_iff subset_mset_imp_subset_add_mset) next case m_irred: False have "L \ set C'" and "is_reducible_lit Ds (M # C @ remove1 L C') L" using l_in l_red m_irred is_reducible_lit_remove1_Cons_iff by auto then show ?thesis by (simp add: ih[of "M # C"] m_irred) qed qed simp lemma is_reducible_imp_is_reducible_lit: "reduce Ds C C' \ C' \ \L \ set C'. is_reducible_lit Ds (C @ remove1 L C') L" proof (induct C' arbitrary: C) case (Cons M C') note ih = this(1) and mc'_red = this(2) show ?case proof (cases "is_reducible_lit Ds (C @ C') M") case m_irred: False show ?thesis using ih[of "M # C"] mc'_red[simplified, simplified m_irred, simplified] m_irred is_reducible_lit_remove1_Cons_iff by auto qed simp qed simp lemma is_irreducible_iff_nexists_is_reducible_lit: "reduce Ds C C' = C' \ \ (\L \ set C'. is_reducible_lit Ds (C @ remove1 L C') L)" using is_reducible_imp_is_reducible_lit is_reducible_lit_imp_is_reducible by blast lemma is_irreducible_mset_iff: "mset E = mset E' \ reduce Ds C E = E \ reduce Ds C E' = E'" unfolding is_irreducible_iff_nexists_is_reducible_lit by (metis (full_types) is_reducible_lit_mset_iff mset_remove1 set_mset_mset union_code) lemma select_min_weight_clause_min_weight: assumes "Ci = select_min_weight_clause P0 P" shows "weight (apfst mset Ci) = Min ((weight \ apfst mset) ` set (P0 # P))" using assms proof (induct P arbitrary: P0 Ci) case (Cons P1 P) note ih = this(1) and ci = this(2) show ?case proof (cases "weight (apfst mset P1) < weight (apfst mset P0)") case True then have min: "Min ((weight \ apfst mset) ` set (P0 # P1 # P)) = Min ((weight \ apfst mset) ` set (P1 # P))" by (simp add: min_def) show ?thesis unfolding min by (rule ih[of Ci P1]) (simp add: ih[of Ci P1] ci True) next case False have "Min ((weight \ apfst mset) ` set (P0 # P1 # P)) = Min ((weight \ apfst mset) ` set (P1 # P0 # P))" by (rule arg_cong[of _ _ Min]) auto then have min: "Min ((weight \ apfst mset) ` set (P0 # P1 # P)) = Min ((weight \ apfst mset) ` set (P0 # P))" by (simp add: min_def) (use False eq_iff in fastforce) show ?thesis unfolding min by (rule ih[of Ci P0]) (simp add: ih[of Ci P1] ci False) qed qed simp lemma remdups_clss_Nil_iff: "remdups_clss Cs = [] \ Cs = []" by (cases Cs, simp, hypsubst, subst remdups_clss.simps(2), simp add: Let_def) lemma empty_N_if_Nil_in_P_or_Q: assumes nil_in: "[] \ fst ` set (P @ Q)" shows "wstate_of_dstate (N, P, Q, n) \\<^sub>w\<^sup>* wstate_of_dstate ([], P, Q, n)" proof (induct N) case ih: (Cons N0 N) have "wstate_of_dstate (N0 # N, P, Q, n) \\<^sub>w wstate_of_dstate (N, P, Q, n)" by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w)", OF _ _ wrp.forward_subsumption[of "{#}" "mset (map (apfst mset) P)" "mset (map (apfst mset) Q)" "mset (fst N0)" "mset (map (apfst mset) N)" "snd N0" n]]) (use nil_in in \force simp: image_def apfst_fst_snd\)+ then show ?case using ih by (rule converse_rtranclp_into_rtranclp) qed simp lemma remove_strictly_subsumed_clauses_in_P: assumes c_in: "C \ fst ` set N" and p_nsubs: "\D \ fst ` set P. \ strictly_subsume [C] D" shows "wstate_of_dstate (N, P @ P', Q, n) \\<^sub>w\<^sup>* wstate_of_dstate (N, P @ filter (Not \ strictly_subsume [C] \ fst) P', Q, n)" using p_nsubs proof (induct "length P'" arbitrary: P P' rule: less_induct) case less note ih = this(1) and p_nsubs = this(2) show ?case proof (cases "length P'") case Suc let ?Dj = "hd P'" let ?P'' = "tl P'" have p': "P' = hd P' # tl P'" using Suc by (metis length_Suc_conv list.distinct(1) list.exhaust_sel) show ?thesis proof (cases "strictly_subsume [C] (fst ?Dj)") case subs: True have p_filtered: "{#(E, k) \# image_mset (apfst mset) (mset P). E \ mset (fst ?Dj)#} = image_mset (apfst mset) (mset P)" by (rule filter_mset_cong[OF refl, of _ _ "\_. True", simplified], use subs p_nsubs in \auto simp: strictly_subsume_def\) have "{#(E, k) \# image_mset (apfst mset) (mset P'). E \ mset (fst ?Dj)#} = {#(E, k) \# image_mset (apfst mset) (mset ?P''). E \ mset (fst ?Dj)#}" by (subst (2) p') (simp add: case_prod_beta) also have "\ = image_mset (apfst mset) (mset (filter (\(E, l). mset E \ mset (fst ?Dj)) ?P''))" by (auto simp: image_mset_filter_swap[symmetric] mset_filter case_prod_beta) finally have p'_filtered: "{#(E, k) \# image_mset (apfst mset) (mset P'). E \ mset (fst ?Dj)#} = image_mset (apfst mset) (mset (filter (\(E, l). mset E \ mset (fst ?Dj)) ?P''))" . have "wstate_of_dstate (N, P @ P', Q, n) \\<^sub>w wstate_of_dstate (N, P @ filter (\(E, l). mset E \ mset (fst ?Dj)) ?P'', Q, n)" by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w)", OF _ _ wrp.backward_subsumption_P[of "mset C" "mset (map (apfst mset) N)" "mset (fst ?Dj)" "mset (map (apfst mset) (P @ P'))" "mset (map (apfst mset) Q)" n]], use c_in subs in \auto simp add: p_filtered p'_filtered arg_cong[OF p', of set] strictly_subsume_def\) also have "\ \\<^sub>w\<^sup>* wstate_of_dstate (N, P @ filter (Not \ strictly_subsume [C] \ fst) P', Q, n)" apply (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w\<^sup>*)", OF _ _ ih[of "filter (\(E, l). mset E \ mset (fst ?Dj)) ?P''" P]]) apply simp_all apply (subst (3) p') using subs apply (simp add: case_prod_beta) apply (rule arg_cong[of _ _ "\x. image_mset (apfst mset) x"]) apply (metis (no_types, hide_lams) strictly_subsume_def) apply (subst (3) p') apply (subst list.size) apply (metis (no_types, lifting) less_Suc0 less_add_same_cancel1 linorder_neqE_nat not_add_less1 sum_length_filter_compl trans_less_add1) using p_nsubs by fast ultimately show ?thesis by (rule converse_rtranclp_into_rtranclp) next case nsubs: False show ?thesis apply (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w\<^sup>*)", OF _ _ ih[of ?P'' "P @ [?Dj]"]]) using nsubs p_nsubs apply (simp_all add: arg_cong[OF p', of mset] arg_cong[OF p', of "filter f" for f]) apply (subst (1 2) p') by simp qed qed simp qed lemma remove_strictly_subsumed_clauses_in_Q: assumes c_in: "C \ fst ` set N" shows "wstate_of_dstate (N, P, Q @ Q', n) \\<^sub>w\<^sup>* wstate_of_dstate (N, P, Q @ filter (Not \ strictly_subsume [C] \ fst) Q', n)" proof (induct Q' arbitrary: Q) case ih: (Cons Dj Q') have "wstate_of_dstate (N, P, Q @ Dj # Q', n) \\<^sub>w\<^sup>* wstate_of_dstate (N, P, Q @ filter (Not \ strictly_subsume [C] \ fst) [Dj] @ Q', n)" proof (cases "strictly_subsume [C] (fst Dj)") case subs: True have "wstate_of_dstate (N, P, Q @ Dj # Q', n) \\<^sub>w wstate_of_dstate (N, P, Q @ Q', n)" by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w)", OF _ _ wrp.backward_subsumption_Q[of "mset C" "mset (map (apfst mset) N)" "mset (fst Dj)" "mset (map (apfst mset) P)" "mset (map (apfst mset) (Q @ Q'))" "snd Dj" n]]) (use c_in subs in \auto simp: apfst_fst_snd strictly_subsume_def\) then show ?thesis by auto qed simp then show ?case using ih[of "Q @ filter (Not \ strictly_subsume [C] \ fst) [Dj]"] by force qed simp lemma reduce_clause_in_P: assumes c_in: "C \ fst ` set N" and p_irred: "\(E, k) \ set (P @ P'). k > j \ is_irreducible [C] E" shows "wstate_of_dstate (N, P @ (D @ D', j) # P', Q, n) \\<^sub>w\<^sup>* wstate_of_dstate (N, P @ (D @ reduce [C] D D', j) # P', Q, n)" proof (induct D' arbitrary: D) case ih: (Cons L D') show ?case proof (cases "is_reducible_lit [C] (D @ D') L") case l_red: True then obtain L' :: "'a literal" and \ :: 's where l'_in: "L' \ set C" and not_l: "- L = L' \l \" and subs: "mset (remove1 L' C) \ \ \# mset (D @ D')" unfolding is_reducible_lit_def by force have ldd'_red: "is_reducible [C] (L # D @ D')" apply (rule is_reducible_lit_imp_is_reducible) using l_red by auto have lt_imp_neq: "\(E, k) \ set (P @ P'). j < k \ mset E \ mset (L # D @ D')" using p_irred ldd'_red is_irreducible_mset_iff by fast have "wstate_of_dstate (N, P @ (D @ L # D', j) # P', Q, n) \\<^sub>w wstate_of_dstate (N, P @ (D @ D', j) # P', Q, n)" apply (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w)", OF _ _ wrp.backward_reduction_P[of "mset C - {#L'#}" L' "mset (map (apfst mset) N)" L \ "mset (D @ D')" "mset (map (apfst mset) (P @ P'))" j "mset (map (apfst mset) Q)" n]]) using l'_in not_l subs c_in lt_imp_neq by (simp_all add: case_prod_beta) force+ then show ?thesis using ih[of D] l_red by simp next case False then show ?thesis using ih[of "D @ [L]"] by simp qed qed simp lemma reduce_clause_in_Q: assumes c_in: "C \ fst ` set N" and p_irred: "\(E, k) \ set P. k > j \ is_irreducible [C] E" and d'_red: "reduce [C] D D' \ D'" shows "wstate_of_dstate (N, P, Q @ (D @ D', j) # Q', n) \\<^sub>w\<^sup>* wstate_of_dstate (N, (D @ reduce [C] D D', j) # P, Q @ Q', n)" using d'_red proof (induct D' arbitrary: D) case (Cons L D') note ih = this(1) and ld'_red = this(2) then show ?case proof (cases "is_reducible_lit [C] (D @ D') L") case l_red: True then obtain L' :: "'a literal" and \ :: 's where l'_in: "L' \ set C" and not_l: "- L = L' \l \" and subs: "mset (remove1 L' C) \ \ \# mset (D @ D')" unfolding is_reducible_lit_def by force have "wstate_of_dstate (N, P, Q @ (D @ L # D', j) # Q', n) \\<^sub>w wstate_of_dstate (N, (D @ D', j) # P, Q @ Q', n)" by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w)", OF _ _ wrp.backward_reduction_Q[of "mset C - {#L'#}" L' "mset (map (apfst mset) N)" L \ "mset (D @ D')" "mset (map (apfst mset) P)" "mset (map (apfst mset) (Q @ Q'))" j n]], use l'_in not_l subs c_in in auto) then show ?thesis using l_red p_irred reduce_clause_in_P[OF c_in, of "[]" P j D D' "Q @ Q'" n] by simp next case l_nred: False then have d'_red: "reduce [C] (D @ [L]) D' \ D'" using ld'_red by simp show ?thesis using ih[OF d'_red] l_nred by simp qed qed simp lemma reduce_clauses_in_P: assumes c_in: "C \ fst ` set N" and p_irred: "\(E, k) \ set P. is_irreducible [C] E" shows "wstate_of_dstate (N, P @ P', Q, n) \\<^sub>w\<^sup>* wstate_of_dstate (N, P @ reduce_all C P', Q, n)" unfolding reduce_all_def using p_irred proof (induct "length P'" arbitrary: P P') case (Suc l) note ih = this(1) and suc_l = this(2) and p_irred = this(3) have p'_nnil: "P' \ []" using suc_l by auto define j :: nat where "j = Max (snd ` set P')" obtain Dj :: "'a dclause" where dj_in: "Dj \ set P'" and snd_dj: "snd Dj = j" using Max_in[of "snd ` set P'", unfolded image_def, simplified] by (metis image_def j_def length_Suc_conv list.set_intros(1) suc_l) have "\k \ snd ` set P'. k \ j" unfolding j_def using p'_nnil by simp then have j_max: "\(E, k) \ set P'. j \ k" unfolding image_def by fastforce obtain P1' P2' :: "'a dclause list" where p': "P' = P1' @ Dj # P2'" using split_list[OF dj_in] by blast have "wstate_of_dstate (N, P @ P1' @ Dj # P2', Q, n) \\<^sub>w\<^sup>* wstate_of_dstate (N, P @ P1' @ apfst (reduce [C] []) Dj # P2', Q, n)" unfolding append_assoc[symmetric] apply (subst (1 2) surjective_pairing[of Dj, unfolded snd_dj]) apply (simp only: apfst_conv) apply (rule reduce_clause_in_P[of _ _ _ _ _ "[]", unfolded append_Nil, OF c_in]) using p_irred j_max[unfolded p'] by (force simp: case_prod_beta) moreover have "wstate_of_dstate (N, P @ P1' @ apfst (reduce [C] []) Dj # P2', Q, n) \\<^sub>w\<^sup>* wstate_of_dstate (N, P @ map (apfst (reduce [C] [])) (P1' @ Dj # P2'), Q, n)" apply (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w\<^sup>*)", OF _ _ ih[of "P1' @ P2'" "apfst (reduce [C] []) Dj # P"]]) using suc_l reduce_idem p_irred unfolding p' by (auto simp: case_prod_beta) ultimately show ?case unfolding p' by simp qed simp lemma reduce_clauses_in_Q: assumes c_in: "C \ fst ` set N" and p_irred: "\(E, k) \ set P. is_irreducible [C] E" shows "wstate_of_dstate (N, P, Q @ Q', n) \\<^sub>w\<^sup>* wstate_of_dstate (N, fst (reduce_all2 C Q') @ P, Q @ snd (reduce_all2 C Q'), n)" using p_irred proof (induct Q' arbitrary: P Q) case (Cons Dj Q') note ih = this(1) and p_irred = this(2) show ?case proof (cases "is_irreducible [C] (fst Dj)") case True then show ?thesis using ih[of _ "Q @ [Dj]"] p_irred by (simp add: case_prod_beta) next case d_red: False have "wstate_of_dstate (N, P, Q @ Dj # Q', n) \\<^sub>w\<^sup>* wstate_of_dstate (N, (reduce [C] [] (fst Dj), snd Dj) # P, Q @ Q', n)" using p_irred reduce_clause_in_Q[of _ _ P "snd Dj" "[]" _ Q Q' n, OF c_in _ d_red] by (cases Dj) force then show ?thesis using ih[of "(reduce [C] [] (fst Dj), snd Dj) # P" Q] d_red p_irred reduce_idem by (force simp: case_prod_beta) qed qed simp lemma eligible_iff: "eligible S \ As DA \ As = [] \ length As = 1 \ maximal_wrt (hd As \a \) (DA \ \)" unfolding eligible.simps S_empty by (fastforce dest: hd_conv_nth) lemma ord_resolve_one_side_prem: "ord_resolve S CAs DA AAs As \ E \ length CAs = 1 \ length AAs = 1 \ length As = 1" by (force elim!: ord_resolve.cases simp: eligible_iff) lemma ord_resolve_rename_one_side_prem: "ord_resolve_rename S CAs DA AAs As \ E \ length CAs = 1 \ length AAs = 1 \ length As = 1" by (force elim!: ord_resolve_rename.cases dest: ord_resolve_one_side_prem) abbreviation Bin_ord_resolve :: "'a clause \ 'a clause \ 'a clause set" where "Bin_ord_resolve C D \ {E. \AA A \. ord_resolve S [C] D [AA] [A] \ E}" abbreviation Bin_ord_resolve_rename :: "'a clause \ 'a clause \ 'a clause set" where "Bin_ord_resolve_rename C D \ {E. \AA A \. ord_resolve_rename S [C] D [AA] [A] \ E}" lemma resolve_on_eq_UNION_Bin_ord_resolve: "mset ` set (resolve_on A D CA) = {E. \AA \. ord_resolve S [mset CA] ({#Neg A#} + mset D) [AA] [A] \ E}" proof { fix E :: "'a literal list" assume "E \ set (resolve_on A D CA)" then have "E \ resolvent D A CA ` {Ls. subseq Ls CA \ resolvable A D CA Ls}" unfolding resolve_on_def by simp then obtain Ls where Ls_p: "resolvent D A CA Ls = E" "subseq Ls CA \ resolvable A D CA Ls" by auto define \ where "\ = the (mgu {insert A (atms_of (mset Ls))})" then have \_p: "mgu {insert A (atms_of (mset Ls))} = Some \" "Ls \ []" "eligible S \ [A] (add_mset (Neg A) (mset D))" "strictly_maximal_wrt (A \a \) ((mset CA - mset Ls) \ \)" "\L \ set Ls. is_pos L" using Ls_p unfolding resolvable_def unfolding Let_def eligible.simps using S_empty by auto from \_p have \_p2: "the (mgu {insert A (atms_of (mset Ls))}) = \" by auto have Ls_sub_CA: "mset Ls \# mset CA" using subseq_mset_subseteq_mset Ls_p by auto then have "mset (resolvent D A CA Ls) = sum_list [mset CA - mset Ls] \ \ + mset D \ \" unfolding resolvent_def \_p2 subst_cls_def using remove_all_mset_minus[of Ls CA] by auto moreover have "length [mset CA - mset Ls] = Suc 0" by auto moreover have "\L \ set Ls. is_pos L" using \_p(5) list_all_iff[of is_pos] by auto then have "{#Pos (atm_of x). x \# mset Ls#} = mset Ls" by (induction Ls) auto then have "mset CA = [mset CA - mset Ls] ! 0 + {#Pos (atm_of x). x \# mset Ls#}" using Ls_sub_CA by auto moreover have "Ls \ []" using \_p by - moreover have "Some \ = mgu {insert A (atm_of ` set Ls)}" using \_p unfolding atms_of_def by auto moreover have "eligible S \ [A] (add_mset (Neg A) (mset D))" using \_p by - moreover have "strictly_maximal_wrt (A \a \) ([mset CA - mset Ls] ! 0 \ \)" using \_p(4) by auto moreover have "S (mset CA) = {#}" by (simp add: S_empty) ultimately have "\Cs. mset (resolvent D A CA Ls) = sum_list Cs \ \ + mset D \ \ \ length Cs = Suc 0 \ mset CA = Cs ! 0 + {#Pos (atm_of x). x \# mset Ls#} \ Ls \ [] \ Some \ = mgu {insert A (atm_of ` set Ls)} \ eligible S \ [A] (add_mset (Neg A) (mset D)) \ strictly_maximal_wrt (A \a \) (Cs ! 0 \ \) \ S (mset CA) = {#}" by blast then have "ord_resolve S [mset CA] (add_mset (Neg A) (mset D)) [image_mset atm_of (mset Ls)] [A] \ (mset (resolvent D A CA Ls))" unfolding ord_resolve.simps by auto then have "\AA \. ord_resolve S [mset CA] (add_mset (Neg A) (mset D)) [AA] [A] \ (mset E)" using Ls_p by auto } then show "mset ` set (resolve_on A D CA) \ {E. \AA \. ord_resolve S [mset CA] ({#Neg A#} + mset D) [AA] [A] \ E}" by auto next { fix E AA \ assume "ord_resolve S [mset CA] (add_mset (Neg A) (mset D)) [AA] [A] \ E" then obtain Cs where res': "E = sum_list Cs \ \ + mset D \ \" "length Cs = Suc 0" "mset CA = Cs ! 0 + poss AA" "AA \ {#}" "Some \ = mgu {insert A (set_mset AA)}" "eligible S \ [A] (add_mset (Neg A) (mset D))" "strictly_maximal_wrt (A \a \) (Cs ! 0 \ \)" "S (Cs ! 0 + poss AA) = {#}" unfolding ord_resolve.simps by auto moreover define C where "C = Cs ! 0" ultimately have res: "E = sum_list Cs \ \ + mset D \ \" "mset CA = C + poss AA" "AA \ {#}" "Some \ = mgu {insert A (set_mset AA)}" "eligible S \ [A] (add_mset (Neg A) (mset D))" "strictly_maximal_wrt (A \a \) (C \ \)" "S (C + poss AA) = {#}" unfolding ord_resolve.simps by auto from this(1) have "E = C \ \ + mset D \ \" unfolding C_def using res'(2) by (cases Cs) auto note res' = this res(2-7) have "\Al. mset Al = AA \ subseq (map Pos Al) CA" using res(2) proof (induction CA arbitrary: AA C) case Nil then show ?case by auto next case (Cons L CA) then show ?case proof (cases "L \# poss AA ") case True then have pos_L: "is_pos L" by auto have rem: "\A'. Pos A' \# poss AA \ remove1_mset (Pos A') (C + poss AA) = C + poss (remove1_mset A' AA)" by (induct AA) auto have "mset CA = C + (poss (AA - {#atm_of L#}))" using True Cons(2) by (metis add_mset_remove_trivial rem literal.collapse(1) mset.simps(2) pos_L) then have "\Al. mset Al = remove1_mset (atm_of L) AA \ subseq (map Pos Al) CA" using Cons(1)[of _ "((AA - {#atm_of L#}))"] by metis then obtain Al where "mset Al = remove1_mset (atm_of L) AA \ subseq (map Pos Al) CA" by auto then have "mset (atm_of L # Al) = AA" and "subseq (map Pos (atm_of L # Al)) (L # CA)" using True by (auto simp add: pos_L) then show ?thesis by blast next case False then have "mset CA = remove1_mset L C + poss AA" using Cons(2) by (metis Un_iff add_mset_remove_trivial mset.simps(2) set_mset_union single_subset_iff subset_mset.add_diff_assoc2 union_single_eq_member) then have "\Al. mset Al = AA \ subseq (map Pos Al) CA" using Cons(1)[of "C - {#L#}" AA] Cons(2) by auto then show ?thesis by auto qed qed then obtain Al where Al_p: "mset Al = AA" "subseq (map Pos Al) CA" by auto define Ls :: "'a lclause" where "Ls = map Pos Al" have diff: "mset CA - mset Ls = C" unfolding Ls_def using res(2) Al_p(1) by auto have ls_subq_ca: "subseq Ls CA" unfolding Ls_def using Al_p by - moreover { have "\y. mgu {insert A (atms_of (mset Ls))} = Some y" unfolding Ls_def using res(4) Al_p by (metis atms_of_poss mset_map) moreover have "Ls \ []" using Al_p(1) Ls_def res'(3) by auto moreover have \_p: "the (mgu {insert A (set Al)}) = \" using res'(4) Al_p(1) by (metis option.sel set_mset_mset) then have "eligible S (the (mgu {insert A (atms_of (mset Ls))})) [A] (add_mset (Neg A) (mset D))" unfolding Ls_def using res by auto moreover have "strictly_maximal_wrt (A \a the (mgu {insert A (atms_of (mset Ls))})) ((mset CA - mset Ls) \ the (mgu {insert A (atms_of (mset Ls))}))" unfolding Ls_def using res \_p Al_p by auto moreover have "\L \ set Ls. is_pos L" by (simp add: Ls_def) ultimately have "resolvable A D CA Ls" unfolding resolvable_def unfolding eligible.simps using S_empty by simp } moreover have ls_sub_ca: "mset Ls \# mset CA" using ls_subq_ca subseq_mset_subseteq_mset[of Ls CA] by simp have "{#x \l \. x \# mset CA - mset Ls#} + {#M \l \. M \# mset D#} = C \ \ + mset D \ \" using diff unfolding subst_cls_def by simp then have "{#x \l \. x \# mset CA - mset Ls#} + {#M \l \. M \# mset D#} = E" using res'(1) by auto then have "{#M \l \. M \# mset (remove_all CA Ls)#} + {#M \l \ . M \# mset D#} = E" using remove_all_mset_minus[of Ls CA] ls_sub_ca by auto then have "mset (resolvent D A CA Ls) = E" unfolding resolvable_def Let_def resolvent_def using Al_p(1) Ls_def atms_of_poss res'(4) by (metis image_mset_union mset_append mset_map option.sel) ultimately have "E \ mset ` set (resolve_on A D CA)" unfolding resolve_on_def by auto } then show "{E. \AA \. ord_resolve S [mset CA] ({#Neg A#} + mset D) [AA] [A] \ E} \ mset ` set (resolve_on A D CA)" by auto qed lemma set_resolve_eq_UNION_set_resolve_on: "set (resolve C D) = (\L \ set D. (case L of Pos _ \ {} | Neg A \ if maximal_wrt A (mset D) then set (resolve_on A (remove1 L D) C) else {}))" unfolding resolve_def by (fastforce split: literal.splits if_splits) lemma resolve_eq_Bin_ord_resolve: "mset ` set (resolve C D) = Bin_ord_resolve (mset C) (mset D)" unfolding set_resolve_eq_UNION_set_resolve_on apply (unfold image_UN literal.case_distrib if_distrib) apply (subst resolve_on_eq_UNION_Bin_ord_resolve) apply (rule order_antisym) apply (force split: literal.splits if_splits) apply (clarsimp split: literal.splits if_splits) apply (rule_tac x = "Neg A" in bexI) apply (rule conjI) apply blast apply clarify apply (rule conjI) apply clarify apply (rule_tac x = AA in exI) apply (rule_tac x = \ in exI) apply (frule ord_resolve.simps[THEN iffD1]) apply force apply (drule ord_resolve.simps[THEN iffD1]) apply (clarsimp simp: eligible_iff simp del: subst_cls_add_mset subst_cls_union) apply (drule maximal_wrt_subst) apply sat apply (drule ord_resolve.simps[THEN iffD1]) using set_mset_mset by fastforce lemma poss_in_map_clauseD: "poss AA \# map_clause f C \ \AA0. poss AA0 \# C \ AA = {#f A. A \# AA0#}" proof (induct AA arbitrary: C) case (add A AA) note ih = this(1) and aaa_sub = this(2) have "Pos A \# map_clause f C" using aaa_sub by auto then obtain A0 where pa0_in: "Pos A0 \# C" and a: "A = f A0" by clarify (metis literal.distinct(1) literal.exhaust literal.inject(1) literal.simps(9,10)) have "poss AA \# map_clause f (C - {#Pos A0#})" using pa0_in aaa_sub[unfolded a] by (simp add: image_mset_remove1_mset_if insert_subset_eq_iff) then obtain AA0 where paa0_sub: "poss AA0 \# C - {#Pos A0#}" and aa: "AA = image_mset f AA0" using ih by meson have "poss (add_mset A0 AA0) \# C" using pa0_in paa0_sub by (simp add: insert_subset_eq_iff) moreover have "add_mset A AA = image_mset f (add_mset A0 AA0)" unfolding a aa by simp ultimately show ?case by blast qed simp lemma poss_subset_filterD: "poss AA \# {#L \l \. L \# mset C#} \ \AA0. poss AA0 \# mset C \ AA = AA0 \am \" unfolding subst_atm_mset_def subst_lit_def by (rule poss_in_map_clauseD) lemma neg_in_map_literalD: "Neg A \ map_literal f ` D \ \A0. Neg A0 \ D \ A = f A0" unfolding image_def by (clarify, case_tac x, auto) lemma neg_in_filterD: "Neg A \# {#L \l \'. L \# mset D#} \ \A0. Neg A0 \# mset D \ A = A0 \a \'" unfolding subst_lit_def image_def by (rule neg_in_map_literalD) simp lemma resolve_rename_eq_Bin_ord_resolve_rename: "mset ` set (resolve_rename C D) = Bin_ord_resolve_rename (mset C) (mset D)" proof (intro order_antisym subsetI) let ?\s = "renamings_apart [mset D, mset C]" define \' :: 's where "\' = hd ?\s" define \ :: 's where "\ = last ?\s" have tl_\s: "tl ?\s = [\]" unfolding \_def using renamings_apart_length Nitpick.size_list_simp(2) Suc_length_conv last.simps by (smt length_greater_0_conv list.sel(3)) { fix E assume e_in: "E \ mset ` set (resolve_rename C D)" from e_in obtain AA :: "'a multiset" and A :: 'a and \ :: 's where aa_sub: "poss AA \# mset C \ \" and a_in: "Neg A \# mset D \ \'" and res_e: "ord_resolve S [mset C \ \] {#L \l \'. L \# mset D#} [AA] [A] \ E" unfolding \'_def \_def apply atomize_elim using e_in unfolding resolve_rename_def Let_def resolve_eq_Bin_ord_resolve apply clarsimp apply (frule ord_resolve_one_side_prem) apply (frule ord_resolve.simps[THEN iffD1]) apply (rule_tac x = AA in exI) apply (clarsimp simp: subst_cls_def) apply (rule_tac x = A in exI) by (metis (full_types) Melem_subst_cls set_mset_mset subst_cls_def union_single_eq_member) obtain AA0 :: "'a multiset" where aa0_sub: "poss AA0 \# mset C" and aa: "AA = AA0 \am \" using aa_sub apply atomize_elim apply (rule ord_resolve.cases[OF res_e]) by (rule poss_subset_filterD[OF aa_sub[unfolded subst_cls_def]]) obtain A0 :: 'a where a0_in: "Neg A0 \ set D" and a: "A = A0 \a \'" apply atomize_elim apply (rule ord_resolve.cases[OF res_e]) using neg_in_filterD[OF a_in[unfolded subst_cls_def]] by simp show "E \ Bin_ord_resolve_rename (mset C) (mset D)" unfolding ord_resolve_rename.simps using res_e apply clarsimp apply (rule_tac x = AA0 in exI) apply (intro conjI) apply (rule aa0_sub) apply (rule_tac x = A0 in exI) apply (intro conjI) apply (rule a0_in) apply (rule_tac x = \ in exI) unfolding aa a \'_def[symmetric] \_def[symmetric] tl_\s by (simp add: subst_cls_def) } { fix E assume e_in: "E \ Bin_ord_resolve_rename (mset C) (mset D)" show "E \ mset ` set (resolve_rename C D)" using e_in unfolding resolve_rename_def Let_def resolve_eq_Bin_ord_resolve ord_resolve_rename.simps apply clarsimp apply (rule_tac x = "AA \am \" in exI) apply (rule_tac x = "A \a \'" in exI) apply (rule_tac x = \ in exI) unfolding tl_\s \'_def \_def by (simp add: subst_cls_def subst_cls_lists_def) } qed lemma bin_ord_FO_\_def: "ord_FO_\ S = {Infer {#CA#} DA E | CA DA AA A \ E. ord_resolve_rename S [CA] DA [AA] [A] \ E}" unfolding ord_FO_\_def apply (rule order.antisym) apply clarify apply (frule ord_resolve_rename_one_side_prem) apply simp apply (metis Suc_length_conv length_0_conv) by blast lemma ord_FO_\_side_prem: "\ \ ord_FO_\ S \ side_prems_of \ = {#THE D. D \# side_prems_of \#}" unfolding bin_ord_FO_\_def by clarsimp lemma ord_FO_\_infer_from_Collect_eq: "{\ \ ord_FO_\ S. infer_from (DD \ {C}) \ \ C \# prems_of \} = {\ \ ord_FO_\ S. \D \ DD \ {C}. prems_of \ = {#C, D#}}" unfolding infer_from_def apply (rule set_eq_subset[THEN iffD2]) apply (rule conjI) apply clarify apply (subst (asm) (1 2) ord_FO_\_side_prem, assumption, assumption) apply (subst (1) ord_FO_\_side_prem, assumption) apply force apply clarify apply (subst (asm) (1) ord_FO_\_side_prem, assumption) apply (subst (1 2) ord_FO_\_side_prem, assumption) by force lemma inferences_between_eq_UNION: "inference_system.inferences_between (ord_FO_\ S) Q C = inference_system.inferences_between (ord_FO_\ S) {C} C \ (\D \ Q. inference_system.inferences_between (ord_FO_\ S) {D} C)" unfolding ord_FO_\_infer_from_Collect_eq inference_system.inferences_between_def by auto lemma concls_of_inferences_between_singleton_eq_Bin_ord_resolve_rename: "concls_of (inference_system.inferences_between (ord_FO_\ S) {D} C) = Bin_ord_resolve_rename C C \ Bin_ord_resolve_rename C D \ Bin_ord_resolve_rename D C" proof (intro order_antisym subsetI) fix E assume e_in: "E \ concls_of (inference_system.inferences_between (ord_FO_\ S) {D} C)" then show "E \ Bin_ord_resolve_rename C C \ Bin_ord_resolve_rename C D \ Bin_ord_resolve_rename D C" unfolding inference_system.inferences_between_def ord_FO_\_infer_from_Collect_eq bin_ord_FO_\_def infer_from_def by (fastforce simp: add_mset_eq_add_mset) qed (force simp: inference_system.inferences_between_def infer_from_def ord_FO_\_def) lemma concls_of_inferences_between_eq_Bin_ord_resolve_rename: "concls_of (inference_system.inferences_between (ord_FO_\ S) Q C) = Bin_ord_resolve_rename C C \ (\D \ Q. Bin_ord_resolve_rename C D \ Bin_ord_resolve_rename D C)" by (subst inferences_between_eq_UNION) (auto simp: image_Un image_UN concls_of_inferences_between_singleton_eq_Bin_ord_resolve_rename) lemma resolve_rename_either_way_eq_congls_of_inferences_between: "mset ` set (resolve_rename C C) \ (\D \ Q. mset ` set (resolve_rename_either_way C D)) = concls_of (inference_system.inferences_between (ord_FO_\ S) (mset ` Q) (mset C))" by (simp add: resolve_rename_either_way_def image_Un resolve_rename_eq_Bin_ord_resolve_rename concls_of_inferences_between_eq_Bin_ord_resolve_rename UN_Un_distrib) lemma compute_inferences: assumes ci_in: "(C, i) \ set P" and ci_min: "\(D, j) \# mset (map (apfst mset) P). weight (mset C, i) \ weight (D, j)" shows "wstate_of_dstate ([], P, Q, n) \\<^sub>w wstate_of_dstate (map (\D. (D, n)) (remdups_gen mset (resolve_rename C C @ concat (map (resolve_rename_either_way C \ fst) Q))), filter (\(D, j). mset D \ mset C) P, (C, i) # Q, Suc n)" (is "_ \\<^sub>w wstate_of_dstate (?N, _)") proof - have ms_ci_in: "(mset C, i) \# image_mset (apfst mset) (mset P)" using ci_in by force show ?thesis apply (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w)", OF _ _ wrp.inference_computation[of "mset (map (apfst mset) P) - {#(mset C, i)#}" "mset C" i "mset (map (apfst mset) ?N)" n "mset (map (apfst mset) Q)"]]) apply (simp add: add_mset_remove_trivial_eq[THEN iffD2, OF ms_ci_in, symmetric]) using ms_ci_in apply (simp add: ci_in image_mset_remove1_mset_if) apply (smt apfst_conv case_prodE case_prodI2 case_prod_conv filter_mset_cong image_mset_filter_swap mset_filter) apply (metis ci_min in_diffD) apply (simp only: list.map_comp apfst_comp_rpair_const) apply (simp only: list.map_comp[symmetric]) apply (subst mset_map) apply (unfold mset_map_remdups_gen mset_remdups_gen_ident) apply (subst image_mset_mset_set) apply (simp add: inj_on_def) apply (subst mset_set_eq_iff) apply simp apply (simp add: finite_ord_FO_resolution_inferences_between) apply (rule arg_cong[of _ _ "\N. (\D. (D, n)) ` N"]) apply (simp only: map_concat list.map_comp image_comp) using resolve_rename_either_way_eq_congls_of_inferences_between[of C "fst ` set Q", symmetric] by (simp add: image_comp comp_def image_UN) qed lemma nonfinal_deterministic_RP_step: assumes nonfinal: "\ is_final_dstate St" and step: "St' = deterministic_RP_step St" shows "wstate_of_dstate St \\<^sub>w\<^sup>+ wstate_of_dstate St'" proof - obtain N P Q :: "'a dclause list" and n :: nat where st: "St = (N, P, Q, n)" by (cases St) blast note step = step[unfolded st deterministic_RP_step.simps, simplified] show ?thesis proof (cases "\Ci \ set P \ set Q. fst Ci = []") case nil_in: True note step = step[simplified nil_in, simplified] have nil_in': "[] \ fst ` set (P @ Q)" using nil_in by (force simp: image_def) have star: "[] \ fst ` set (P @ Q) \ wstate_of_dstate (N, P, Q, n) \\<^sub>w\<^sup>* wstate_of_dstate ([], [], remdups_clss P @ Q, n + length (remdups_clss P))" proof (induct "length (remdups_clss P)" arbitrary: N P Q n) case 0 note len_p = this(1) and nil_in' = this(2) have p_nil: "P = []" using len_p remdups_clss_Nil_iff by simp have "wstate_of_dstate (N, [], Q, n) \\<^sub>w\<^sup>* wstate_of_dstate ([], [], Q, n)" by (rule empty_N_if_Nil_in_P_or_Q[OF nil_in'[unfolded p_nil]]) then show ?case unfolding p_nil by simp next case (Suc k) note ih = this(1) and suc_k = this(2) and nil_in' = this(3) have "P \ []" using suc_k remdups_clss_Nil_iff by force hence p_cons: "P = hd P # tl P" by simp obtain C :: "'a lclause" and i :: nat where ci: "(C, i) = select_min_weight_clause (hd P) (tl P)" by (metis prod.exhaust) have ci_in: "(C, i) \ set P" unfolding ci using p_cons select_min_weight_clause_in[of "hd P" "tl P"] by simp have ci_min: "\(D, j) \# mset (map (apfst mset) P). weight (mset C, i) \ weight (D, j)" by (subst p_cons) (simp add: select_min_weight_clause_min_weight[OF ci, simplified]) let ?P' = "filter (\(D, j). mset D \ mset C) P" have ms_p'_ci_q_eq: "mset (remdups_clss ?P' @ (C, i) # Q) = mset (remdups_clss P @ Q)" apply (subst (2) p_cons) apply (subst remdups_clss.simps(2)) by (auto simp: Let_def case_prod_beta p_cons[symmetric] ci[symmetric]) then have len_p: "length (remdups_clss P) = length (remdups_clss ?P') + 1" by (smt Suc_eq_plus1_left add.assoc add_right_cancel length_Cons length_append mset_eq_length) have "wstate_of_dstate (N, P, Q, n) \\<^sub>w\<^sup>* wstate_of_dstate ([], P, Q, n)" by (rule empty_N_if_Nil_in_P_or_Q[OF nil_in']) also obtain N' :: "'a dclause list" where "\ \\<^sub>w wstate_of_dstate (N', ?P', (C, i) # Q, Suc n)" by (atomize_elim, rule exI, rule compute_inferences[OF ci_in], use ci_min in fastforce) also have "\ \\<^sub>w\<^sup>* wstate_of_dstate ([], [], remdups_clss P @ Q, n + length (remdups_clss P))" apply (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w\<^sup>*)", OF _ _ ih[of ?P' "(C, i) # Q" N' "Suc n"], OF refl]) using ms_p'_ci_q_eq suc_k nil_in' ci_in apply (simp_all add: len_p) apply (metis (no_types) apfst_conv image_mset_add_mset) by force finally show ?case . qed show ?thesis unfolding st step using star[OF nil_in'] nonfinal[unfolded st is_final_dstate.simps] by cases simp_all next case nil_ni: False note step = step[simplified nil_ni, simplified] show ?thesis proof (cases N) case n_nil: Nil note step = step[unfolded n_nil, simplified] show ?thesis proof (cases P) case Nil then have False using n_nil nonfinal[unfolded st] by (simp add: is_final_dstate.simps) then show ?thesis using step by simp next case p_cons: (Cons P0 P') note step = step[unfolded p_cons list.case, folded p_cons] obtain C :: "'a lclause" and i :: nat where ci: "(C, i) = select_min_weight_clause P0 P'" by (metis prod.exhaust) note step = step[unfolded select, simplified] have ci_in: "(C, i) \ set P" by (rule select_min_weight_clause_in[of P0 P', folded ci p_cons]) show ?thesis unfolding st n_nil step p_cons[symmetric] ci[symmetric] prod.case by (rule tranclp.r_into_trancl, rule compute_inferences[OF ci_in]) (simp add: select_min_weight_clause_min_weight[OF ci, simplified] p_cons) qed next case n_cons: (Cons Ci N') note step = step[unfolded n_cons, simplified] obtain C :: "'a lclause" and i :: nat where ci: "Ci = (C, i)" by (cases Ci) simp note step = step[unfolded ci, simplified] define C' :: "'a lclause" where "C' = reduce (map fst P @ map fst Q) [] C" note step = step[unfolded ci C'_def[symmetric], simplified] have "wstate_of_dstate ((E @ C, i) # N', P, Q, n) \\<^sub>w\<^sup>* wstate_of_dstate ((E @ reduce (map fst P @ map fst Q) E C, i) # N', P, Q, n)" for E unfolding C'_def proof (induct C arbitrary: E) case (Cons L C) note ih = this(1) show ?case proof (cases "is_reducible_lit (map fst P @ map fst Q) (E @ C) L") case l_red: True then have red_lc: "reduce (map fst P @ map fst Q) E (L # C) = reduce (map fst P @ map fst Q) E C" by simp obtain D D' :: "'a literal list" and L' :: "'a literal" and \ :: 's where "D \ set (map fst P @ map fst Q)" and "D' = remove1 L' D" and "L' \ set D" and "- L = L' \l \" and "mset D' \ \ \# mset (E @ C)" using l_red unfolding is_reducible_lit_def comp_def by blast then have \: "mset D' + {#L'#} \ set (map (mset \ fst) (P @ Q))" "- L = L' \l \ \ mset D' \ \ \# mset (E @ C)" unfolding is_reducible_lit_def by (auto simp: comp_def) have "wstate_of_dstate ((E @ L # C, i) # N', P, Q, n) \\<^sub>w wstate_of_dstate ((E @ C, i) # N', P, Q, n)" by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w)", OF _ _ wrp.forward_reduction[of "mset D'" L' "mset (map (apfst mset) P)" "mset (map (apfst mset) Q)" L \ "mset (E @ C)" "mset (map (apfst mset) N')" i n]]) (use \ in \auto simp: comp_def\) then show ?thesis unfolding red_lc using ih[of E] by (rule converse_rtranclp_into_rtranclp) next case False then show ?thesis using ih[of "L # E"] by simp qed qed simp then have red_C: "wstate_of_dstate ((C, i) # N', P, Q, n) \\<^sub>w\<^sup>* wstate_of_dstate ((C', i) # N', P, Q, n)" unfolding C'_def by (metis self_append_conv2) have proc_C: "wstate_of_dstate ((C', i) # N', P', Q', n') \\<^sub>w wstate_of_dstate (N', (C', i) # P', Q', n')" for P' Q' n' by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w)", OF _ _ wrp.clause_processing[of "mset (map (apfst mset) N')" "mset C'" i "mset (map (apfst mset) P')" "mset (map (apfst mset) Q')" n']], simp+) show ?thesis proof (cases "C' = []") case True note c'_nil = this note step = step[simplified c'_nil, simplified] have filter_p: "filter (Not \ strictly_subsume [[]] \ fst) P = []" and filter_q: "filter (Not \ strictly_subsume [[]] \ fst) Q = []" using nil_ni unfolding strictly_subsume_def filter_empty_conv find_None_iff by force+ note red_C[unfolded c'_nil] also have "wstate_of_dstate (([], i) # N', P, Q, n) \\<^sub>w\<^sup>* wstate_of_dstate (([], i) # N', [], Q, n)" by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w\<^sup>*)", OF _ _ remove_strictly_subsumed_clauses_in_P[of "[]" _ "[]", unfolded append_Nil], OF refl]) (auto simp: filter_p) also have "\ \\<^sub>w\<^sup>* wstate_of_dstate (([], i) # N', [], [], n)" by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w\<^sup>*)", OF _ _ remove_strictly_subsumed_clauses_in_Q[of "[]" _ _ "[]", unfolded append_Nil], OF refl]) (auto simp: filter_q) also note proc_C[unfolded c'_nil, THEN tranclp.r_into_trancl[of "(\\<^sub>w)"]] also have "wstate_of_dstate (N', [([], i)], [], n) \\<^sub>w\<^sup>* wstate_of_dstate ([], [([], i)], [], n)" by (rule empty_N_if_Nil_in_P_or_Q) simp also have "\ \\<^sub>w wstate_of_dstate ([], [], [([], i)], Suc n)" by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w)", OF _ _ wrp.inference_computation[of "{#}" "{#}" i "{#}" n "{#}"]]) (auto simp: ord_FO_resolution_inferences_between_empty_empty) finally show ?thesis unfolding step st n_cons ci . next case c'_nnil: False note step = step[simplified c'_nnil, simplified] show ?thesis proof (cases "is_tautology C' \ subsume (map fst P @ map fst Q) C'") case taut_or_subs: True note step = step[simplified taut_or_subs, simplified] have "wstate_of_dstate ((C', i) # N', P, Q, n) \\<^sub>w wstate_of_dstate (N', P, Q, n)" proof (cases "is_tautology C'") case True then obtain A :: 'a where neg_a: "Neg A \ set C'" and pos_a: "Pos A \ set C'" unfolding is_tautology_def by blast show ?thesis by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w)", OF _ _ wrp.tautology_deletion[of A "mset C'" "mset (map (apfst mset) N')" i "mset (map (apfst mset) P)" "mset (map (apfst mset) Q)" n]]) (use neg_a pos_a in simp_all) next case False then have "subsume (map fst P @ map fst Q) C'" using taut_or_subs by blast then obtain D :: "'a lclause" where d_in: "D \ set (map fst P @ map fst Q)" and subs: "subsumes (mset D) (mset C')" unfolding subsume_def by blast show ?thesis by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\\<^sub>w)", OF _ _ wrp.forward_subsumption[of "mset D" "mset (map (apfst mset) P)" "mset (map (apfst mset) Q)" "mset C'" "mset (map (apfst mset) N')" i n]], use d_in subs in \auto simp: subsume_def\) qed then show ?thesis unfolding step st n_cons ci using red_C by (rule rtranclp_into_tranclp1[rotated]) next case not_taut_or_subs: False note step = step[simplified not_taut_or_subs, simplified] define P' :: "('a literal list \ nat) list" where "P' = reduce_all C' P" obtain back_to_P Q' :: "'a dclause list" where red_Q: "(back_to_P, Q') = reduce_all2 C' Q" by (metis prod.exhaust) note step = step[unfolded red_Q[symmetric], simplified] define Q'' :: "('a literal list \ nat) list" where "Q'' = filter (Not \ strictly_subsume [C'] \ fst) Q'" define P'' :: "('a literal list \ nat) list" where "P'' = filter (Not \ strictly_subsume [C'] \ fst) (back_to_P @ P')" note step = step[unfolded P'_def[symmetric] Q''_def[symmetric] P''_def[symmetric], simplified] note red_C also have "wstate_of_dstate ((C', i) # N', P, Q, n) \\<^sub>w\<^sup>* wstate_of_dstate ((C', i) # N', P', Q, n)" unfolding P'_def by (rule reduce_clauses_in_P[of _ _ "[]", unfolded append_Nil]) simp+ also have "\ \\<^sub>w\<^sup>* wstate_of_dstate ((C', i) # N', back_to_P @ P', Q', n)" unfolding P'_def by (rule reduce_clauses_in_Q[of C' _ _ "[]" Q, folded red_Q, unfolded append_Nil prod.sel]) (auto intro: reduce_idem simp: reduce_all_def) also have "\ \\<^sub>w\<^sup>* wstate_of_dstate ((C', i) # N', back_to_P @ P', Q'', n)" unfolding Q''_def by (rule remove_strictly_subsumed_clauses_in_Q[of _ _ _ "[]", unfolded append_Nil]) simp also have "\ \\<^sub>w\<^sup>* wstate_of_dstate ((C', i) # N', P'', Q'', n)" unfolding P''_def by (rule remove_strictly_subsumed_clauses_in_P[of _ _ "[]", unfolded append_Nil]) auto also note proc_C[THEN tranclp.r_into_trancl[of "(\\<^sub>w)"]] finally show ?thesis unfolding step st n_cons ci P''_def by simp qed qed qed qed qed lemma final_deterministic_RP_step: "is_final_dstate St \ deterministic_RP_step St = St" by (cases St) (auto simp: deterministic_RP_step.simps is_final_dstate.simps) lemma deterministic_RP_SomeD: assumes "deterministic_RP (N, P, Q, n) = Some R" shows "\N' P' Q' n'. (\k. (deterministic_RP_step ^^ k) (N, P, Q, n) = (N', P', Q', n')) \ is_final_dstate (N', P', Q', n') \ R = map fst Q'" proof (induct rule: deterministic_RP.raw_induct[OF _ assms]) case (1 self_call St R) note ih = this(1) and step = this(2) obtain N P Q :: "'a dclause list" and n :: nat where st: "St = (N, P, Q, n)" by (cases St) blast note step = step[unfolded st, simplified] show ?case proof (cases "is_final_dstate (N, P, Q, n)") case True then have "(deterministic_RP_step ^^ 0) (N, P, Q, n) = (N, P, Q, n) \ is_final_dstate (N, P, Q, n) \ R = map fst Q" using step by simp then show ?thesis unfolding st by blast next case nonfinal: False note step = step[simplified nonfinal, simplified] obtain N' P' Q' :: "'a dclause list" and n' k :: nat where "(deterministic_RP_step ^^ k) (deterministic_RP_step (N, P, Q, n)) = (N', P', Q', n')" and "is_final_dstate (N', P', Q', n')" "R = map fst Q'" using ih[OF step] by blast then show ?thesis unfolding st funpow_Suc_right[symmetric, THEN fun_cong, unfolded comp_apply] by blast qed qed context fixes N0 :: "'a dclause list" and n0 :: nat and R :: "'a lclause list" begin abbreviation St0 :: "'a dstate" where "St0 \ (N0, [], [], n0)" abbreviation grounded_N0 where "grounded_N0 \ grounding_of_clss (set (map (mset \ fst) N0))" abbreviation grounded_R :: "'a clause set" where "grounded_R \ grounding_of_clss (set (map mset R))" primcorec derivation_from :: "'a dstate \ 'a dstate llist" where "derivation_from St = LCons St (if is_final_dstate St then LNil else derivation_from (deterministic_RP_step St))" abbreviation Sts :: "'a dstate llist" where "Sts \ derivation_from St0" abbreviation wSts :: "'a wstate llist" where "wSts \ lmap wstate_of_dstate Sts" lemma full_deriv_wSts_trancl_weighted_RP: "full_chain (\\<^sub>w\<^sup>+) wSts" proof - have "Sts' = derivation_from St0' \ full_chain (\\<^sub>w\<^sup>+) (lmap wstate_of_dstate Sts')" for St0' Sts' proof (coinduction arbitrary: St0' Sts' rule: full_chain.coinduct) case sts': full_chain show ?case proof (cases "is_final_dstate St0'") case True then have "ltl (lmap wstate_of_dstate Sts') = LNil" unfolding sts' by simp then have "lmap wstate_of_dstate Sts' = LCons (wstate_of_dstate St0') LNil" unfolding sts' by (subst derivation_from.code, subst (asm) derivation_from.code, auto) moreover have "\St''. \ wstate_of_dstate St0' \\<^sub>w St''" using True by (rule is_final_dstate_imp_not_weighted_RP) ultimately show ?thesis by (meson tranclpD) next case nfinal: False have "lmap wstate_of_dstate Sts' = LCons (wstate_of_dstate St0') (lmap wstate_of_dstate (ltl Sts'))" unfolding sts' by (subst derivation_from.code) simp moreover have "ltl Sts' = derivation_from (deterministic_RP_step St0')" unfolding sts' using nfinal by (subst derivation_from.code) simp moreover have "wstate_of_dstate St0' \\<^sub>w\<^sup>+ wstate_of_dstate (lhd (ltl Sts'))" unfolding sts' using nonfinal_deterministic_RP_step[OF nfinal refl] nfinal by (subst derivation_from.code) simp ultimately show ?thesis by fastforce qed qed then show ?thesis by blast qed lemmas deriv_wSts_trancl_weighted_RP = full_chain_imp_chain[OF full_deriv_wSts_trancl_weighted_RP] definition sswSts :: "'a wstate llist" where "sswSts = (SOME wSts'. full_chain (\\<^sub>w) wSts' \ emb wSts wSts' \ lhd wSts' = lhd wSts \ llast wSts' = llast wSts)" lemma sswSts: "full_chain (\\<^sub>w) sswSts \ emb wSts sswSts \ lhd sswSts = lhd wSts \ llast sswSts = llast wSts" unfolding sswSts_def by (rule someI_ex[OF full_chain_tranclp_imp_exists_full_chain[OF full_deriv_wSts_trancl_weighted_RP]]) lemmas full_deriv_sswSts_weighted_RP = sswSts[THEN conjunct1] lemmas emb_sswSts = sswSts[THEN conjunct2, THEN conjunct1] lemmas lfinite_sswSts_iff = emb_lfinite[OF emb_sswSts] lemmas lhd_sswSts = sswSts[THEN conjunct2, THEN conjunct2, THEN conjunct1] lemmas llast_sswSts = sswSts[THEN conjunct2, THEN conjunct2, THEN conjunct2] lemmas deriv_sswSts_weighted_RP = full_chain_imp_chain[OF full_deriv_sswSts_weighted_RP] lemma not_lnull_sswSts: "\ lnull sswSts" using deriv_sswSts_weighted_RP by (cases rule: chain.cases) auto lemma empty_ssgP0: "wrp.P_of_wstate (lhd sswSts) = {}" unfolding lhd_sswSts by (subst derivation_from.code) simp lemma empty_ssgQ0: "wrp.Q_of_wstate (lhd sswSts) = {}" unfolding lhd_sswSts by (subst derivation_from.code) simp lemmas sswSts_thms = full_deriv_sswSts_weighted_RP empty_ssgP0 empty_ssgQ0 abbreviation S_ssgQ :: "'a clause \ 'a clause" where "S_ssgQ \ wrp.S_gQ sswSts" abbreviation ord_\ :: "'a inference set" where "ord_\ \ ground_resolution_with_selection.ord_\ S_ssgQ" abbreviation Rf :: "'a clause set \ 'a clause set" where "Rf \ standard_redundancy_criterion.Rf" abbreviation Ri :: "'a clause set \ 'a inference set" where "Ri \ standard_redundancy_criterion.Ri ord_\" abbreviation saturated_upto :: "'a clause set \ bool" where "saturated_upto \ redundancy_criterion.saturated_upto ord_\ Rf Ri" context assumes drp_some: "deterministic_RP St0 = Some R" begin lemma lfinite_Sts: "lfinite Sts" proof (induct rule: deterministic_RP.raw_induct[OF _ drp_some]) case (1 self_call St St') note ih = this(1) and step = this(2) show ?case using step by (subst derivation_from.code, auto intro: ih) qed lemma lfinite_wSts: "lfinite wSts" by (rule lfinite_lmap[THEN iffD2, OF lfinite_Sts]) lemmas lfinite_sswSts = lfinite_sswSts_iff[THEN iffD2, OF lfinite_wSts] theorem deterministic_RP_saturated: "saturated_upto grounded_R" (is ?saturated) and deterministic_RP_model: "I \s grounded_N0 \ I \s grounded_R" (is ?model) proof - obtain N' P' Q' :: "'a dclause list" and n' k :: nat where k_steps: "(deterministic_RP_step ^^ k) St0 = (N', P', Q', n')" (is "_ = ?Stk") and final: "is_final_dstate (N', P', Q', n')" and r: "R = map fst Q'" using deterministic_RP_SomeD[OF drp_some] by blast have wrp: "wstate_of_dstate St0 \\<^sub>w\<^sup>* wstate_of_dstate (llast Sts)" using lfinite_chain_imp_rtranclp_lhd_llast by (metis (no_types) deriv_sswSts_weighted_RP derivation_from.disc_iff derivation_from.simps(2) lfinite_Sts lfinite_sswSts llast_lmap llist.map_sel(1) sswSts) have last_sts: "llast Sts = ?Stk" proof - have "(deterministic_RP_step ^^ k') St0' = ?Stk \ llast (derivation_from St0') = ?Stk" for St0' k' proof (induct k' arbitrary: St0') case 0 then show ?case using final by (subst derivation_from.code) simp next case (Suc k') note ih = this(1) and suc_k'_steps = this(2) show ?case proof (cases "is_final_dstate St0'") case True then show ?thesis using ih[of "deterministic_RP_step St0'"] suc_k'_steps final_deterministic_RP_step funpow_fixpoint[of deterministic_RP_step] by auto next case False then show ?thesis using ih[of "deterministic_RP_step St0'"] suc_k'_steps by (subst derivation_from.code) (simp add: llast_LCons funpow_swap1[symmetric]) qed qed then show ?thesis using k_steps by blast qed have fin_gr_fgsts: "lfinite (lmap wrp.grounding_of_wstate sswSts)" by (rule lfinite_lmap[THEN iffD2, OF lfinite_sswSts]) have lim_last: "Liminf_llist (lmap wrp.grounding_of_wstate sswSts) = wrp.grounding_of_wstate (llast sswSts)" unfolding lfinite_Liminf_llist[OF fin_gr_fgsts] llast_lmap[OF lfinite_sswSts not_lnull_sswSts] using not_lnull_sswSts by simp have gr_st0: "wrp.grounding_of_wstate (wstate_of_dstate St0) = grounded_N0" - by (simp add: clss_of_state_def comp_def) + unfolding comp_def by simp have "?saturated \ ?model" proof (cases "[] \ set R") case True then have emp_in: "{#} \ grounded_R" unfolding grounding_of_clss_def grounding_of_cls_def by (auto intro: ex_ground_subst) have "grounded_R \ wrp.grounding_of_wstate (llast sswSts)" unfolding r llast_sswSts - by (simp add: last_sts llast_lmap[OF lfinite_Sts] clss_of_state_def grounding_of_clss_def) + by (simp add: last_sts llast_lmap[OF lfinite_Sts] grounding_of_clss_def) then have gr_last_st: "grounded_R \ wrp.grounding_of_wstate (wstate_of_dstate (llast Sts))" by (simp add: lfinite_Sts llast_lmap llast_sswSts) have gr_r_fls: "\ I \s grounded_R" using emp_in unfolding true_clss_def by force then have gr_last_fls: "\ I \s wrp.grounding_of_wstate (wstate_of_dstate (llast Sts))" using gr_last_st unfolding true_clss_def by auto have ?saturated unfolding wrp.ord_\_saturated_upto_def[OF sswSts_thms] wrp.ord_\_contradiction_Rf[OF sswSts_thms emp_in] inference_system.inferences_from_def by auto moreover have ?model unfolding gr_r_fls[THEN eq_False[THEN iffD2]] by (rule rtranclp_imp_eq_image[of "(\\<^sub>w)" "\St. I \s wrp.grounding_of_wstate St", OF _ wrp, unfolded gr_st0 gr_last_fls[THEN eq_False[THEN iffD2]]]) (use wrp.weighted_RP_model[OF sswSts_thms] in blast) ultimately show ?thesis by blast next case False then have gr_last: "wrp.grounding_of_wstate (llast sswSts) = grounded_R" using final unfolding r llast_sswSts - by (simp add: last_sts llast_lmap[OF lfinite_Sts] clss_of_state_def comp_def - is_final_dstate.simps) + by (simp add: last_sts llast_lmap[OF lfinite_Sts] comp_def is_final_dstate.simps) then have gr_last_st: "wrp.grounding_of_wstate (wstate_of_dstate (llast Sts)) = grounded_R" by (simp add: lfinite_Sts llast_lmap llast_sswSts) have ?saturated using wrp.weighted_RP_saturated[OF sswSts_thms, unfolded gr_last lim_last] by auto moreover have ?model by (rule rtranclp_imp_eq_image[of "(\\<^sub>w)" "\St. I \s wrp.grounding_of_wstate St", OF _ wrp, unfolded gr_st0 gr_last_st]) (use wrp.weighted_RP_model[OF sswSts_thms] in blast) ultimately show ?thesis by blast qed then show ?saturated and ?model by blast+ qed corollary deterministic_RP_refutation: "\ satisfiable grounded_N0 \ {#} \ grounded_R" (is "?lhs \ ?rhs") proof assume ?rhs then have "\ satisfiable grounded_R" unfolding true_clss_def true_cls_def by force then show ?lhs using deterministic_RP_model[THEN iffD1] by blast next assume ?lhs then have "\ satisfiable grounded_R" using deterministic_RP_model[THEN iffD2] by blast then show ?rhs unfolding wrp.ord_\_saturated_upto_complete[OF sswSts_thms deterministic_RP_saturated] . qed end context assumes drp_none: "deterministic_RP St0 = None" begin theorem deterministic_RP_complete: "satisfiable grounded_N0" proof (rule ccontr) assume unsat: "\ satisfiable grounded_N0" have unsat_wSts0: "\ satisfiable (wrp.grounding_of_wstate (lhd wSts))" - using unsat by (subst derivation_from.code) (simp add: clss_of_state_def comp_def) + using unsat by (subst derivation_from.code) (simp add: comp_def) have bot_in_ss: "{#} \ Q_of_state (wrp.Liminf_wstate sswSts)" by (rule wrp.weighted_RP_complete[OF sswSts_thms unsat_wSts0[folded lhd_sswSts]]) have bot_in_lim: "{#} \ Q_of_state (wrp.Liminf_wstate wSts)" proof (cases "lfinite Sts") case fin: True have "wrp.Liminf_wstate sswSts = wrp.Liminf_wstate wSts" by (rule Liminf_state_fin, simp_all add: fin lfinite_sswSts_iff not_lnull_sswSts, subst (1 2) llast_lmap, simp_all add: lfinite_sswSts_iff fin not_lnull_sswSts llast_sswSts) then show ?thesis using bot_in_ss by simp next case False then show ?thesis using bot_in_ss Q_of_Liminf_state_inf[OF _ emb_lmap[OF emb_sswSts]] by auto qed then obtain k :: nat where k_lt: "enat k < llength Sts" and emp_in: "{#} \ wrp.Q_of_wstate (lnth wSts k)" unfolding Liminf_state_def Liminf_llist_def by auto have emp_in: "{#} \ Q_of_state (state_of_dstate ((deterministic_RP_step ^^ k) St0))" proof - have "enat k < llength Sts' \ Sts' = derivation_from St0' \ {#} \ wrp.Q_of_wstate (lnth (lmap wstate_of_dstate Sts') k) \ {#} \ Q_of_state (state_of_dstate ((deterministic_RP_step ^^ k) St0'))" for St0' Sts' k proof (induction k arbitrary: St0' Sts') case 0 then show ?case by (subst (asm) derivation_from.code, cases St0', auto simp: comp_def) next case (Suc k) note ih = this(1) and sk_lt = this(2) and sts' = this(3) and emp_in_sk = this(4) have k_lt: "enat k < llength (ltl Sts')" using sk_lt by (cases Sts') (auto simp: Suc_ile_eq) moreover have "ltl Sts' = derivation_from (deterministic_RP_step St0')" using sts' k_lt by (cases Sts') auto moreover have "{#} \ wrp.Q_of_wstate (lnth (lmap wstate_of_dstate (ltl Sts')) k)" using emp_in_sk k_lt by (cases Sts') auto ultimately show ?case using ih[of "ltl Sts'" "deterministic_RP_step St0'"] by (simp add: funpow_swap1) qed then show ?thesis using k_lt emp_in by blast qed have "deterministic_RP St0 \ None" by (rule is_final_dstate_funpow_imp_deterministic_RP_neq_None[of "Suc k"], cases "(deterministic_RP_step ^^ k) St0", use emp_in in \force simp: deterministic_RP_step.simps is_final_dstate.simps\) then show False using drp_none .. qed end end end end diff --git a/thys/Functional_Ordered_Resolution_Prover/Weighted_FO_Ordered_Resolution_Prover.thy b/thys/Functional_Ordered_Resolution_Prover/Weighted_FO_Ordered_Resolution_Prover.thy --- a/thys/Functional_Ordered_Resolution_Prover/Weighted_FO_Ordered_Resolution_Prover.thy +++ b/thys/Functional_Ordered_Resolution_Prover/Weighted_FO_Ordered_Resolution_Prover.thy @@ -1,891 +1,889 @@ (* Title: A Fair Ordered Resolution Prover for First-Order Clauses with Weights Author: Anders Schlichtkrull , 2017 Author: Jasmin Blanchette , 2017 Maintainer: Anders Schlichtkrull *) section \A Fair Ordered Resolution Prover for First-Order Clauses with Weights\ text \ The \weighted_RP\ prover introduced below operates on finite multisets of clauses and organizes the multiset of processed clauses as a priority queue to ensure that inferences are performed in a fair manner, to guarantee completeness. \ theory Weighted_FO_Ordered_Resolution_Prover imports Ordered_Resolution_Prover.FO_Ordered_Resolution_Prover begin subsection \Library\ (* TODO: Move to "Coinductive"? *) lemma ldrop_Suc_conv_ltl: "ldrop (enat (Suc k)) xs = ltl (ldrop (enat k) xs)" by (metis eSuc_enat ldrop_eSuc_conv_ltl) (* TODO: Move to "Coinductive"? *) lemma lhd_ldrop': assumes "enat k < llength xs" shows "lhd (ldrop (enat k) xs) = lnth xs k" using assms by (simp add: lhd_ldrop) (* TODO: Move to "Multiset_More.thy". *) lemma filter_mset_empty_if_finite_and_filter_set_empty: assumes "{x \ X. P x} = {}" and "finite X" shows "{#x \# mset_set X. P x#} = {#}" proof - have empty_empty: "\Y. set_mset Y = {} \ Y = {#}" by auto from assms have "set_mset {#x \# mset_set X. P x#} = {}" by auto then show ?thesis by (rule empty_empty) qed (* TODO: Move to "Lazy_List_Chain.thy". *) lemma inf_chain_ltl_chain: "chain R xs \ llength xs = \ \ chain R (ltl xs)" unfolding chain.simps[of R xs] llength_eq_infty_conv_lfinite by (metis lfinite_code(1) lfinite_ltl llist.sel(3)) (* TODO: Move to "Lazy_List_Chain.thy". *) lemma inf_chain_ldrop_chain: assumes chain: "chain R xs" and inf: "\ lfinite xs" shows "chain R (ldrop (enat k) xs)" proof (induction k) case 0 then show ?case using zero_enat_def chain by auto next case (Suc k) have "llength (ldrop (enat k) xs) = \" using inf by (simp add: not_lfinite_llength) with Suc have "chain R (ltl (ldrop (enat k) xs))" using inf_chain_ltl_chain[of R "(ldrop (enat k) xs)"] by auto then show ?case using ldrop_Suc_conv_ltl[of k xs] by auto qed subsection \Prover\ type_synonym 'a wclause = "'a clause \ nat" type_synonym 'a wstate = "'a wclause multiset \ 'a wclause multiset \ 'a wclause multiset \ nat" fun state_of_wstate :: "'a wstate \ 'a state" where "state_of_wstate (N, P, Q, n) = (set_mset (image_mset fst N), set_mset (image_mset fst P), set_mset (image_mset fst Q))" locale weighted_FO_resolution_prover = FO_resolution_prover S subst_atm id_subst comp_subst renamings_apart atm_of_atms mgu less_atm for S :: "('a :: wellorder) clause \ 'a clause" and subst_atm :: "'a \ 's \ 'a" and id_subst :: "'s" and comp_subst :: "'s \ 's \ 's" and renamings_apart :: "'a clause list \ 's list" and atm_of_atms :: "'a list \ 'a" and mgu :: "'a set set \ 's option" and less_atm :: "'a \ 'a \ bool" + fixes weight :: "'a clause \ nat \ nat" assumes weight_mono: "i < j \ weight (C, i) < weight (C, j)" begin abbreviation clss_of_wstate :: "'a wstate \ 'a clause set" where "clss_of_wstate St \ clss_of_state (state_of_wstate St)" abbreviation N_of_wstate :: "'a wstate \ 'a clause set" where "N_of_wstate St \ N_of_state (state_of_wstate St)" abbreviation P_of_wstate :: "'a wstate \ 'a clause set" where "P_of_wstate St \ P_of_state (state_of_wstate St)" abbreviation Q_of_wstate :: "'a wstate \ 'a clause set" where "Q_of_wstate St \ Q_of_state (state_of_wstate St)" fun wN_of_wstate :: "'a wstate \ 'a wclause multiset" where "wN_of_wstate (N, P, Q, n) = N" fun wP_of_wstate :: "'a wstate \ 'a wclause multiset" where "wP_of_wstate (N, P, Q, n) = P" fun wQ_of_wstate :: "'a wstate \ 'a wclause multiset" where "wQ_of_wstate (N, P, Q, n) = Q" fun n_of_wstate :: "'a wstate \ nat" where "n_of_wstate (N, P, Q, n) = n" lemma of_wstate_split[simp]: "(wN_of_wstate St, wP_of_wstate St, wQ_of_wstate St, n_of_wstate St) = St" by (cases St) auto abbreviation grounding_of_wstate :: "'a wstate \ 'a clause set" where "grounding_of_wstate St \ grounding_of_state (state_of_wstate St)" abbreviation Liminf_wstate :: "'a wstate llist \ 'a state" where "Liminf_wstate Sts \ Liminf_state (lmap state_of_wstate Sts)" lemma timestamp_le_weight: "n \ weight (C, n)" by (induct n, simp, metis weight_mono[of k "Suc k" for k] Suc_le_eq le_less le_trans) inductive weighted_RP :: "'a wstate \ 'a wstate \ bool" (infix "\\<^sub>w" 50) where tautology_deletion: "Neg A \# C \ Pos A \# C \ (N + {#(C, i)#}, P, Q, n) \\<^sub>w (N, P, Q, n)" | forward_subsumption: "D \# image_mset fst (P + Q) \ subsumes D C \ (N + {#(C, i)#}, P, Q, n) \\<^sub>w (N, P, Q, n)" | backward_subsumption_P: "D \# image_mset fst N \ C \# image_mset fst P \ strictly_subsumes D C \ (N, P, Q, n) \\<^sub>w (N, {#(E, k) \# P. E \ C#}, Q, n)" | backward_subsumption_Q: "D \# image_mset fst N \ strictly_subsumes D C \ (N, P, Q + {#(C, i)#}, n) \\<^sub>w (N, P, Q, n)" | forward_reduction: "D + {#L'#} \# image_mset fst (P + Q) \ - L = L' \l \ \ D \ \ \# C \ (N + {#(C + {#L#}, i)#}, P, Q, n) \\<^sub>w (N + {#(C, i)#}, P, Q, n)" | backward_reduction_P: "D + {#L'#} \# image_mset fst N \ - L = L' \l \ \ D \ \ \# C \ (\j. (C + {#L#}, j) \# P \ j \ i) \ (N, P + {#(C + {#L#}, i)#}, Q, n) \\<^sub>w (N, P + {#(C, i)#}, Q, n)" | backward_reduction_Q: "D + {#L'#} \# image_mset fst N \ - L = L' \l \ \ D \ \ \# C \ (N, P, Q + {#(C + {#L#}, i)#}, n) \\<^sub>w (N, P + {#(C, i)#}, Q, n)" | clause_processing: "(N + {#(C, i)#}, P, Q, n) \\<^sub>w (N, P + {#(C, i)#}, Q, n)" | inference_computation: "(\(D, j) \# P. weight (C, i) \ weight (D, j)) \ N = mset_set ((\D. (D, n)) ` concls_of (inference_system.inferences_between (ord_FO_\ S) (set_mset (image_mset fst Q)) C)) \ ({#}, P + {#(C, i)#}, Q, n) \\<^sub>w (N, {#(D, j) \# P. D \ C#}, Q + {#(C, i)#}, Suc n)" lemma weighted_RP_imp_RP: "St \\<^sub>w St' \ state_of_wstate St \ state_of_wstate St'" proof (induction rule: weighted_RP.induct) case (backward_subsumption_P D N C P Q n) show ?case by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\)", OF _ _ RP.backward_subsumption_P[of D "fst ` set_mset N" C "fst ` set_mset P - {C}" "fst ` set_mset Q"]]) (use backward_subsumption_P in auto) next case (inference_computation P C i N n Q) show ?case by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\)", OF _ _ RP.inference_computation[of "fst ` set_mset N" "fst ` set_mset Q" C "fst ` set_mset P - {C}"]], use inference_computation(2) finite_ord_FO_resolution_inferences_between in \auto simp: comp_def image_comp inference_system.inferences_between_def\) qed (use RP.intros in simp_all) lemma final_weighted_RP: "\ ({#}, {#}, Q, n) \\<^sub>w St" by (auto elim: weighted_RP.cases) context fixes Sts :: "'a wstate llist" assumes full_deriv: "full_chain (\\<^sub>w) Sts" and empty_P0: "P_of_wstate (lhd Sts) = {}" and empty_Q0: "Q_of_wstate (lhd Sts) = {}" begin lemma finite_Sts0: "finite (clss_of_wstate (lhd Sts))" - unfolding clss_of_state_def by (cases "lhd Sts") auto + by (cases "lhd Sts") auto lemmas deriv = full_chain_imp_chain[OF full_deriv] lemmas lhd_lmap_Sts = llist.map_sel(1)[OF chain_not_lnull[OF deriv]] lemma deriv_RP: "chain (\) (lmap state_of_wstate Sts)" using deriv weighted_RP_imp_RP by (metis chain_lmap) lemma finite_Sts0_RP: "finite (clss_of_state (lhd (lmap state_of_wstate Sts)))" using finite_Sts0 chain_length_pos[OF deriv] by auto lemma empty_P0_RP: "P_of_state (lhd (lmap state_of_wstate Sts)) = {}" using empty_P0 chain_length_pos[OF deriv] by auto lemma empty_Q0_RP: "Q_of_state (lhd (lmap state_of_wstate Sts)) = {}" using empty_Q0 chain_length_pos[OF deriv] by auto lemmas Sts_thms = deriv_RP finite_Sts0_RP empty_P0_RP empty_Q0_RP theorem weighted_RP_model: "St \\<^sub>w St' \ I \s grounding_of_wstate St' \ I \s grounding_of_wstate St" using RP_model Sts_thms weighted_RP_imp_RP by (simp only: comp_def) abbreviation S_gQ :: "'a clause \ 'a clause" where "S_gQ \ S_Q (lmap state_of_wstate Sts)" interpretation sq: selection S_gQ - unfolding S_Q_def[OF deriv_RP empty_Q0_RP] - using S_M_selects_subseteq S_M_selects_neg_lits selection_axioms + unfolding S_Q_def[OF deriv_RP] using S_M_selects_subseteq S_M_selects_neg_lits selection_axioms by unfold_locales auto interpretation gd: ground_resolution_with_selection S_gQ by unfold_locales interpretation src: standard_redundancy_criterion_reductive gd.ord_\ by unfold_locales interpretation src: standard_redundancy_criterion_counterex_reducing gd.ord_\ "ground_resolution_with_selection.INTERP S_gQ" by unfold_locales lemmas ord_\_saturated_upto_def = src.saturated_upto_def lemmas ord_\_saturated_upto_complete = src.saturated_upto_complete lemmas ord_\_contradiction_Rf = src.contradiction_Rf theorem weighted_RP_sound: assumes "{#} \ clss_of_state (Liminf_wstate Sts)" shows "\ satisfiable (grounding_of_wstate (lhd Sts))" - by (rule RP_sound[OF deriv_RP empty_Q0_RP assms, unfolded lhd_lmap_Sts]) + by (rule RP_sound[OF deriv_RP assms, unfolded lhd_lmap_Sts]) abbreviation RP_filtered_measure :: "('a wclause \ bool) \ 'a wstate \ nat \ nat \ nat" where "RP_filtered_measure \ \p (N, P, Q, n). (sum_mset (image_mset (\(C, i). Suc (size C)) {#Di \# N + P + Q. p Di#}), size {#Di \# N. p Di#}, size {#Di \# P. p Di#})" abbreviation RP_combined_measure :: "nat \ 'a wstate \ nat \ (nat \ nat \ nat) \ (nat \ nat \ nat)" where "RP_combined_measure \ \w St. (w + 1 - n_of_wstate St, RP_filtered_measure (\(C, i). i \ w) St, RP_filtered_measure (\Ci. True) St)" abbreviation (input) RP_filtered_relation :: "((nat \ nat \ nat) \ (nat \ nat \ nat)) set" where "RP_filtered_relation \ natLess <*lex*> natLess <*lex*> natLess" abbreviation (input) RP_combined_relation :: "((nat \ ((nat \ nat \ nat) \ (nat \ nat \ nat))) \ (nat \ ((nat \ nat \ nat) \ (nat \ nat \ nat)))) set" where "RP_combined_relation \ natLess <*lex*> RP_filtered_relation <*lex*> RP_filtered_relation" abbreviation "(fst3 :: 'b * 'c * 'd \ 'b) \ fst" abbreviation "(snd3 :: 'b * 'c * 'd \ 'c) \ \x. fst (snd x)" abbreviation "(trd3 :: 'b * 'c * 'd \ 'd) \ \x. snd (snd x)" lemma wf_RP_filtered_relation: "wf RP_filtered_relation" and wf_RP_combined_relation: "wf RP_combined_relation" unfolding natLess_def using wf_less wf_mult by auto lemma multiset_sum_of_Suc_f_monotone: "N \# M \ (\x \# N. Suc (f x)) < (\x \# M. Suc (f x))" proof (induction N arbitrary: M) case empty then obtain y where "y \# M" by force then have "(\x \# M. 1) = (\x \# M - {#y#} + {#y#}. 1)" by auto also have "... = (\x \# M - {#y#}. 1) + (\x \# {#y#}. 1)" by (metis image_mset_union sum_mset.union) also have "... > (0 :: nat)" by auto finally have "0 < (\x \# M. Suc (f x))" by (fastforce intro: gr_zeroI) then show ?case using empty by auto next case (add x N) from this(2) have "(\y \# N. Suc (f y)) < (\y \# M - {#x#}. Suc (f y))" using add(1)[of "M - {#x#}"] by (simp add: insert_union_subset_iff) moreover have "add_mset x (remove1_mset x M) = M" by (meson add.prems add_mset_remove_trivial_If mset_subset_insertD) ultimately show ?case by (metis (no_types) add.commute add_less_cancel_right sum_mset.insert) qed lemma multiset_sum_monotone_f': assumes "CC \# DD" shows "(\(C, i) \# CC. Suc (f C)) < (\(C, i) \# DD. Suc (f C))" using multiset_sum_of_Suc_f_monotone[OF assms, of "f \ fst"] by (metis (mono_tags) comp_apply image_mset_cong2 split_beta) lemma filter_mset_strict_subset: assumes "x \# M" and "\ p x" shows "{#y \# M. p y#} \# M" proof - have subseteq: "{#E \# M. p E#} \# M" by auto have "count {#E \# M. p E#} x = 0" using assms by auto moreover have "0 < count M x" using assms by auto ultimately have lt_count: "count {#y \# M. p y#} x < count M x" by auto then show ?thesis using subseteq by (metis less_not_refl2 subset_mset.le_neq_trans) qed lemma weighted_RP_measure_decreasing_N: assumes "St \\<^sub>w St'" and "(C, l) \# wN_of_wstate St" shows "(RP_filtered_measure (\Ci. True) St', RP_filtered_measure (\Ci. True) St) \ RP_filtered_relation" using assms proof (induction rule: weighted_RP.induct) case (backward_subsumption_P D N C' P Q n) then obtain i' where "(C', i') \# P" by auto then have "{#(E, k) \# P. E \ C'#} \# P" using filter_mset_strict_subset[of "(C', i')" P "\X. \fst X = C'"] by (metis (mono_tags, lifting) filter_mset_cong fst_conv prod.case_eq_if) then have "(\(C, i) \# {#(E, k) \# P. E \ C'#}. Suc (size C)) < (\(C, i) \# P. Suc (size C))" using multiset_sum_monotone_f'[of "{#(E, k) \# P. E \ C'#}" P size] by metis then show ?case unfolding natLess_def by auto qed (auto simp: natLess_def) lemma weighted_RP_measure_decreasing_P: assumes "St \\<^sub>w St'" and "(C, i) \# wP_of_wstate St" shows "(RP_combined_measure (weight (C, i)) St', RP_combined_measure (weight (C, i)) St) \ RP_combined_relation" using assms proof (induction rule: weighted_RP.induct) case (backward_subsumption_P D N C' P Q n) define St where "St = (N, P, Q, n)" define P' where "P' = {#(E, k) \# P. E \ C'#}" define St' where "St' = (N, P', Q, n)" from backward_subsumption_P obtain i' where "(C', i') \# P" by auto then have P'_sub_P: "P' \# P" unfolding P'_def using filter_mset_strict_subset[of "(C', i')" P "\Dj. fst Dj \ C'"] by (metis (no_types, lifting) filter_mset_cong fst_conv prod.case_eq_if) have P'_subeq_P_filter: "{#(Ca, ia) \# P'. ia \ weight (C, i)#} \# {#(Ca, ia) \# P. ia \ weight (C, i)#}" using P'_sub_P by (auto intro: multiset_filter_mono) have "fst3 (RP_combined_measure (weight (C, i)) St') \ fst3 (RP_combined_measure (weight (C, i)) St)" unfolding St'_def St_def by auto moreover have "(\(C, i) \# {#(Ca, ia) \# P'. ia \ weight (C, i)#}. Suc (size C)) \ (\x \# {#(Ca, ia) \# P. ia \ weight (C, i)#}. case x of (C, i) \ Suc (size C))" using P'_subeq_P_filter by (rule sum_image_mset_mono) then have "fst3 (snd3 (RP_combined_measure (weight (C, i)) St')) \ fst3 (snd3 (RP_combined_measure (weight (C, i)) St))" unfolding St'_def St_def by auto moreover have "snd3 (snd3 (RP_combined_measure (weight (C, i)) St')) \ snd3 (snd3 (RP_combined_measure (weight (C, i)) St))" unfolding St'_def St_def by auto moreover from P'_subeq_P_filter have "size {#(Ca, ia) \# P'. ia \ weight (C, i)#} \ size {#(Ca, ia) \# P. ia \ weight (C, i)#}" by (simp add: size_mset_mono) then have "trd3 (snd3 (RP_combined_measure (weight (C, i)) St')) \ trd3 (snd3 (RP_combined_measure (weight (C, i)) St))" unfolding St'_def St_def unfolding fst_def snd_def by auto moreover from P'_sub_P have "(\(C, i) \# P'. Suc (size C)) < (\(C, i) \# P. Suc (size C))" using multiset_sum_monotone_f'[of "{#(E, k) \# P. E \ C'#}" P size] unfolding P'_def by metis then have "fst3 (trd3 (RP_combined_measure (weight (C, i)) St')) < fst3 (trd3 (RP_combined_measure (weight (C, i)) St))" unfolding P'_def St'_def St_def by auto ultimately show ?case unfolding natLess_def P'_def St'_def St_def by auto next case (inference_computation P C' i' N n Q) then show ?case proof (cases "n \ weight (C, i)") case True then have "weight (C, i) + 1 - n > weight (C, i) + 1 - Suc n" by auto then show ?thesis unfolding natLess_def by auto next case n_nle_w: False define St :: "'a wstate" where "St = ({#}, P + {#(C', i')#}, Q, n)" define St' :: "'a wstate" where "St' = (N, {#(D, j) \# P. D \ C'#}, Q + {#(C', i')#}, Suc n)" define concls :: "'a wclause set" where "concls = (\D. (D, n)) ` concls_of (inference_system.inferences_between (ord_FO_\ S) (fst ` set_mset Q) C')" have fin: "finite concls" unfolding concls_def using finite_ord_FO_resolution_inferences_between by auto have "{(D, ia) \ concls. ia \ weight (C, i)} = {}" unfolding concls_def using n_nle_w by auto then have "{#(D, ia) \# mset_set concls. ia \ weight (C, i)#} = {#}" using fin filter_mset_empty_if_finite_and_filter_set_empty[of concls] by auto then have n_low_weight_empty: "{#(D, ia) \# N. ia \ weight (C, i)#} = {#}" unfolding inference_computation unfolding concls_def by auto have "weight (C', i') \ weight (C, i)" using inference_computation by auto then have i'_le_w_Ci: "i' \ weight (C, i)" using timestamp_le_weight[of i' C'] by auto have subs: "{#(D, ia) \# N + {#(D, j) \# P. D \ C'#} + (Q + {#(C', i')#}). ia \ weight (C, i)#} \# {#(D, ia) \# {#} + (P + {#(C', i')#}) + Q. ia \ weight (C, i)#}" using n_low_weight_empty by (auto simp: multiset_filter_mono) have "fst3 (RP_combined_measure (weight (C, i)) St') \ fst3 (RP_combined_measure (weight (C, i)) St)" unfolding St'_def St_def by auto moreover have "fst (RP_filtered_measure ((\(D, ia). ia \ weight (C, i))) St') = (\(C, i) \# {#(D, ia) \# N + {#(D, j) \# P. D \ C'#} + (Q + {#(C', i')#}). ia \ weight (C, i)#}. Suc (size C))" unfolding St'_def by auto also have "... \ (\(C, i) \# {#(D, ia) \# {#} + (P + {#(C', i')#}) + Q. ia \ weight (C, i)#}. Suc (size C))" using subs sum_image_mset_mono by blast also have "... = fst (RP_filtered_measure (\(D, ia). ia \ weight (C, i)) St)" unfolding St_def by auto finally have "fst3 (snd3 (RP_combined_measure (weight (C, i)) St')) \ fst3 (snd3 (RP_combined_measure (weight (C, i)) St))" by auto moreover have "snd3 (snd3 (RP_combined_measure (weight (C, i)) St')) = snd3 (snd3 (RP_combined_measure (weight (C, i)) St))" unfolding St_def St'_def using n_low_weight_empty by auto moreover have "trd3 (snd3 (RP_combined_measure (weight (C, i)) St')) < trd3 (snd3 (RP_combined_measure (weight (C, i)) St))" unfolding St_def St'_def using i'_le_w_Ci by (simp add: le_imp_less_Suc multiset_filter_mono size_mset_mono) ultimately show ?thesis unfolding natLess_def St'_def St_def lex_prod_def by force qed qed (auto simp: natLess_def) lemma preserve_min_or_delete_completely: assumes "St \\<^sub>w St'" "(C, i) \# wP_of_wstate St" "\k. (C, k) \# wP_of_wstate St \ i \ k" shows "(C, i) \# wP_of_wstate St' \ (\j. (C, j) \# wP_of_wstate St')" using assms proof (induction rule: weighted_RP.induct) case (backward_reduction_P D L' N L \ C' P i' Q n) show ?case proof (cases "C = C' + {#L#}") case True_outer: True then have C_i_in: "(C, i) \# P + {#(C, i')#}" using backward_reduction_P by auto then have max: "\k. (C, k) \# P + {#(C, i')#} \ k \ i'" using backward_reduction_P unfolding True_outer[symmetric] by auto then have "count (P + {#(C, i')#}) (C, i') \ 1" by auto moreover { assume asm: "count (P + {#(C, i')#}) (C, i') = 1" then have nin_P: "(C, i') \# P" using not_in_iff by force have ?thesis proof (cases "(C, i) = (C, i')") case True then have "i = i'" by auto then have "\j. (C, j) \# P + {#(C, i')#} \ j = i'" using max backward_reduction_P(6) unfolding True_outer[symmetric] by force then show ?thesis using True_outer[symmetric] nin_P by auto next case False then show ?thesis using C_i_in by auto qed } moreover { assume "count (P + {#(C, i')#}) (C, i') > 1" then have ?thesis using C_i_in by auto } ultimately show ?thesis by (cases "count (P + {#(C, i')#}) (C, i') = 1") auto next case False then show ?thesis using backward_reduction_P by auto qed qed auto lemma preserve_min_P: assumes "St \\<^sub>w St'" "(C, j) \# wP_of_wstate St'" and "(C, i) \# wP_of_wstate St" and "\k. (C, k) \# wP_of_wstate St \ i \ k" shows "(C, i) \# wP_of_wstate St'" using assms preserve_min_or_delete_completely by blast lemma preserve_min_P_Sts: assumes "enat (Suc k) < llength Sts" and "(C, i) \# wP_of_wstate (lnth Sts k)" and "(C, j) \# wP_of_wstate (lnth Sts (Suc k))" and "\j. (C, j) \# wP_of_wstate (lnth Sts k) \ i \ j" shows "(C, i) \# wP_of_wstate (lnth Sts (Suc k))" using deriv assms chain_lnth_rel preserve_min_P by metis lemma in_lnth_in_Supremum_ldrop: assumes "i < llength xs" and "x \# (lnth xs i)" shows "x \ Sup_llist (lmap set_mset (ldrop (enat i) xs))" using assms by (metis (no_types) ldrop_eq_LConsD ldropn_0 llist.simps(13) contra_subsetD ldrop_enat ldropn_Suc_conv_ldropn lnth_0 lnth_lmap lnth_subset_Sup_llist) lemma persistent_wclause_in_P_if_persistent_clause_in_P: assumes "C \ Liminf_llist (lmap P_of_state (lmap state_of_wstate Sts))" shows "\i. (C, i) \ Liminf_llist (lmap (set_mset \ wP_of_wstate) Sts)" proof - obtain t_C where t_C_p: "enat t_C < llength Sts" "\t. t_C \ t \ t < llength Sts \ C \ P_of_state (state_of_wstate (lnth Sts t))" using assms unfolding Liminf_llist_def by auto then obtain i where i_p: "(C, i) \# wP_of_wstate (lnth Sts t_C)" using t_C_p by (cases "lnth Sts t_C") force have Ci_in_nth_wP: "\i. (C, i) \# wP_of_wstate (lnth Sts (t_C + t))" if "t_C + t < llength Sts" for t using that t_C_p(2)[of "t_C + _"] by (cases "lnth Sts (t_C + t)") force define in_Sup_wP :: "nat \ bool" where "in_Sup_wP = (\i. (C, i) \ Sup_llist (lmap (set_mset \ wP_of_wstate) (ldrop t_C Sts)))" have "in_Sup_wP i" using i_p assms(1) in_lnth_in_Supremum_ldrop[of t_C "lmap wP_of_wstate Sts" "(C, i)"] t_C_p by (simp add: in_Sup_wP_def llist.map_comp) then obtain j where j_p: "is_least in_Sup_wP j" unfolding in_Sup_wP_def[symmetric] using least_exists by metis then have "\i. (C, i) \ Sup_llist (lmap (set_mset \ wP_of_wstate) (ldrop t_C Sts)) \ j \ i" unfolding is_least_def in_Sup_wP_def using not_less by blast then have j_smallest: "\i t. enat (t_C + t) < llength Sts \ (C, i) \# wP_of_wstate (lnth Sts (t_C + t)) \ j \ i" unfolding comp_def by (smt add.commute ldrop_enat ldrop_eq_LConsD ldrop_ldrop ldropn_Suc_conv_ldropn plus_enat_simps(1) lnth_ldropn Sup_llist_def UN_I ldrop_lmap llength_lmap lnth_lmap mem_Collect_eq) from j_p have "\t_Cj. t_Cj < llength (ldrop (enat t_C) Sts) \ (C, j) \# wP_of_wstate (lnth (ldrop t_C Sts) t_Cj)" unfolding in_Sup_wP_def Sup_llist_def is_least_def by simp then obtain t_Cj where j_p: "(C,j) \# wP_of_wstate (lnth Sts (t_C + t_Cj))" "enat (t_C + t_Cj) < llength Sts" by (smt add.commute ldrop_enat ldrop_eq_LConsD ldrop_ldrop ldropn_Suc_conv_ldropn plus_enat_simps(1) lhd_ldropn) have Ci_stays: "t_C + t_Cj + t < llength Sts \ (C,j) \# wP_of_wstate (lnth Sts (t_C + t_Cj + t))" for t proof (induction t) case 0 then show ?case using j_p by (simp add: add.commute) next case (Suc t) have any_Ck_in_wP: "j \ k" if "(C, k) \# wP_of_wstate (lnth Sts (t_C + t_Cj + t))" for k using that j_p j_smallest Suc by (smt Suc_ile_eq add.commute add.left_commute add_Suc less_imp_le plus_enat_simps(1) the_enat.simps) from Suc have Cj_in_wP: "(C, j) \# wP_of_wstate (lnth Sts (t_C + t_Cj + t))" by (metis (no_types, hide_lams) Suc_ile_eq add.commute add_Suc_right less_imp_le) moreover have "C \ P_of_state (state_of_wstate (lnth Sts (Suc (t_C + t_Cj + t))))" using t_C_p(2) Suc.prems by auto then have "\k. (C, k) \# wP_of_wstate (lnth Sts (Suc (t_C + t_Cj + t)))" by (smt Suc.prems Ci_in_nth_wP add.commute add.left_commute add_Suc_right enat_ord_code(4)) ultimately have "(C, j) \# wP_of_wstate (lnth Sts (Suc (t_C + t_Cj + t)))" using preserve_min_P_Sts Cj_in_wP any_Ck_in_wP Suc.prems by force then have "(C, j) \# lnth (lmap wP_of_wstate Sts) (Suc (t_C + t_Cj + t))" using Suc.prems by auto then show ?case by (smt Suc.prems add.commute add_Suc_right lnth_lmap) qed then have "(\t. t_C + t_Cj \ t \ t < llength (lmap (set_mset \ wP_of_wstate) Sts) \ (C, j) \# wP_of_wstate (lnth Sts t))" using Ci_stays[of "_ - (t_C + t_Cj)"] by (metis le_add_diff_inverse llength_lmap) then have "(C, j) \ Liminf_llist (lmap (set_mset \ wP_of_wstate) Sts)" unfolding Liminf_llist_def using j_p by auto then show "\i. (C, i) \ Liminf_llist (lmap (set_mset \ wP_of_wstate) Sts)" by auto qed lemma lfinite_not_LNil_nth_llast: assumes "lfinite Sts" and "Sts \ LNil" shows "\i < llength Sts. lnth Sts i = llast Sts \ (\j < llength Sts. j \ i)" using assms proof (induction rule: lfinite.induct) case (lfinite_LConsI xs x) then show ?case proof (cases "xs = LNil") case True show ?thesis using True zero_enat_def by auto next case False then obtain i where i_p: "enat i < llength xs \ lnth xs i = llast xs \ (\j < llength xs. j \ enat i)" using lfinite_LConsI by auto then have "enat (Suc i) < llength (LCons x xs)" by (simp add: Suc_ile_eq) moreover from i_p have "lnth (LCons x xs) (Suc i) = llast (LCons x xs)" by (metis gr_implies_not_zero llast_LCons llength_lnull lnth_Suc_LCons) moreover from i_p have "\j < llength (LCons x xs). j \ enat (Suc i)" by (metis antisym_conv2 eSuc_enat eSuc_ile_mono ileI1 iless_Suc_eq llength_LCons) ultimately show ?thesis by auto qed qed auto lemma fair_if_finite: assumes fin: "lfinite Sts" shows "fair_state_seq (lmap state_of_wstate Sts)" proof (rule ccontr) assume unfair: "\ fair_state_seq (lmap state_of_wstate Sts)" have no_inf_from_last: "\y. \ llast Sts \\<^sub>w y" using fin full_chain_iff_chain[of "(\\<^sub>w)" Sts] full_deriv by auto from unfair obtain C where "C \ Liminf_llist (lmap N_of_state (lmap state_of_wstate Sts)) \ Liminf_llist (lmap P_of_state (lmap state_of_wstate Sts))" unfolding fair_state_seq_def Liminf_state_def by auto then obtain i where i_p: "enat i < llength Sts" "\j. i \ j \ enat j < llength Sts \ C \ N_of_state (state_of_wstate (lnth Sts j)) \ P_of_state (state_of_wstate (lnth Sts j))" unfolding Liminf_llist_def by auto have C_in_llast: "C \ N_of_state (state_of_wstate (llast Sts)) \ P_of_state (state_of_wstate (llast Sts))" proof - obtain l where l_p: "enat l < llength Sts \ lnth Sts l = llast Sts \ (\j < llength Sts. j \ enat l)" using fin lfinite_not_LNil_nth_llast i_p(1) by fastforce then have "C \ N_of_state (state_of_wstate (lnth Sts l)) \ P_of_state (state_of_wstate (lnth Sts l))" using i_p(1) i_p(2)[of l] by auto then show ?thesis using l_p by auto qed define N :: "'a wclause multiset" where "N = wN_of_wstate (llast Sts)" define P :: "'a wclause multiset" where "P = wP_of_wstate (llast Sts)" define Q :: "'a wclause multiset" where "Q = wQ_of_wstate (llast Sts)" define n :: nat where "n = n_of_wstate (llast Sts)" { assume "N_of_state (state_of_wstate (llast Sts)) \ {}" then obtain D j where "(D, j) \# N" unfolding N_def by (cases "llast Sts") auto then have "llast Sts \\<^sub>w (N - {#(D, j)#}, P + {#(D, j)#}, Q, n)" using weighted_RP.clause_processing[of "N - {#(D, j)#}" D j P Q n] unfolding N_def P_def Q_def n_def by auto then have "\St'. llast Sts \\<^sub>w St'" by auto } moreover { assume a: "N_of_state (state_of_wstate (llast Sts)) = {}" then have b: "N = {#}" unfolding N_def by (cases "llast Sts") auto from a have "C \ P_of_state (state_of_wstate (llast Sts))" using C_in_llast by auto then obtain D j where "(D, j) \# P" unfolding P_def by (cases "llast Sts") auto then have "weight (D, j) \ weight ` set_mset P" by auto then have "\w. is_least (\w. w \ (weight ` set_mset P)) w" using least_exists by auto then have "\D j. (\(D', j') \# P. weight (D, j) \ weight (D', j')) \ (D, j) \# P" using assms linorder_not_less unfolding is_least_def by (auto 6 0) then obtain D j where min: "(\(D', j') \# P. weight (D, j) \ weight (D', j'))" and Dj_in_p: "(D, j) \# P" by auto from min have min: "(\(D', j') \# P - {#(D, j)#}. weight (D, j) \ weight (D', j'))" using mset_subset_diff_self[OF Dj_in_p] by auto define N' where "N' = mset_set ((\D'. (D', n)) ` concls_of (inference_system.inferences_between (ord_FO_\ S) (set_mset (image_mset fst Q)) D))" have "llast Sts \\<^sub>w (N', {#(D', j') \# P - {#(D, j)#}. D' \ D#}, Q + {#(D,j)#}, Suc n)" using weighted_RP.inference_computation[of "P - {#(D, j)#}" D j N' n Q, OF min N'_def] of_wstate_split[symmetric, of "llast Sts"] Dj_in_p unfolding N_def[symmetric] P_def[symmetric] Q_def[symmetric] n_def[symmetric] b by auto then have "\St'. llast Sts \\<^sub>w St'" by auto } ultimately have "\St'. llast Sts \\<^sub>w St'" by auto then show False using no_inf_from_last by metis qed lemma N_of_state_state_of_wstate_wN_of_wstate: assumes "C \ N_of_state (state_of_wstate St)" shows "\i. (C, i) \# wN_of_wstate St" by (smt N_of_state.elims assms eq_fst_iff fstI fst_conv image_iff of_wstate_split set_image_mset state_of_wstate.simps) lemma in_wN_of_wstate_in_N_of_wstate: "(C, i) \# wN_of_wstate St \ C \ N_of_wstate St" by (metis (mono_guards_query_query) N_of_state.simps fst_conv image_eqI of_wstate_split set_image_mset state_of_wstate.simps) lemma in_wP_of_wstate_in_P_of_wstate: "(C, i) \# wP_of_wstate St \ C \ P_of_wstate St" by (metis (mono_guards_query_query) P_of_state.simps fst_conv image_eqI of_wstate_split set_image_mset state_of_wstate.simps) lemma in_wQ_of_wstate_in_Q_of_wstate: "(C, i) \# wQ_of_wstate St \ C \ Q_of_wstate St" by (metis (mono_guards_query_query) Q_of_state.simps fst_conv image_eqI of_wstate_split set_image_mset state_of_wstate.simps) lemma n_of_wstate_weighted_RP_increasing: "St \\<^sub>w St' \ n_of_wstate St \ n_of_wstate St'" by (induction rule: weighted_RP.induct) auto lemma nth_of_wstate_monotonic: assumes "j < llength Sts" and "i \ j" shows "n_of_wstate (lnth Sts i) \ n_of_wstate (lnth Sts j)" using assms proof (induction "j - i" arbitrary: i) case (Suc x) then have "x = j - (i + 1)" by auto then have "n_of_wstate (lnth Sts (i + 1)) \ n_of_wstate (lnth Sts j)" using Suc by auto moreover have "i < j" using Suc by auto then have "Suc i < llength Sts" using Suc by (metis enat_ord_simps(2) le_less_Suc_eq less_le_trans not_le) then have "lnth Sts i \\<^sub>w lnth Sts (Suc i)" using deriv chain_lnth_rel[of "(\\<^sub>w)" Sts i] by auto then have "n_of_wstate (lnth Sts i) \ n_of_wstate (lnth Sts (i + 1))" using n_of_wstate_weighted_RP_increasing[of "lnth Sts i" "lnth Sts (i + 1)"] by auto ultimately show ?case by auto qed auto lemma infinite_chain_relation_measure: assumes measure_decreasing: "\St St'. P St \ R St St' \ (m St', m St) \ mR" and non_infer_chain: "chain R (ldrop (enat k) Sts)" and inf: "llength Sts = \" and P: "\i. P (lnth (ldrop (enat k) Sts) i)" shows "chain (\x y. (x, y) \ mR)\\ (lmap m (ldrop (enat k) Sts))" proof (rule lnth_rel_chain) show "\ lnull (lmap m (ldrop (enat k) Sts))" using assms by auto next from inf have ldrop_inf: "llength (ldrop (enat k) Sts) = \ \ \ lfinite (ldrop (enat k) Sts)" using inf by (auto simp: llength_eq_infty_conv_lfinite) { fix j :: "nat" define St where "St = lnth (ldrop (enat k) Sts) j" define St' where "St' = lnth (ldrop (enat k) Sts) (j + 1)" have P': "P St \ P St'" unfolding St_def St'_def using P by auto from ldrop_inf have "R St St'" unfolding St_def St'_def using non_infer_chain infinite_chain_lnth_rel[of "ldrop (enat k) Sts" R j] by auto then have "(m St', m St) \ mR" using measure_decreasing P' by auto then have "(lnth (lmap m (ldrop (enat k) Sts)) (j + 1), lnth (lmap m (ldrop (enat k) Sts)) j) \ mR" unfolding St_def St'_def using lnth_lmap by (smt enat.distinct(1) enat_add_left_cancel enat_ord_simps(4) inf ldrop_lmap llength_lmap lnth_ldrop plus_enat_simps(3)) } then show "\j. enat (j + 1) < llength (lmap m (ldrop (enat k) Sts)) \ (\x y. (x, y) \ mR)\\ (lnth (lmap m (ldrop (enat k) Sts)) j) (lnth (lmap m (ldrop (enat k) Sts)) (j + 1))" by blast qed theorem weighted_RP_fair: "fair_state_seq (lmap state_of_wstate Sts)" proof (rule ccontr) assume asm: "\ fair_state_seq (lmap state_of_wstate Sts)" then have inff: "\ lfinite Sts" using fair_if_finite by auto then have inf: "llength Sts = \" using llength_eq_infty_conv_lfinite by auto from asm obtain C where "C \ Liminf_llist (lmap N_of_state (lmap state_of_wstate Sts)) \ Liminf_llist (lmap P_of_state (lmap state_of_wstate Sts))" unfolding fair_state_seq_def Liminf_state_def by auto then show False proof assume "C \ Liminf_llist (lmap N_of_state (lmap state_of_wstate Sts))" then obtain x where "enat x < llength Sts" "\xa. x \ xa \ enat xa < llength Sts \ C \ N_of_state (state_of_wstate (lnth Sts xa))" unfolding Liminf_llist_def by auto then have "\k. \j. k \ j \ (\i. (C, i) \# wN_of_wstate (lnth Sts j))" unfolding Liminf_llist_def by (force simp add: inf N_of_state_state_of_wstate_wN_of_wstate) then obtain k where k_p: "\j. k \ j \ \i. (C, i) \# wN_of_wstate (lnth Sts j)" unfolding Liminf_llist_def by auto have chain_drop_Sts: "chain (\\<^sub>w) (ldrop k Sts)" using deriv inf inff inf_chain_ldrop_chain by auto have in_N_j: "\j. \i. (C, i) \# wN_of_wstate (lnth (ldrop k Sts) j)" using k_p by (simp add: add.commute inf) then have "chain (\x y. (x, y) \ RP_filtered_relation)\\ (lmap (RP_filtered_measure (\Ci. True)) (ldrop k Sts))" using inff inf weighted_RP_measure_decreasing_N chain_drop_Sts infinite_chain_relation_measure[of "\St. \i. (C, i) \# wN_of_wstate St" "(\\<^sub>w)"] by blast then show False using wfP_iff_no_infinite_down_chain_llist[of "\x y. (x, y) \ RP_filtered_relation"] wf_RP_filtered_relation inff by (metis (no_types, lifting) inf_llist_lnth ldrop_enat_inf_llist lfinite_inf_llist lfinite_lmap wfPUNIVI wf_induct_rule) next assume asm: "C \ Liminf_llist (lmap P_of_state (lmap state_of_wstate Sts))" from asm obtain i where i_p: "enat i < llength Sts" "\j. i \ j \ enat j < llength Sts \ C \ P_of_state (state_of_wstate (lnth Sts j))" unfolding Liminf_llist_def by auto then obtain i where "(C, i) \ Liminf_llist (lmap (set_mset \ wP_of_wstate) Sts)" using persistent_wclause_in_P_if_persistent_clause_in_P[of C] using asm inf by auto then have "\l. \k \ l. (C, i) \ (set_mset \ wP_of_wstate) (lnth Sts k)" unfolding Liminf_llist_def using inff inf by auto then obtain k where k_p: "(\k'\k. (C, i) \ (set_mset \ wP_of_wstate) (lnth Sts k'))" by blast have Ci_in: "\k'. (C, i) \ (set_mset \ wP_of_wstate) (lnth (ldrop k Sts) k')" using k_p lnth_ldrop[of k _ Sts] inf inff by force then have Ci_inn: "\k'. (C, i) \# (wP_of_wstate) (lnth (ldrop k Sts) k')" by auto have "chain (\\<^sub>w) (ldrop k Sts)" using deriv inf_chain_ldrop_chain inf inff by auto then have "chain (\x y. (x, y) \ RP_combined_relation)\\ (lmap (RP_combined_measure (weight (C, i))) (ldrop k Sts))" using inff inf Ci_in weighted_RP_measure_decreasing_P infinite_chain_relation_measure[of "\St. (C, i) \# wP_of_wstate St" "(\\<^sub>w)" "RP_combined_measure (weight (C, i))" ] by auto then show False using wfP_iff_no_infinite_down_chain_llist[of "\x y. (x, y) \ RP_combined_relation"] wf_RP_combined_relation inff by (smt inf_llist_lnth ldrop_enat_inf_llist lfinite_inf_llist lfinite_lmap wfPUNIVI wf_induct_rule) qed qed corollary weighted_RP_saturated: "src.saturated_upto (Liminf_llist (lmap grounding_of_wstate Sts))" - using RP_saturated_if_fair[OF deriv_RP empty_Q0_RP weighted_RP_fair, unfolded llist.map_comp] + using RP_saturated_if_fair[OF deriv_RP weighted_RP_fair empty_Q0_RP, unfolded llist.map_comp] by simp corollary weighted_RP_complete: "\ satisfiable (grounding_of_wstate (lhd Sts)) \ {#} \ Q_of_state (Liminf_wstate Sts)" - using RP_complete_if_fair[OF deriv_RP empty_Q0_RP weighted_RP_fair, simplified lhd_lmap_Sts] - by simp + using RP_complete_if_fair[OF deriv_RP weighted_RP_fair empty_Q0_RP, simplified lhd_lmap_Sts] . end end locale weighted_FO_resolution_prover_with_size_timestamp_factors = FO_resolution_prover S subst_atm id_subst comp_subst renamings_apart atm_of_atms mgu less_atm for S :: "('a :: wellorder) clause \ 'a clause" and subst_atm :: "'a \ 's \ 'a" and id_subst :: "'s" and comp_subst :: "'s \ 's \ 's" and renamings_apart :: "'a literal multiset list \ 's list" and atm_of_atms :: "'a list \ 'a" and mgu :: "'a set set \ 's option" and less_atm :: "'a \ 'a \ bool" + fixes size_atm :: "'a \ nat" and size_factor :: nat and timestamp_factor :: nat assumes timestamp_factor_pos: "timestamp_factor > 0" begin fun weight :: "'a wclause \ nat" where "weight (C, i) = size_factor * size_multiset (size_literal size_atm) C + timestamp_factor * i" lemma weight_mono: "i < j \ weight (C, i) < weight (C, j)" using timestamp_factor_pos by simp declare weight.simps [simp del] sublocale wrp: weighted_FO_resolution_prover _ _ _ _ _ _ _ _ weight by unfold_locales (rule weight_mono) notation wrp.weighted_RP (infix "\\<^sub>w" 50) end end diff --git a/thys/Nested_Multisets_Ordinals/Multiset_More.thy b/thys/Nested_Multisets_Ordinals/Multiset_More.thy --- a/thys/Nested_Multisets_Ordinals/Multiset_More.thy +++ b/thys/Nested_Multisets_Ordinals/Multiset_More.thy @@ -1,986 +1,993 @@ (* Title: More about Multisets Author: Mathias Fleury , 2015 Author: Jasmin Blanchette , 2014, 2015 Author: Anders Schlichtkrull , 2017 Author: Dmitriy Traytel , 2014 Maintainer: Mathias Fleury *) section \More about Multisets\ theory Multiset_More imports "HOL-Library.Multiset_Order" begin text \ Isabelle's theory of finite multisets is not as developed as other areas, such as lists and sets. The present theory introduces some missing concepts and lemmas. Some of it is expected to move to Isabelle's library. \ subsection \Basic Setup\ declare diff_single_trivial [simp] in_image_mset [iff] image_mset.compositionality [simp] (*To have the same rules as the set counter-part*) mset_subset_eqD[dest, intro?] (*@{thm subsetD}*) Multiset.in_multiset_in_set[simp] inter_add_left1[simp] inter_add_left2[simp] inter_add_right1[simp] inter_add_right2[simp] sum_mset_sum_list[simp] subsection \Lemmas about Intersection, Union and Pointwise Inclusion\ lemma subset_add_mset_notin_subset_mset: \A \# add_mset b B \ b \# A \ A \# B\ by (simp add: subset_mset.le_iff_sup) lemma subset_msetE: "\A \# B; \A \# B; \ B \# A\ \ R\ \ R" by (simp add: subset_mset.less_le_not_le) lemma Diff_triv_mset: "M \# N = {#} \ M - N = M" by (metis diff_intersect_left_idem diff_zero) lemma diff_intersect_sym_diff: "(A - B) \# (B - A) = {#}" unfolding multiset_inter_def proof - have "A - (B - (B - A)) = A - B" by (metis diff_intersect_right_idem multiset_inter_def) then show "A - B - (A - B - (B - A)) = {#}" by (metis diff_add diff_cancel diff_subset_eq_self subset_mset.diff_add) qed declare subset_msetE [elim!] subsection \Lemmas about Filter and Image\ lemma count_image_mset_ge_count: "count (image_mset f A) (f b) \ count A b" by (induction A) auto lemma count_image_mset_inj: assumes \inj f\ shows \count (image_mset f M) (f x) = count M x\ by (induct M) (use assms in \auto simp: inj_on_def\) lemma count_image_mset_le_count_inj_on: "inj_on f (set_mset M) \ count (image_mset f M) y \ count M (inv_into (set_mset M) f y)" proof (induct M) case (add x M) note ih = this(1) and inj_xM = this(2) have inj_M: "inj_on f (set_mset M)" using inj_xM by simp show ?case proof (cases "x \# M") case x_in_M: True show ?thesis proof (cases "y = f x") case y_eq_fx: True show ?thesis using x_in_M ih[OF inj_M] unfolding y_eq_fx by (simp add: inj_M insert_absorb) next case y_ne_fx: False show ?thesis using x_in_M ih[OF inj_M] y_ne_fx insert_absorb by fastforce qed next case x_ni_M: False show ?thesis proof (cases "y = f x") case y_eq_fx: True have "f x \# image_mset f M" using x_ni_M inj_xM by force thus ?thesis unfolding y_eq_fx by (metis (no_types) inj_xM count_add_mset count_greater_eq_Suc_zero_iff count_inI image_mset_add_mset inv_into_f_f union_single_eq_member) next case y_ne_fx: False show ?thesis proof (rule ccontr) assume neg_conj: "\ count (image_mset f (add_mset x M)) y \ count (add_mset x M) (inv_into (set_mset (add_mset x M)) f y)" have cnt_y: "count (add_mset (f x) (image_mset f M)) y = count (image_mset f M) y" using y_ne_fx by simp have "inv_into (set_mset M) f y \# add_mset x M \ inv_into (set_mset (add_mset x M)) f (f (inv_into (set_mset M) f y)) = inv_into (set_mset M) f y" by (meson inj_xM inv_into_f_f) hence "0 < count (image_mset f (add_mset x M)) y \ count M (inv_into (set_mset M) f y) = 0 \ x = inv_into (set_mset M) f y" using neg_conj cnt_y ih[OF inj_M] by (metis (no_types) count_add_mset count_greater_zero_iff count_inI f_inv_into_f image_mset_add_mset set_image_mset) thus False using neg_conj cnt_y x_ni_M ih[OF inj_M] by (metis (no_types) count_greater_zero_iff count_inI eq_iff image_mset_add_mset less_imp_le) qed qed qed qed simp lemma mset_filter_compl: "mset (filter p xs) + mset (filter (Not \ p) xs) = mset xs" by (induction xs) (auto simp: mset_filter ac_simps) text \Near duplicate of @{thm [source] filter_eq_replicate_mset}: @{thm filter_eq_replicate_mset}.\ lemma filter_mset_eq: "filter_mset ((=) L) A = replicate_mset (count A L) L" by (auto simp: multiset_eq_iff) lemma filter_mset_cong[fundef_cong]: assumes "M = M'" "\a. a \# M \ P a = Q a" shows "filter_mset P M = filter_mset Q M" proof - have "M - filter_mset Q M = filter_mset (\a. \Q a) M" by (metis multiset_partition add_diff_cancel_left') then show ?thesis by (auto simp: filter_mset_eq_conv assms) qed lemma image_mset_filter_swap: "image_mset f {# x \# M. P (f x)#} = {# x \# image_mset f M. P x#}" by (induction M) auto lemma image_mset_cong2: "(\x. x \# M \ f x = g x) \ M = N \ image_mset f M = image_mset g N" by (hypsubst, rule image_mset_cong) lemma filter_mset_empty_conv: \(filter_mset P M = {#}) = (\L\#M. \ P L)\ by (induction M) auto lemma multiset_filter_mono2: \filter_mset P A \# filter_mset Q A \ (\a\#A. P a \ Q a)\ by (induction A) (auto intro: subset_mset.order.trans) lemma image_filter_cong: assumes \\C. C \# M \ P C \ f C = g C\ shows \{#f C. C \# {#C \# M. P C#}#} = {#g C | C\# M. P C#}\ using assms by (induction M) auto lemma image_mset_filter_swap2: \{#C \# {#P x. x \# D#}. Q C #} = {#P x. x \# {#C| C \# D. Q (P C)#}#}\ by (simp add: image_mset_filter_swap) declare image_mset_cong2 [cong] subsection \Lemmas about Sum\ lemma sum_image_mset_sum_map[simp]: "sum_mset (image_mset f (mset xs)) = sum_list (map f xs)" by (metis mset_map sum_mset_sum_list) lemma sum_image_mset_mono: fixes f :: "'a \ 'b::canonically_ordered_monoid_add" assumes sub: "A \# B" shows "(\m \# A. f m) \ (\m \# B. f m)" by (metis image_mset_union le_iff_add sub subset_mset.add_diff_inverse sum_mset.union) lemma sum_image_mset_mono_mem: "n \# M \ f n \ (\m \# M. f m)" for f :: "'a \ 'b::canonically_ordered_monoid_add" using le_iff_add multi_member_split by fastforce lemma count_sum_mset_if_1_0: \count M a = (\x\#M. if x = a then 1 else 0)\ by (induction M) auto lemma sum_mset_dvd: fixes k :: "'a::comm_semiring_1_cancel" assumes "\m \# M. k dvd f m" shows "k dvd (\m \# M. f m)" using assms by (induct M) auto lemma sum_mset_distrib_div_if_dvd: fixes k :: "'a::unique_euclidean_semiring" assumes "\m \# M. k dvd f m" shows "(\m \# M. f m) div k = (\m \# M. f m div k)" using assms by (induct M) (auto simp: div_plus_div_distrib_dvd_left) subsection \Lemmas about Remove\ lemma set_mset_minus_replicate_mset[simp]: "n \ count A a \ set_mset (A - replicate_mset n a) = set_mset A - {a}" "n < count A a \ set_mset (A - replicate_mset n a) = set_mset A" unfolding set_mset_def by (auto split: if_split simp: not_in_iff) abbreviation removeAll_mset :: "'a \ 'a multiset \ 'a multiset" where "removeAll_mset C M \ M - replicate_mset (count M C) C" lemma mset_removeAll[simp, code]: "removeAll_mset C (mset L) = mset (removeAll C L)" by (induction L) (auto simp: ac_simps multiset_eq_iff split: if_split_asm) lemma removeAll_mset_filter_mset: "removeAll_mset C M = filter_mset ((\) C) M" by (induction M) (auto simp: ac_simps multiset_eq_iff) abbreviation remove1_mset :: "'a \ 'a multiset \ 'a multiset" where "remove1_mset C M \ M - {#C#}" lemma removeAll_subseteq_remove1_mset: "removeAll_mset x M \# remove1_mset x M" by (auto simp: subseteq_mset_def) lemma in_remove1_mset_neq: assumes ab: "a \ b" shows "a \# remove1_mset b C \ a \# C" by (metis assms diff_single_trivial in_diffD insert_DiffM insert_noteq_member) lemma size_mset_removeAll_mset_le_iff: "size (removeAll_mset x M) < size M \ x \# M" by (auto intro: count_inI mset_subset_size simp: subset_mset_def multiset_eq_iff) lemma size_remove1_mset_If: \size (remove1_mset x M) = size M - (if x \# M then 1 else 0)\ by (auto simp: size_Diff_subset_Int) lemma size_mset_remove1_mset_le_iff: "size (remove1_mset x M) < size M \ x \# M" using less_irrefl by (fastforce intro!: mset_subset_size elim: in_countE simp: subset_mset_def multiset_eq_iff) lemma remove_1_mset_id_iff_notin: "remove1_mset a M = M \ a \# M" by (meson diff_single_trivial multi_drop_mem_not_eq) lemma id_remove_1_mset_iff_notin: "M = remove1_mset a M \ a \# M" using remove_1_mset_id_iff_notin by metis lemma remove1_mset_eqE: "remove1_mset L x1 = M \ (L \# x1 \ x1 = M + {#L#} \ P) \ (L \# x1 \ x1 = M \ P) \ P" by (cases "L \# x1") auto lemma image_filter_ne_mset[simp]: "image_mset f {#x \# M. f x \ y#} = removeAll_mset y (image_mset f M)" by (induction M) simp_all lemma image_mset_remove1_mset_if: "image_mset f (remove1_mset a M) = (if a \# M then remove1_mset (f a) (image_mset f M) else image_mset f M)" by (auto simp: image_mset_Diff) lemma filter_mset_neq: "{#x \# M. x \ y#} = removeAll_mset y M" by (metis add_diff_cancel_left' filter_eq_replicate_mset multiset_partition) lemma filter_mset_neq_cond: "{#x \# M. P x \ x \ y#} = removeAll_mset y {# x\#M. P x#}" by (metis filter_filter_mset filter_mset_neq) lemma remove1_mset_add_mset_If: "remove1_mset L (add_mset L' C) = (if L = L' then C else remove1_mset L C + {#L'#})" by (auto simp: multiset_eq_iff) lemma minus_remove1_mset_if: "A - remove1_mset b B = (if b \# B \ b \# A \ count A b \ count B b then {#b#} + (A - B) else A - B)" by (auto simp: multiset_eq_iff count_greater_zero_iff[symmetric] simp del: count_greater_zero_iff) lemma add_mset_eq_add_mset_ne: "a \ b \ add_mset a A = add_mset b B \ a \# B \ b \# A \ A = add_mset b (B - {#a#})" by (metis (no_types, lifting) diff_single_eq_union diff_union_swap multi_self_add_other_not_self remove_1_mset_id_iff_notin union_single_eq_diff) lemma add_mset_eq_add_mset: \add_mset a M = add_mset b M' \ (a = b \ M = M') \ (a \ b \ b \# M \ add_mset a (M - {#b#}) = M')\ by (metis add_mset_eq_add_mset_ne add_mset_remove_trivial union_single_eq_member) (* TODO move to Multiset: could replace add_mset_remove_trivial_eq? *) lemma add_mset_remove_trivial_iff: \N = add_mset a (N - {#b#}) \ a \# N \ a = b\ by (metis add_left_cancel add_mset_remove_trivial insert_DiffM2 single_eq_single size_mset_remove1_mset_le_iff union_single_eq_member) lemma trivial_add_mset_remove_iff: \add_mset a (N - {#b#}) = N \ a \# N \ a = b\ by (subst eq_commute) (fact add_mset_remove_trivial_iff) lemma remove1_single_empty_iff[simp]: \remove1_mset L {#L'#} = {#} \ L = L'\ using add_mset_remove_trivial_iff by fastforce lemma add_mset_less_imp_less_remove1_mset: assumes xM_lt_N: "add_mset x M < N" shows "M < remove1_mset x N" proof - have "M < N" using assms le_multiset_right_total mset_le_trans by blast then show ?thesis by (metis add_less_cancel_right add_mset_add_single diff_single_trivial insert_DiffM2 xM_lt_N) qed subsection \Lemmas about Replicate\ lemma replicate_mset_minus_replicate_mset_same[simp]: "replicate_mset m x - replicate_mset n x = replicate_mset (m - n) x" by (induct m arbitrary: n, simp, metis left_diff_repeat_mset_distrib' repeat_mset_replicate_mset) lemma replicate_mset_subset_iff_lt[simp]: "replicate_mset m x \# replicate_mset n x \ m < n" by (induct n m rule: diff_induct) (auto intro: subset_mset.gr_zeroI) lemma replicate_mset_subseteq_iff_le[simp]: "replicate_mset m x \# replicate_mset n x \ m \ n" by (induct n m rule: diff_induct) auto lemma replicate_mset_lt_iff_lt[simp]: "replicate_mset m x < replicate_mset n x \ m < n" by (induct n m rule: diff_induct) (auto intro: subset_mset.gr_zeroI gr_zeroI) lemma replicate_mset_le_iff_le[simp]: "replicate_mset m x \ replicate_mset n x \ m \ n" by (induct n m rule: diff_induct) auto lemma replicate_mset_eq_iff[simp]: "replicate_mset m x = replicate_mset n y \ m = n \ (m \ 0 \ x = y)" by (cases m; cases n; simp) (metis in_replicate_mset insert_noteq_member size_replicate_mset union_single_eq_diff) lemma replicate_mset_plus: "replicate_mset (a + b) C = replicate_mset a C + replicate_mset b C" by (induct a) (auto simp: ac_simps) lemma mset_replicate_replicate_mset: "mset (replicate n L) = replicate_mset n L" by (induction n) auto lemma set_mset_single_iff_replicate_mset: "set_mset U = {a} \ (\n > 0. U = replicate_mset n a)" by (rule, metis count_greater_zero_iff count_replicate_mset insertI1 multi_count_eq singletonD zero_less_iff_neq_zero, force) lemma ex_replicate_mset_if_all_elems_eq: assumes "\x \# M. x = y" shows "\n. M = replicate_mset n y" using assms by (metis count_replicate_mset mem_Collect_eq multiset_eqI neq0_conv set_mset_def) subsection \Multiset and Set Conversions\ lemma count_mset_set_if: "count (mset_set A) a = (if a \ A \ finite A then 1 else 0)" by auto lemma mset_set_set_mset_empty_mempty[iff]: "mset_set (set_mset D) = {#} \ D = {#}" by (simp add: mset_set_empty_iff) lemma count_mset_set_le_one: "count (mset_set A) x \ 1" by (simp add: count_mset_set_if) lemma mset_set_set_mset_subseteq[simp]: "mset_set (set_mset A) \# A" by (simp add: mset_set_set_mset_msubset) lemma mset_sorted_list_of_set[simp]: "mset (sorted_list_of_set A) = mset_set A" by (metis mset_sorted_list_of_multiset sorted_list_of_mset_set) lemma sorted_sorted_list_of_multiset[simp]: "sorted (sorted_list_of_multiset (M :: 'a::linorder multiset))" by (metis mset_sorted_list_of_multiset sorted_list_of_multiset_mset sorted_sort) lemma mset_take_subseteq: "mset (take n xs) \# mset xs" apply (induct xs arbitrary: n) apply simp by (case_tac n) simp_all lemma sorted_list_of_multiset_eq_Nil[simp]: "sorted_list_of_multiset M = [] \ M = {#}" by (metis mset_sorted_list_of_multiset sorted_list_of_multiset_empty) subsection \Duplicate Removal\ (* TODO: use abbreviation? *) definition remdups_mset :: "'v multiset \ 'v multiset" where "remdups_mset S = mset_set (set_mset S)" lemma set_mset_remdups_mset[simp]: \set_mset (remdups_mset A) = set_mset A\ unfolding remdups_mset_def by auto lemma count_remdups_mset_eq_1: "a \# remdups_mset A \ count (remdups_mset A) a = 1" unfolding remdups_mset_def by (auto simp: count_eq_zero_iff intro: count_inI) lemma remdups_mset_empty[simp]: "remdups_mset {#} = {#}" unfolding remdups_mset_def by auto lemma remdups_mset_singleton[simp]: "remdups_mset {#a#} = {#a#}" unfolding remdups_mset_def by auto lemma remdups_mset_eq_empty[iff]: "remdups_mset D = {#} \ D = {#}" unfolding remdups_mset_def by blast lemma remdups_mset_singleton_sum[simp]: "remdups_mset (add_mset a A) = (if a \# A then remdups_mset A else add_mset a (remdups_mset A))" unfolding remdups_mset_def by (simp_all add: insert_absorb) lemma mset_remdups_remdups_mset[simp]: "mset (remdups D) = remdups_mset (mset D)" by (induction D) (auto simp add: ac_simps) declare mset_remdups_remdups_mset[symmetric, code] definition distinct_mset :: "'a multiset \ bool" where "distinct_mset S \ (\a. a \# S \ count S a = 1)" lemma distinct_mset_count_less_1: "distinct_mset S \ (\a. count S a \ 1)" using eq_iff nat_le_linear unfolding distinct_mset_def by fastforce lemma distinct_mset_empty[simp]: "distinct_mset {#}" unfolding distinct_mset_def by auto lemma distinct_mset_singleton: "distinct_mset {#a#}" unfolding distinct_mset_def by auto lemma distinct_mset_union: assumes dist: "distinct_mset (A + B)" shows "distinct_mset A" unfolding distinct_mset_count_less_1 proof (rule allI) fix a have \count A a \ count (A + B) a\ by auto moreover have \count (A + B) a \ 1\ using dist unfolding distinct_mset_count_less_1 by auto ultimately show \count A a \ 1\ by simp qed lemma distinct_mset_minus[simp]: "distinct_mset A \ distinct_mset (A - B)" by (metis diff_subset_eq_self mset_subset_eq_exists_conv distinct_mset_union) lemma count_remdups_mset_If: \count (remdups_mset A) a = (if a \# A then 1 else 0)\ unfolding remdups_mset_def by auto lemma distinct_mset_rempdups_union_mset: assumes "distinct_mset A" and "distinct_mset B" shows "A \# B = remdups_mset (A + B)" using assms nat_le_linear unfolding remdups_mset_def by (force simp add: multiset_eq_iff max_def count_mset_set_if distinct_mset_def not_in_iff) lemma distinct_mset_add_mset[simp]: "distinct_mset (add_mset a L) \ a \# L \ distinct_mset L" unfolding distinct_mset_def apply (rule iffI) apply (auto split: if_split_asm; fail)[] by (auto simp: not_in_iff; fail) lemma distinct_mset_size_eq_card: "distinct_mset C \ size C = card (set_mset C)" by (induction C) auto lemma distinct_mset_add: "distinct_mset (L + L') \ distinct_mset L \ distinct_mset L' \ L \# L' = {#}" by (induction L arbitrary: L') auto lemma distinct_mset_set_mset_ident[simp]: "distinct_mset M \ mset_set (set_mset M) = M" by (induction M) auto lemma distinct_finite_set_mset_subseteq_iff[iff]: assumes "distinct_mset M" "finite N" shows "set_mset M \ N \ M \# mset_set N" by (metis assms distinct_mset_set_mset_ident finite_set_mset msubset_mset_set_iff) lemma distinct_mem_diff_mset: assumes dist: "distinct_mset M" and mem: "x \ set_mset (M - N)" shows "x \ set_mset N" proof - have "count M x = 1" using dist mem by (meson distinct_mset_def in_diffD) then show ?thesis using mem by (metis count_greater_eq_one_iff in_diff_count not_less) qed lemma distinct_set_mset_eq: assumes "distinct_mset M" "distinct_mset N" "set_mset M = set_mset N" shows "M = N" using assms distinct_mset_set_mset_ident by fastforce lemma distinct_mset_union_mset[simp]: \distinct_mset (D \# C) \ distinct_mset D \ distinct_mset C\ unfolding distinct_mset_count_less_1 by force lemma distinct_mset_inter_mset: "distinct_mset C \ distinct_mset (C \# D)" "distinct_mset D \ distinct_mset (C \# D)" by (simp_all add: multiset_inter_def, metis distinct_mset_minus multiset_inter_commute multiset_inter_def) lemma distinct_mset_remove1_All: "distinct_mset C \ remove1_mset L C = removeAll_mset L C" by (auto simp: multiset_eq_iff distinct_mset_count_less_1) lemma distinct_mset_size_2: "distinct_mset {#a, b#} \ a \ b" by auto lemma distinct_mset_filter: "distinct_mset M \ distinct_mset {# L \# M. P L#}" by (simp add: distinct_mset_def) lemma distinct_mset_mset_distinct[simp]: \distinct_mset (mset xs) = distinct xs\ by (induction xs) auto lemma distinct_image_mset_inj: \inj_on f (set_mset M) \ distinct_mset (image_mset f M) \ distinct_mset M\ by (induction M) (auto simp: inj_on_def) subsection \Repeat Operation\ lemma repeat_mset_compower: "repeat_mset n A = (((+) A) ^^ n) {#}" by (induction n) auto lemma repeat_mset_prod: "repeat_mset (m * n) A = (((+) (repeat_mset n A)) ^^ m) {#}" by (induction m) (auto simp: repeat_mset_distrib) subsection \Cartesian Product\ text \Definition of the cartesian products over multisets. The construction mimics of the cartesian product on sets and use the same theorem names (adding only the suffix \_mset\ to Sigma and Times). See file @{file \~~/src/HOL/Product_Type.thy\}\ definition Sigma_mset :: "'a multiset \ ('a \ 'b multiset) \ ('a \ 'b) multiset" where "Sigma_mset A B \ \# {#{#(a, b). b \# B a#}. a \# A #}" abbreviation Times_mset :: "'a multiset \ 'b multiset \ ('a \ 'b) multiset" (infixr "\#" 80) where "Times_mset A B \ Sigma_mset A (\_. B)" hide_const (open) Times_mset text \Contrary to the set version @{term \SIGMA x:A. B\}, we use the non-ASCII symbol \\#\.\ syntax "_Sigma_mset" :: "[pttrn, 'a multiset, 'b multiset] => ('a * 'b) multiset" ("(3SIGMAMSET _\#_./ _)" [0, 0, 10] 10) translations "SIGMAMSET x\#A. B" == "CONST Sigma_mset A (\x. B)" text \Link between the multiset and the set cartesian product:\ lemma Times_mset_Times: "set_mset (A \# B) = set_mset A \ set_mset B" unfolding Sigma_mset_def by auto lemma Sigma_msetI [intro!]: "\a \# A; b \# B a\ \ (a, b) \# Sigma_mset A B" by (unfold Sigma_mset_def) auto lemma Sigma_msetE[elim!]: "\c \# Sigma_mset A B; \x y. \x \# A; y \# B x; c = (x, y)\ \ P\ \ P" by (unfold Sigma_mset_def) auto text \Elimination of @{term "(a, b) \# A \# B"} -- introduces no eigenvariables.\ lemma Sigma_msetD1: "(a, b) \# Sigma_mset A B \ a \# A" by blast lemma Sigma_msetD2: "(a, b) \# Sigma_mset A B \ b \# B a" by blast lemma Sigma_msetE2: "\(a, b) \# Sigma_mset A B; \a \# A; b \# B a\ \ P\ \ P" by blast lemma Sigma_mset_cong: "\A = B; \x. x \# B \ C x = D x\ \ (SIGMAMSET x \# A. C x) = (SIGMAMSET x \# B. D x)" by (metis (mono_tags, lifting) Sigma_mset_def image_mset_cong) lemma count_sum_mset: "count (\# M) b = (\P \# M. count P b)" by (induction M) auto lemma Sigma_mset_plus_distrib1[simp]: "Sigma_mset (A + B) C = Sigma_mset A C + Sigma_mset B C" unfolding Sigma_mset_def by auto lemma Sigma_mset_plus_distrib2[simp]: "Sigma_mset A (\i. B i + C i) = Sigma_mset A B + Sigma_mset A C" unfolding Sigma_mset_def by (induction A) (auto simp: multiset_eq_iff) lemma Times_mset_single_left: "{#a#} \# B = image_mset (Pair a) B" unfolding Sigma_mset_def by auto lemma Times_mset_single_right: "A \# {#b#} = image_mset (\a. Pair a b) A" unfolding Sigma_mset_def by (induction A) auto lemma Times_mset_single_single[simp]: "{#a#} \# {#b#} = {#(a, b)#}" unfolding Sigma_mset_def by simp lemma count_image_mset_Pair: "count (image_mset (Pair a) B) (x, b) = (if x = a then count B b else 0)" by (induction B) auto lemma count_Sigma_mset: "count (Sigma_mset A B) (a, b) = count A a * count (B a) b" by (induction A) (auto simp: Sigma_mset_def count_image_mset_Pair) lemma Sigma_mset_empty1[simp]: "Sigma_mset {#} B = {#}" unfolding Sigma_mset_def by auto lemma Sigma_mset_empty2[simp]: "A \# {#} = {#}" by (auto simp: multiset_eq_iff count_Sigma_mset) lemma Sigma_mset_mono: assumes "A \# C" and "\x. x \# A \ B x \# D x" shows "Sigma_mset A B \# Sigma_mset C D" proof - have "count A a * count (B a) b \ count C a * count (D a) b" for a b using assms unfolding subseteq_mset_def by (metis count_inI eq_iff mult_eq_0_iff mult_le_mono) then show ?thesis by (auto simp: subseteq_mset_def count_Sigma_mset) qed lemma mem_Sigma_mset_iff[iff]: "((a,b) \# Sigma_mset A B) = (a \# A \ b \# B a)" by blast lemma mem_Times_mset_iff: "x \# A \# B \ fst x \# A \ snd x \# B" by (induct x) simp lemma Sigma_mset_empty_iff: "(SIGMAMSET i\#I. X i) = {#} \ (\i\#I. X i = {#})" by (auto simp: Sigma_mset_def) lemma Times_mset_subset_mset_cancel1: "x \# A \ (A \# B \# A \# C) = (B \# C)" by (auto simp: subseteq_mset_def count_Sigma_mset) lemma Times_mset_subset_mset_cancel2: "x \# C \ (A \# C \# B \# C) = (A \# B)" by (auto simp: subseteq_mset_def count_Sigma_mset) lemma Times_mset_eq_cancel2: "x \# C \ (A \# C = B \# C) = (A = B)" by (auto simp: multiset_eq_iff count_Sigma_mset dest!: in_countE) lemma split_paired_Ball_mset_Sigma_mset[simp]: "(\z\#Sigma_mset A B. P z) \ (\x\#A. \y\#B x. P (x, y))" by blast lemma split_paired_Bex_mset_Sigma_mset[simp]: "(\z\#Sigma_mset A B. P z) \ (\x\#A. \y\#B x. P (x, y))" by blast lemma sum_mset_if_eq_constant: "(\x\#M. if a = x then (f x) else 0) = (((+) (f a)) ^^ (count M a)) 0" by (induction M) (auto simp: ac_simps) lemma iterate_op_plus: "(((+) k) ^^ m) 0 = k * m" by (induction m) auto lemma untion_image_mset_Pair_distribute: "\#{#image_mset (Pair x) (C x). x \# J - I#} = \# {#image_mset (Pair x) (C x). x \# J#} - \#{#image_mset (Pair x) (C x). x \# I#}" by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant iterate_op_plus diff_mult_distrib2) lemma Sigma_mset_Un_distrib1: "Sigma_mset (I \# J) C = Sigma_mset I C \# Sigma_mset J C" by (auto simp: Sigma_mset_def sup_subset_mset_def untion_image_mset_Pair_distribute) lemma Sigma_mset_Un_distrib2: "(SIGMAMSET i\#I. A i \# B i) = Sigma_mset I A \# Sigma_mset I B" by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant Sigma_mset_def diff_mult_distrib2 iterate_op_plus max_def not_in_iff) lemma Sigma_mset_Int_distrib1: "Sigma_mset (I \# J) C = Sigma_mset I C \# Sigma_mset J C" by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant Sigma_mset_def iterate_op_plus min_def not_in_iff) lemma Sigma_mset_Int_distrib2: "(SIGMAMSET i\#I. A i \# B i) = Sigma_mset I A \# Sigma_mset I B" by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant Sigma_mset_def iterate_op_plus min_def not_in_iff) lemma Sigma_mset_Diff_distrib1: "Sigma_mset (I - J) C = Sigma_mset I C - Sigma_mset J C" by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant Sigma_mset_def iterate_op_plus min_def not_in_iff diff_mult_distrib2) lemma Sigma_mset_Diff_distrib2: "(SIGMAMSET i\#I. A i - B i) = Sigma_mset I A - Sigma_mset I B" by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant Sigma_mset_def iterate_op_plus min_def not_in_iff diff_mult_distrib) lemma Sigma_mset_Union: "Sigma_mset (\#X) B = (\# (image_mset (\A. Sigma_mset A B) X))" by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant Sigma_mset_def iterate_op_plus min_def not_in_iff sum_mset_distrib_left) lemma Times_mset_Un_distrib1: "(A \# B) \# C = A \# C \# B \# C" by (fact Sigma_mset_Un_distrib1) lemma Times_mset_Int_distrib1: "(A \# B) \# C = A \# C \# B \# C" by (fact Sigma_mset_Int_distrib1) lemma Times_mset_Diff_distrib1: "(A - B) \# C = A \# C - B \# C" by (fact Sigma_mset_Diff_distrib1) lemma Times_mset_empty[simp]: "A \# B = {#} \ A = {#} \ B = {#}" by (auto simp: Sigma_mset_empty_iff) lemma Times_insert_left: "A \# add_mset x B = A \# B + image_mset (\a. Pair a x) A" unfolding add_mset_add_single[of x B] Sigma_mset_plus_distrib2 by (simp add: Times_mset_single_right) lemma Times_insert_right: "add_mset a A \# B = A \# B + image_mset (Pair a) B" unfolding add_mset_add_single[of a A] Sigma_mset_plus_distrib1 by (simp add: Times_mset_single_left) lemma fst_image_mset_times_mset [simp]: "image_mset fst (A \# B) = (if B = {#} then {#} else repeat_mset (size B) A)" by (induct B) (auto simp: Times_mset_single_right ac_simps Times_insert_left) lemma snd_image_mset_times_mset [simp]: "image_mset snd (A \# B) = (if A = {#} then {#} else repeat_mset (size A) B)" by (induct B) (auto simp add: Times_mset_single_right Times_insert_left image_mset_const_eq) lemma product_swap_mset: "image_mset prod.swap (A \# B) = B \# A" by (induction A) (auto simp add: Times_mset_single_left Times_mset_single_right Times_insert_right Times_insert_left) context begin qualified definition product_mset :: "'a multiset \ 'b multiset \ ('a \ 'b) multiset" where [code_abbrev]: "product_mset A B = A \# B" lemma member_product_mset: "x \# product_mset A B \ x \# A \# B" by (simp add: Multiset_More.product_mset_def) end lemma count_Sigma_mset_abs_def: "count (Sigma_mset A B) = (\(a, b) \ count A a * count (B a) b)" by (auto simp: fun_eq_iff count_Sigma_mset) lemma Times_mset_image_mset1: "image_mset f A \# B = image_mset (\(a, b). (f a, b)) (A \# B)" by (induct B) (auto simp: Times_insert_left) lemma Times_mset_image_mset2: "A \# image_mset f B = image_mset (\(a, b). (a, f b)) (A \# B)" by (induct A) (auto simp: Times_insert_right) lemma sum_le_singleton: "A \ {x} \ sum f A = (if x \ A then f x else 0)" by (auto simp: subset_singleton_iff elim: finite_subset) lemma Times_mset_assoc: "(A \# B) \# C = image_mset (\(a, b, c). ((a, b), c)) (A \# B \# C)" by (auto simp: multiset_eq_iff count_Sigma_mset count_image_mset vimage_def Times_mset_Times Int_commute count_eq_zero_iff intro!: trans[OF _ sym[OF sum_le_singleton[of _ "(_, _, _)"]]] cong: sum.cong if_cong) subsection \Transfer Rules\ lemma plus_multiset_transfer[transfer_rule]: "(rel_fun (rel_mset R) (rel_fun (rel_mset R) (rel_mset R))) (+) (+)" by (unfold rel_fun_def rel_mset_def) (force dest: list_all2_appendI intro: exI[of _ "_ @ _"] conjI[rotated]) lemma minus_multiset_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique R" shows "(rel_fun (rel_mset R) (rel_fun (rel_mset R) (rel_mset R))) (-) (-)" proof (unfold rel_fun_def rel_mset_def, safe) fix xs ys xs' ys' assume [transfer_rule]: "list_all2 R xs ys" "list_all2 R xs' ys'" have "list_all2 R (fold remove1 xs' xs) (fold remove1 ys' ys)" by transfer_prover moreover have "mset (fold remove1 xs' xs) = mset xs - mset xs'" by (induct xs' arbitrary: xs) auto moreover have "mset (fold remove1 ys' ys) = mset ys - mset ys'" by (induct ys' arbitrary: ys) auto ultimately show "\xs'' ys''. mset xs'' = mset xs - mset xs' \ mset ys'' = mset ys - mset ys' \ list_all2 R xs'' ys''" by blast qed declare rel_mset_Zero[transfer_rule] lemma count_transfer[transfer_rule]: assumes "bi_unique R" shows "(rel_fun (rel_mset R) (rel_fun R (=))) count count" unfolding rel_fun_def rel_mset_def proof safe fix x y xs ys assume "list_all2 R xs ys" "R x y" then show "count (mset xs) x = count (mset ys) y" proof (induct xs ys rule: list.rel_induct) case (Cons x' xs y' ys) then show ?case using assms unfolding bi_unique_alt_def2 by (auto simp: rel_fun_def) qed simp qed lemma subseteq_multiset_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique R" "right_total R" shows "(rel_fun (rel_mset R) (rel_fun (rel_mset R) (=))) (\M N. filter_mset (Domainp R) M \# filter_mset (Domainp R) N) (\#)" proof - have count_filter_mset_less: "(\a. count (filter_mset (Domainp R) M) a \ count (filter_mset (Domainp R) N) a) \ (\a \ {x. Domainp R x}. count M a \ count N a)" for M and N by auto show ?thesis unfolding subseteq_mset_def count_filter_mset_less by transfer_prover qed lemma sum_mset_transfer[transfer_rule]: "R 0 0 \ rel_fun R (rel_fun R R) (+) (+) \ (rel_fun (rel_mset R) R) sum_mset sum_mset" using sum_list_transfer[of R] unfolding rel_fun_def rel_mset_def by auto lemma Sigma_mset_transfer[transfer_rule]: "(rel_fun (rel_mset R) (rel_fun (rel_fun R (rel_mset S)) (rel_mset (rel_prod R S)))) Sigma_mset Sigma_mset" by (unfold Sigma_mset_def) transfer_prover subsection \Even More about Multisets\ -subsubsection \Multisets and functions\ +subsubsection \Multisets and Functions\ lemma range_image_mset: assumes "set_mset Ds \ range f" shows "Ds \ range (image_mset f)" proof - have "\D. D \# Ds \ (\C. f C = D)" using assms by blast then obtain f_i where f_p: "\D. D \# Ds \ (f (f_i D) = D)" by metis define Cs where "Cs \ image_mset f_i Ds" from f_p Cs_def have "image_mset f Cs = Ds" by auto then show ?thesis by blast qed -subsubsection \Multisets and lists\ +subsubsection \Multisets and Lists\ + +lemma length_sorted_list_of_multiset[simp]: "length (sorted_list_of_multiset A) = size A" + by (metis mset_sorted_list_of_multiset size_mset) definition list_of_mset :: "'a multiset \ 'a list" where "list_of_mset m = (SOME l. m = mset l)" lemma list_of_mset_exi: "\l. m = mset l" using ex_mset by metis -lemma [simp]: "mset (list_of_mset m) = m" +lemma mset_list_of_mset [simp]: "mset (list_of_mset m) = m" by (metis (mono_tags, lifting) ex_mset list_of_mset_def someI_ex) +lemma length_list_of_mset[simp]: "length (list_of_mset A) = size A" + unfolding list_of_mset_def by (metis (mono_tags) ex_mset size_mset someI_ex) + lemma range_mset_map: assumes "set_mset Ds \ range f" shows "Ds \ range (\Cl. mset (map f Cl))" proof - have "Ds \ range (image_mset f)" by (simp add: assms range_image_mset) then obtain Cs where Cs_p: "image_mset f Cs = Ds" by auto define Cl where "Cl = list_of_mset Cs" then have "mset Cl = Cs" by auto then have "image_mset f (mset Cl) = Ds" using Cs_p by auto then have "mset (map f Cl) = Ds" by auto then show ?thesis by auto qed lemma list_of_mset_empty[iff]: "list_of_mset m = [] \ m = {#}" by (metis (mono_tags, lifting) ex_mset list_of_mset_def mset_zero_iff_right someI_ex) lemma in_mset_conv_nth: "(x \# mset xs) = (\i# LL" assumes "LL \ set Ci" shows "L \# sum_list Ci" using assms by (induction Ci) auto lemma in_mset_sum_list2: assumes "L \# sum_list Ci" obtains LL where "LL \ set Ci" "L \# LL" using assms by (induction Ci) auto lemma subseteq_list_Union_mset: assumes "length Ci = n" assumes "length CAi = n" assumes "\i# CAi ! i " shows "\# (mset Ci) \# \# (mset CAi)" using assms proof (induction n arbitrary: Ci CAi) case 0 then show ?case by auto next case (Suc n) from Suc have "\i# tl CAi ! i" by (simp add: nth_tl) hence "\#(mset (tl Ci)) \# \#(mset (tl CAi))" using Suc by auto moreover have "hd Ci \# hd CAi" using Suc by (metis hd_conv_nth length_greater_0_conv zero_less_Suc) ultimately show "\#(mset Ci) \# \#(mset CAi)" using Suc by (cases Ci; cases CAi) (auto intro: subset_mset.add_mono) qed -subsubsection \More on multisets and functions\ +subsubsection \More on Multisets and Functions\ lemma subseteq_mset_size_eql: "X \# Y \ size Y = size X \ X = Y" using mset_subset_size subset_mset_def by fastforce lemma image_mset_of_subset_list: assumes "image_mset \ C' = mset lC" shows "\qC'. map \ qC' = lC \ mset qC' = C'" using assms apply (induction lC arbitrary: C') subgoal by simp subgoal by (fastforce dest!: msed_map_invR intro: exI[of _ \_ # _\]) done lemma image_mset_of_subset: assumes "A \# image_mset \ C'" shows "\A'. image_mset \ A' = A \ A' \# C'" proof - define C where "C = image_mset \ C'" define lA where "lA = list_of_mset A" define lD where "lD = list_of_mset (C-A)" define lC where "lC = lA @ lD" have "mset lC = C" using C_def assms unfolding lD_def lC_def lA_def by auto then have "\qC'. map \ qC' = lC \ mset qC' = C'" using assms image_mset_of_subset_list unfolding C_def by metis then obtain qC' where qC'_p: "map \ qC' = lC \ mset qC' = C'" by auto let ?lA' = "take (length lA) qC'" have m: "map \ ?lA' = lA" using qC'_p lC_def by (metis append_eq_conv_conj take_map) let ?A' = "mset ?lA'" have "image_mset \ ?A' = A" using m using lA_def by (metis (full_types) ex_mset list_of_mset_def mset_map someI_ex) moreover have "?A' \# C'" using qC'_p unfolding lA_def using mset_take_subseteq by blast ultimately show ?thesis by blast qed lemma all_the_same: "\x \# X. x = y \ card (set_mset X) \ Suc 0" by (metis card.empty card.insert card_mono finite.intros(1) finite_insert le_SucI singletonI subsetI) lemma Melem_subseteq_Union_mset[simp]: assumes "x \# T" shows "x \# \#T" using assms sum_mset.remove by force lemma Melem_subset_eq_sum_list[simp]: assumes "x \# mset T" shows "x \# sum_list T" using assms by (metis mset_subset_eq_add_left sum_mset.remove sum_mset_sum_list) lemma less_subset_eq_Union_mset[simp]: assumes "i < length CAi" shows "CAi ! i \# \#(mset CAi)" proof - from assms have "CAi ! i \# mset CAi" by auto then show ?thesis by auto qed lemma less_subset_eq_sum_list[simp]: assumes "i < length CAi" shows "CAi ! i \# sum_list CAi" proof - from assms have "CAi ! i \# mset CAi" by auto then show ?thesis by auto qed -subsubsection \More on multiset order\ + +subsubsection \More on Multiset Order\ 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/Ordered_Resolution_Prover/Abstract_Substitution.thy b/thys/Ordered_Resolution_Prover/Abstract_Substitution.thy --- a/thys/Ordered_Resolution_Prover/Abstract_Substitution.thy +++ b/thys/Ordered_Resolution_Prover/Abstract_Substitution.thy @@ -1,1154 +1,1152 @@ (* Title: Abstract Substitutions Author: Dmitriy Traytel , 2014 Author: Jasmin Blanchette , 2014, 2017 Author: Anders Schlichtkrull , 2016, 2017 Maintainer: Anders Schlichtkrull *) section \Abstract Substitutions\ theory Abstract_Substitution imports Clausal_Logic Map2 begin text \ Atoms and substitutions are abstracted away behind some locales, to avoid having a direct dependency on the IsaFoR library. Conventions: \'s\ substitutions, \'a\ atoms. \ subsection \Library\ -(* TODO: move to Isabelle? *) lemma f_Suc_decr_eventually_const: fixes f :: "nat \ nat" assumes leq: "\i. f (Suc i) \ f i" shows "\l. \l' \ l. f l' = f (Suc l')" proof (rule ccontr) assume a: "\l. \l' \ l. f l' = f (Suc l')" have "\i. \i'. i' > i \ f i' < f i" proof fix i from a have "\l' \ i. f l' \ f (Suc l')" by auto then obtain l' where l'_p: "l' \ i \ f l' \ f (Suc l')" by metis then have "f l' > f (Suc l')" using leq le_eq_less_or_eq by auto moreover have "f i \ f l'" using leq l'_p by (induction l' arbitrary: i) (blast intro: lift_Suc_antimono_le)+ ultimately show "\i' > i. f i' < f i" using l'_p less_le_trans by blast qed then obtain g_sm :: "nat \ nat" where g_sm_p: "\i. g_sm i > i \ f (g_sm i) < f i" by metis define c :: "nat \ nat" where "\n. c n = (g_sm ^^ n) 0" have "f (c i) > f (c (Suc i))" for i by (induction i) (auto simp: c_def g_sm_p) then have "\i. (f \ c) i > (f \ c) (Suc i)" by auto then have "\fc :: nat \ nat. \i. fc i > fc (Suc i)" by metis then show False using wf_less_than by (simp add: wf_iff_no_infinite_down_chain) qed subsection \Substitution Operators\ locale substitution_ops = fixes subst_atm :: "'a \ 's \ 'a" and id_subst :: 's and comp_subst :: "'s \ 's \ 's" begin abbreviation subst_atm_abbrev :: "'a \ 's \ 'a" (infixl "\a" 67) where "subst_atm_abbrev \ subst_atm" abbreviation comp_subst_abbrev :: "'s \ 's \ 's" (infixl "\" 67) where "comp_subst_abbrev \ comp_subst" definition comp_substs :: "'s list \ 's list \ 's list" (infixl "\s" 67) where "\s \s \s = map2 comp_subst \s \s" definition subst_atms :: "'a set \ 's \ 'a set" (infixl "\as" 67) where "AA \as \ = (\A. A \a \) ` AA" definition subst_atmss :: "'a set set \ 's \ 'a set set" (infixl "\ass" 67) where "AAA \ass \ = (\AA. AA \as \) ` AAA" definition subst_atm_list :: "'a list \ 's \ 'a list" (infixl "\al" 67) where "As \al \ = map (\A. A \a \) As" definition subst_atm_mset :: "'a multiset \ 's \ 'a multiset" (infixl "\am" 67) where "AA \am \ = image_mset (\A. A \a \) AA" definition subst_atm_mset_list :: "'a multiset list \ 's \ 'a multiset list" (infixl "\aml" 67) where "AAA \aml \ = map (\AA. AA \am \) AAA" definition subst_atm_mset_lists :: "'a multiset list \ 's list \ 'a multiset list" (infixl "\\aml" 67) where "AAs \\aml \s = map2 (\am) AAs \s" definition subst_lit :: "'a literal \ 's \ 'a literal" (infixl "\l" 67) where "L \l \ = map_literal (\A. A \a \) L" lemma atm_of_subst_lit[simp]: "atm_of (L \l \) = atm_of L \a \" unfolding subst_lit_def by (cases L) simp+ definition subst_cls :: "'a clause \ 's \ 'a clause" (infixl "\" 67) where "AA \ \ = image_mset (\A. A \l \) AA" definition subst_clss :: "'a clause set \ 's \ 'a clause set" (infixl "\cs" 67) where "AA \cs \ = (\A. A \ \) ` AA" definition subst_cls_list :: "'a clause list \ 's \ 'a clause list" (infixl "\cl" 67) where "Cs \cl \ = map (\A. A \ \) Cs" definition subst_cls_lists :: "'a clause list \ 's list \ 'a clause list" (infixl "\\cl" 67) where "Cs \\cl \s = map2 (\) Cs \s" definition subst_cls_mset :: "'a clause multiset \ 's \ 'a clause multiset" (infixl "\cm" 67) where "CC \cm \ = image_mset (\A. A \ \) CC" lemma subst_cls_add_mset[simp]: "add_mset L C \ \ = add_mset (L \l \) (C \ \)" unfolding subst_cls_def by simp lemma subst_cls_mset_add_mset[simp]: "add_mset C CC \cm \ = add_mset (C \ \) (CC \cm \)" unfolding subst_cls_mset_def by simp definition generalizes_atm :: "'a \ 'a \ bool" where "generalizes_atm A B \ (\\. A \a \ = B)" definition strictly_generalizes_atm :: "'a \ 'a \ bool" where "strictly_generalizes_atm A B \ generalizes_atm A B \ \ generalizes_atm B A" definition generalizes_lit :: "'a literal \ 'a literal \ bool" where "generalizes_lit L M \ (\\. L \l \ = M)" definition strictly_generalizes_lit :: "'a literal \ 'a literal \ bool" where "strictly_generalizes_lit L M \ generalizes_lit L M \ \ generalizes_lit M L" definition generalizes_cls :: "'a clause \ 'a clause \ bool" where "generalizes_cls C D \ (\\. C \ \ = D)" definition strictly_generalizes_cls :: "'a clause \ 'a clause \ bool" where "strictly_generalizes_cls C D \ generalizes_cls C D \ \ generalizes_cls D C" definition subsumes :: "'a clause \ 'a clause \ bool" where "subsumes C D \ (\\. C \ \ \# D)" definition strictly_subsumes :: "'a clause \ 'a clause \ bool" where "strictly_subsumes C D \ subsumes C D \ \ subsumes D C" -(* FIXME: define as exists renaming from one to the other? *) definition variants :: "'a clause \ 'a clause \ bool" where "variants C D \ generalizes_cls C D \ generalizes_cls D C" definition is_renaming :: "'s \ bool" where "is_renaming \ \ (\\. \ \ \ = id_subst)" definition is_renaming_list :: "'s list \ bool" where "is_renaming_list \s \ (\\ \ set \s. is_renaming \)" definition inv_renaming :: "'s \ 's" where "inv_renaming \ = (SOME \. \ \ \ = id_subst)" definition is_ground_atm :: "'a \ bool" where "is_ground_atm A \ (\\. A = A \a \)" definition is_ground_atms :: "'a set \ bool" where "is_ground_atms AA = (\A \ AA. is_ground_atm A)" definition is_ground_atm_list :: "'a list \ bool" where "is_ground_atm_list As \ (\A \ set As. is_ground_atm A)" definition is_ground_atm_mset :: "'a multiset \ bool" where "is_ground_atm_mset AA \ (\A. A \# AA \ is_ground_atm A)" definition is_ground_lit :: "'a literal \ bool" where "is_ground_lit L \ is_ground_atm (atm_of L)" definition is_ground_cls :: "'a clause \ bool" where "is_ground_cls C \ (\L. L \# C \ is_ground_lit L)" definition is_ground_clss :: "'a clause set \ bool" where "is_ground_clss CC \ (\C \ CC. is_ground_cls C)" definition is_ground_cls_list :: "'a clause list \ bool" where "is_ground_cls_list CC \ (\C \ set CC. is_ground_cls C)" definition is_ground_subst :: "'s \ bool" where "is_ground_subst \ \ (\A. is_ground_atm (A \a \))" definition is_ground_subst_list :: "'s list \ bool" where "is_ground_subst_list \s \ (\\ \ set \s. is_ground_subst \)" definition grounding_of_cls :: "'a clause \ 'a clause set" where "grounding_of_cls C = {C \ \ | \. is_ground_subst \}" definition grounding_of_clss :: "'a clause set \ 'a clause set" where "grounding_of_clss CC = (\C \ CC. grounding_of_cls C)" definition is_unifier :: "'s \ 'a set \ bool" where "is_unifier \ AA \ card (AA \as \) \ 1" definition is_unifiers :: "'s \ 'a set set \ bool" where "is_unifiers \ AAA \ (\AA \ AAA. is_unifier \ AA)" definition is_mgu :: "'s \ 'a set set \ bool" where "is_mgu \ AAA \ is_unifiers \ AAA \ (\\. is_unifiers \ AAA \ (\\. \ = \ \ \))" definition var_disjoint :: "'a clause list \ bool" where "var_disjoint Cs \ (\\s. length \s = length Cs \ (\\. \i < length Cs. \S. S \# Cs ! i \ S \ \s ! i = S \ \))" end subsection \Substitution Lemmas\ locale substitution = substitution_ops subst_atm id_subst comp_subst for subst_atm :: "'a \ 's \ 'a" and id_subst :: 's and comp_subst :: "'s \ 's \ 's" + fixes renamings_apart :: "'a clause list \ 's list" and atm_of_atms :: "'a list \ 'a" assumes subst_atm_id_subst[simp]: "A \a id_subst = A" and subst_atm_comp_subst[simp]: "A \a (\ \ \) = (A \a \) \a \" and subst_ext: "(\A. A \a \ = A \a \) \ \ = \" and make_ground_subst: "is_ground_cls (C \ \) \ \\. is_ground_subst \ \C \ \ = C \ \" and wf_strictly_generalizes_atm: "wfP strictly_generalizes_atm" and renamings_apart_length: "length (renamings_apart Cs) = length Cs" and renamings_apart_renaming: "\ \ set (renamings_apart Cs) \ is_renaming \" and renamings_apart_var_disjoint: "var_disjoint (Cs \\cl (renamings_apart Cs))" and atm_of_atms_subst: "\As Bs. atm_of_atms As \a \ = atm_of_atms Bs \ map (\A. A \a \) As = Bs" begin lemma subst_ext_iff: "\ = \ \ (\A. A \a \ = A \a \)" by (blast intro: subst_ext) subsubsection \Identity Substitution\ lemma id_subst_comp_subst[simp]: "id_subst \ \ = \" by (rule subst_ext) simp lemma comp_subst_id_subst[simp]: "\ \ id_subst = \" by (rule subst_ext) simp lemma id_subst_comp_substs[simp]: "replicate (length \s) id_subst \s \s = \s" using comp_substs_def by (induction \s) auto lemma comp_substs_id_subst[simp]: "\s \s replicate (length \s) id_subst = \s" using comp_substs_def by (induction \s) auto lemma subst_atms_id_subst[simp]: "AA \as id_subst = AA" unfolding subst_atms_def by simp lemma subst_atmss_id_subst[simp]: "AAA \ass id_subst = AAA" unfolding subst_atmss_def by simp lemma subst_atm_list_id_subst[simp]: "As \al id_subst = As" unfolding subst_atm_list_def by auto lemma subst_atm_mset_id_subst[simp]: "AA \am id_subst = AA" unfolding subst_atm_mset_def by simp lemma subst_atm_mset_list_id_subst[simp]: "AAs \aml id_subst = AAs" unfolding subst_atm_mset_list_def by simp lemma subst_atm_mset_lists_id_subst[simp]: "AAs \\aml replicate (length AAs) id_subst = AAs" unfolding subst_atm_mset_lists_def by (induct AAs) auto lemma subst_lit_id_subst[simp]: "L \l id_subst = L" unfolding subst_lit_def by (simp add: literal.map_ident) lemma subst_cls_id_subst[simp]: "C \ id_subst = C" unfolding subst_cls_def by simp lemma subst_clss_id_subst[simp]: "CC \cs id_subst = CC" unfolding subst_clss_def by simp lemma subst_cls_list_id_subst[simp]: "Cs \cl id_subst = Cs" unfolding subst_cls_list_def by simp lemma subst_cls_lists_id_subst[simp]: "Cs \\cl replicate (length Cs) id_subst = Cs" unfolding subst_cls_lists_def by (induct Cs) auto lemma subst_cls_mset_id_subst[simp]: "CC \cm id_subst = CC" unfolding subst_cls_mset_def by simp subsubsection \Associativity of Composition\ lemma comp_subst_assoc[simp]: "\ \ (\ \ \) = \ \ \ \ \" by (rule subst_ext) simp subsubsection \Compatibility of Substitution and Composition\ lemma subst_atms_comp_subst[simp]: "AA \as (\ \ \) = AA \as \ \as \" unfolding subst_atms_def by auto lemma subst_atmss_comp_subst[simp]: "AAA \ass (\ \ \) = AAA \ass \ \ass \" unfolding subst_atmss_def by auto lemma subst_atm_list_comp_subst[simp]: "As \al (\ \ \) = As \al \ \al \" unfolding subst_atm_list_def by auto lemma subst_atm_mset_comp_subst[simp]: "AA \am (\ \ \) = AA \am \ \am \" unfolding subst_atm_mset_def by auto lemma subst_atm_mset_list_comp_subst[simp]: "AAs \aml (\ \ \) = (AAs \aml \) \aml \" unfolding subst_atm_mset_list_def by auto lemma subst_atm_mset_lists_comp_substs[simp]: "AAs \\aml (\s \s \s) = AAs \\aml \s \\aml \s" unfolding subst_atm_mset_lists_def comp_substs_def map_zip_map map_zip_map2 map_zip_assoc by (simp add: split_def) lemma subst_lit_comp_subst[simp]: "L \l (\ \ \) = L \l \ \l \" unfolding subst_lit_def by (auto simp: literal.map_comp o_def) lemma subst_cls_comp_subst[simp]: "C \ (\ \ \) = C \ \ \ \" unfolding subst_cls_def by auto lemma subst_clsscomp_subst[simp]: "CC \cs (\ \ \) = CC \cs \ \cs \" unfolding subst_clss_def by auto lemma subst_cls_list_comp_subst[simp]: "Cs \cl (\ \ \) = Cs \cl \ \cl \" unfolding subst_cls_list_def by auto lemma subst_cls_lists_comp_substs[simp]: "Cs \\cl (\s \s \s) = Cs \\cl \s \\cl \s" unfolding subst_cls_lists_def comp_substs_def map_zip_map map_zip_map2 map_zip_assoc by (simp add: split_def) lemma subst_cls_mset_comp_subst[simp]: "CC \cm (\ \ \) = CC \cm \ \cm \" unfolding subst_cls_mset_def by auto subsubsection \``Commutativity'' of Membership and Substitution\ lemma Melem_subst_atm_mset[simp]: "A \# AA \am \ \ (\B. B \# AA \ A = B \a \)" unfolding subst_atm_mset_def by auto lemma Melem_subst_cls[simp]: "L \# C \ \ \ (\M. M \# C \ L = M \l \)" unfolding subst_cls_def by auto lemma Melem_subst_cls_mset[simp]: "AA \# CC \cm \ \ (\BB. BB \# CC \ AA = BB \ \)" unfolding subst_cls_mset_def by auto subsubsection \Signs and Substitutions\ lemma subst_lit_is_neg[simp]: "is_neg (L \l \) = is_neg L" unfolding subst_lit_def by auto lemma subst_lit_is_pos[simp]: "is_pos (L \l \) = is_pos L" unfolding subst_lit_def by auto lemma subst_minus[simp]: "(- L) \l \ = - (L \l \)" by (simp add: literal.map_sel subst_lit_def uminus_literal_def) subsubsection \Substitution on Literal(s)\ lemma eql_neg_lit_eql_atm[simp]: "(Neg A' \l \) = Neg A \ A' \a \ = A" by (simp add: subst_lit_def) lemma eql_pos_lit_eql_atm[simp]: "(Pos A' \l \) = Pos A \ A' \a \ = A" by (simp add: subst_lit_def) lemma subst_cls_negs[simp]: "(negs AA) \ \ = negs (AA \am \)" unfolding subst_cls_def subst_lit_def subst_atm_mset_def by auto lemma subst_cls_poss[simp]: "(poss AA) \ \ = poss (AA \am \)" unfolding subst_cls_def subst_lit_def subst_atm_mset_def by auto lemma atms_of_subst_atms: "atms_of C \as \ = atms_of (C \ \)" proof - have "atms_of (C \ \) = set_mset (image_mset atm_of (image_mset (map_literal (\A. A \a \)) C))" unfolding subst_cls_def subst_atms_def subst_lit_def atms_of_def by auto also have "... = set_mset (image_mset (\A. A \a \) (image_mset atm_of C))" by simp (meson literal.map_sel) finally show "atms_of C \as \ = atms_of (C \ \)" unfolding subst_atms_def atms_of_def by auto qed lemma in_image_Neg_is_neg[simp]: "L \l \ \ Neg ` AA \ is_neg L" by (metis bex_imageD literal.disc(2) literal.map_disc_iff subst_lit_def) lemma subst_lit_in_negs_subst_is_neg: "L \l \ \# (negs AA) \ \ \ is_neg L" by simp lemma subst_lit_in_negs_is_neg: "L \l \ \# negs AA \ is_neg L" by simp subsubsection \Substitution on Empty\ lemma subst_atms_empty[simp]: "{} \as \ = {}" unfolding subst_atms_def by auto lemma subst_atmss_empty[simp]: "{} \ass \ = {}" unfolding subst_atmss_def by auto lemma comp_substs_empty_iff[simp]: "\s \s \s = [] \ \s = [] \ \s = []" using comp_substs_def map2_empty_iff by auto lemma subst_atm_list_empty[simp]: "[] \al \ = []" unfolding subst_atm_list_def by auto lemma subst_atm_mset_empty[simp]: "{#} \am \ = {#}" unfolding subst_atm_mset_def by auto lemma subst_atm_mset_list_empty[simp]: "[] \aml \ = []" unfolding subst_atm_mset_list_def by auto lemma subst_atm_mset_lists_empty[simp]: "[] \\aml \s = []" unfolding subst_atm_mset_lists_def by auto lemma subst_cls_empty[simp]: "{#} \ \ = {#}" unfolding subst_cls_def by auto lemma subst_clss_empty[simp]: "{} \cs \ = {}" unfolding subst_clss_def by auto lemma subst_cls_list_empty[simp]: "[] \cl \ = []" unfolding subst_cls_list_def by auto lemma subst_cls_lists_empty[simp]: "[] \\cl \s = []" unfolding subst_cls_lists_def by auto lemma subst_scls_mset_empty[simp]: "{#} \cm \ = {#}" unfolding subst_cls_mset_def by auto lemma subst_atms_empty_iff[simp]: "AA \as \ = {} \ AA = {}" unfolding subst_atms_def by auto lemma subst_atmss_empty_iff[simp]: "AAA \ass \ = {} \ AAA = {}" unfolding subst_atmss_def by auto lemma subst_atm_list_empty_iff[simp]: "As \al \ = [] \ As = []" unfolding subst_atm_list_def by auto lemma subst_atm_mset_empty_iff[simp]: "AA \am \ = {#} \ AA = {#}" unfolding subst_atm_mset_def by auto lemma subst_atm_mset_list_empty_iff[simp]: "AAs \aml \ = [] \ AAs = []" unfolding subst_atm_mset_list_def by auto lemma subst_atm_mset_lists_empty_iff[simp]: "AAs \\aml \s = [] \ (AAs = [] \ \s = [])" using map2_empty_iff subst_atm_mset_lists_def by auto lemma subst_cls_empty_iff[simp]: "C \ \ = {#} \ C = {#}" unfolding subst_cls_def by auto lemma subst_clss_empty_iff[simp]: "CC \cs \ = {} \ CC = {}" unfolding subst_clss_def by auto lemma subst_cls_list_empty_iff[simp]: "Cs \cl \ = [] \ Cs = []" unfolding subst_cls_list_def by auto lemma subst_cls_lists_empty_iff[simp]: "Cs \\cl \s = [] \ (Cs = [] \ \s = [])" using map2_empty_iff subst_cls_lists_def by auto lemma subst_cls_mset_empty_iff[simp]: "CC \cm \ = {#} \ CC = {#}" unfolding subst_cls_mset_def by auto subsubsection \Substitution on a Union\ lemma subst_atms_union[simp]: "(AA \ BB) \as \ = AA \as \ \ BB \as \" unfolding subst_atms_def by auto lemma subst_atmss_union[simp]: "(AAA \ BBB) \ass \ = AAA \ass \ \ BBB \ass \" unfolding subst_atmss_def by auto lemma subst_atm_list_append[simp]: "(As @ Bs) \al \ = As \al \ @ Bs \al \" unfolding subst_atm_list_def by auto lemma subst_atm_mset_union[simp]: "(AA + BB) \am \ = AA \am \ + BB \am \" unfolding subst_atm_mset_def by auto lemma subst_atm_mset_list_append[simp]: "(AAs @ BBs) \aml \ = AAs \aml \ @ BBs \aml \" unfolding subst_atm_mset_list_def by auto lemma subst_cls_union[simp]: "(C + D) \ \ = C \ \ + D \ \" unfolding subst_cls_def by auto lemma subst_clss_union[simp]: "(CC \ DD) \cs \ = CC \cs \ \ DD \cs \" unfolding subst_clss_def by auto lemma subst_cls_list_append[simp]: "(Cs @ Ds) \cl \ = Cs \cl \ @ Ds \cl \" unfolding subst_cls_list_def by auto lemma subst_cls_mset_union[simp]: "(CC + DD) \cm \ = CC \cm \ + DD \cm \" unfolding subst_cls_mset_def by auto subsubsection \Substitution on a Singleton\ lemma subst_atms_single[simp]: "{A} \as \ = {A \a \}" unfolding subst_atms_def by auto lemma subst_atmss_single[simp]: "{AA} \ass \ = {AA \as \}" unfolding subst_atmss_def by auto lemma subst_atm_list_single[simp]: "[A] \al \ = [A \a \]" unfolding subst_atm_list_def by auto lemma subst_atm_mset_single[simp]: "{#A#} \am \ = {#A \a \#}" unfolding subst_atm_mset_def by auto lemma subst_atm_mset_list[simp]: "[AA] \aml \ = [AA \am \]" unfolding subst_atm_mset_list_def by auto lemma subst_cls_single[simp]: "{#L#} \ \ = {#L \l \#}" by simp lemma subst_clss_single[simp]: "{C} \cs \ = {C \ \}" unfolding subst_clss_def by auto lemma subst_cls_list_single[simp]: "[C] \cl \ = [C \ \]" unfolding subst_cls_list_def by auto lemma subst_cls_mset_single[simp]: "{#C#} \cm \ = {#C \ \#}" by simp subsubsection \Substitution on @{term Cons}\ lemma subst_atm_list_Cons[simp]: "(A # As) \al \ = A \a \ # As \al \" unfolding subst_atm_list_def by auto lemma subst_atm_mset_list_Cons[simp]: "(A # As) \aml \ = A \am \ # As \aml \" unfolding subst_atm_mset_list_def by auto lemma subst_atm_mset_lists_Cons[simp]: "(C # Cs) \\aml (\ # \s) = C \am \ # Cs \\aml \s" unfolding subst_atm_mset_lists_def by auto lemma subst_cls_list_Cons[simp]: "(C # Cs) \cl \ = C \ \ # Cs \cl \" unfolding subst_cls_list_def by auto lemma subst_cls_lists_Cons[simp]: "(C # Cs) \\cl (\ # \s) = C \ \ # Cs \\cl \s" unfolding subst_cls_lists_def by auto subsubsection \Substitution on @{term tl}\ lemma subst_atm_list_tl[simp]: "tl (As \al \) = tl As \al \" by (induction As) auto lemma subst_atm_mset_list_tl[simp]: "tl (AAs \aml \) = tl AAs \aml \" by (induction AAs) auto subsubsection \Substitution on @{term nth}\ lemma comp_substs_nth[simp]: "length \s = length \s \ i < length \s \ (\s \s \s) ! i = (\s ! i) \ (\s ! i)" by (simp add: comp_substs_def) lemma subst_atm_list_nth[simp]: "i < length As \ (As \al \) ! i = As ! i \a \" unfolding subst_atm_list_def using less_Suc_eq_0_disj nth_map by force lemma subst_atm_mset_list_nth[simp]: "i < length AAs \ (AAs \aml \) ! i = (AAs ! i) \am \" unfolding subst_atm_mset_list_def by auto lemma subst_atm_mset_lists_nth[simp]: "length AAs = length \s \ i < length AAs \ (AAs \\aml \s) ! i = (AAs ! i) \am (\s ! i)" unfolding subst_atm_mset_lists_def by auto lemma subst_cls_list_nth[simp]: "i < length Cs \ (Cs \cl \) ! i = (Cs ! i) \ \" unfolding subst_cls_list_def using less_Suc_eq_0_disj nth_map by (induction Cs) auto lemma subst_cls_lists_nth[simp]: "length Cs = length \s \ i < length Cs \ (Cs \\cl \s) ! i = (Cs ! i) \ (\s ! i)" unfolding subst_cls_lists_def by auto subsubsection \Substitution on Various Other Functions\ lemma subst_clss_image[simp]: "image f X \cs \ = {f x \ \ | x. x \ X}" unfolding subst_clss_def by auto lemma subst_cls_mset_image_mset[simp]: "image_mset f X \cm \ = {# f x \ \. x \# X #}" unfolding subst_cls_mset_def by auto lemma mset_subst_atm_list_subst_atm_mset[simp]: "mset (As \al \) = mset (As) \am \" unfolding subst_atm_list_def subst_atm_mset_def by auto lemma mset_subst_cls_list_subst_cls_mset: "mset (Cs \cl \) = (mset Cs) \cm \" unfolding subst_cls_mset_def subst_cls_list_def by auto lemma sum_list_subst_cls_list_subst_cls[simp]: "sum_list (Cs \cl \) = sum_list Cs \ \" unfolding subst_cls_list_def by (induction Cs) auto lemma set_mset_subst_cls_mset_subst_clss: "set_mset (CC \cm \) = (set_mset CC) \cs \" by (simp add: subst_cls_mset_def subst_clss_def) lemma Neg_Melem_subst_atm_subst_cls[simp]: "Neg A \# C \ Neg (A \a \) \# C \ \ " by (metis Melem_subst_cls eql_neg_lit_eql_atm) lemma Pos_Melem_subst_atm_subst_cls[simp]: "Pos A \# C \ Pos (A \a \) \# C \ \ " by (metis Melem_subst_cls eql_pos_lit_eql_atm) lemma in_atms_of_subst[simp]: "B \ atms_of C \ B \a \ \ atms_of (C \ \)" by (metis atms_of_subst_atms image_iff subst_atms_def) subsubsection \Renamings\ lemma is_renaming_id_subst[simp]: "is_renaming id_subst" unfolding is_renaming_def by simp lemma is_renamingD: "is_renaming \ \ (\A1 A2. A1 \a \ = A2 \a \ \ A1 = A2)" by (metis is_renaming_def subst_atm_comp_subst subst_atm_id_subst) lemma inv_renaming_cancel_r[simp]: "is_renaming r \ r \ inv_renaming r = id_subst" unfolding inv_renaming_def is_renaming_def by (metis (mono_tags) someI_ex) lemma inv_renaming_cancel_r_list[simp]: "is_renaming_list rs \ rs \s map inv_renaming rs = replicate (length rs) id_subst" unfolding is_renaming_list_def by (induction rs) (auto simp add: comp_substs_def) lemma Nil_comp_substs[simp]: "[] \s s = []" unfolding comp_substs_def by auto lemma comp_substs_Nil[simp]: "s \s [] = []" unfolding comp_substs_def by auto lemma is_renaming_idempotent_id_subst: "is_renaming r \ r \ r = r \ r = id_subst" by (metis comp_subst_assoc comp_subst_id_subst inv_renaming_cancel_r) lemma is_renaming_left_id_subst_right_id_subst: "is_renaming r \ s \ r = id_subst \ r \ s = id_subst" by (metis comp_subst_assoc comp_subst_id_subst is_renaming_def) lemma is_renaming_closure: "is_renaming r1 \ is_renaming r2 \ is_renaming (r1 \ r2)" unfolding is_renaming_def by (metis comp_subst_assoc comp_subst_id_subst) lemma is_renaming_inv_renaming_cancel_atm[simp]: "is_renaming \ \ A \a \ \a inv_renaming \ = A" by (metis inv_renaming_cancel_r subst_atm_comp_subst subst_atm_id_subst) lemma is_renaming_inv_renaming_cancel_atms[simp]: "is_renaming \ \ AA \as \ \as inv_renaming \ = AA" by (metis inv_renaming_cancel_r subst_atms_comp_subst subst_atms_id_subst) lemma is_renaming_inv_renaming_cancel_atmss[simp]: "is_renaming \ \ AAA \ass \ \ass inv_renaming \ = AAA" by (metis inv_renaming_cancel_r subst_atmss_comp_subst subst_atmss_id_subst) lemma is_renaming_inv_renaming_cancel_atm_list[simp]: "is_renaming \ \ As \al \ \al inv_renaming \ = As" by (metis inv_renaming_cancel_r subst_atm_list_comp_subst subst_atm_list_id_subst) lemma is_renaming_inv_renaming_cancel_atm_mset[simp]: "is_renaming \ \ AA \am \ \am inv_renaming \ = AA" by (metis inv_renaming_cancel_r subst_atm_mset_comp_subst subst_atm_mset_id_subst) lemma is_renaming_inv_renaming_cancel_atm_mset_list[simp]: "is_renaming \ \ (AAs \aml \) \aml inv_renaming \ = AAs" by (metis inv_renaming_cancel_r subst_atm_mset_list_comp_subst subst_atm_mset_list_id_subst) lemma is_renaming_list_inv_renaming_cancel_atm_mset_lists[simp]: "length AAs = length \s \ is_renaming_list \s \ AAs \\aml \s \\aml map inv_renaming \s = AAs" by (metis inv_renaming_cancel_r_list subst_atm_mset_lists_comp_substs subst_atm_mset_lists_id_subst) lemma is_renaming_inv_renaming_cancel_lit[simp]: "is_renaming \ \ (L \l \) \l inv_renaming \ = L" by (metis inv_renaming_cancel_r subst_lit_comp_subst subst_lit_id_subst) lemma is_renaming_inv_renaming_cancel_cls[simp]: "is_renaming \ \ C \ \ \ inv_renaming \ = C" by (metis inv_renaming_cancel_r subst_cls_comp_subst subst_cls_id_subst) lemma is_renaming_inv_renaming_cancel_clss[simp]: "is_renaming \ \ CC \cs \ \cs inv_renaming \ = CC" by (metis inv_renaming_cancel_r subst_clss_id_subst subst_clsscomp_subst) lemma is_renaming_inv_renaming_cancel_cls_list[simp]: "is_renaming \ \ Cs \cl \ \cl inv_renaming \ = Cs" by (metis inv_renaming_cancel_r subst_cls_list_comp_subst subst_cls_list_id_subst) lemma is_renaming_list_inv_renaming_cancel_cls_list[simp]: "length Cs = length \s \ is_renaming_list \s \ Cs \\cl \s \\cl map inv_renaming \s = Cs" by (metis inv_renaming_cancel_r_list subst_cls_lists_comp_substs subst_cls_lists_id_subst) lemma is_renaming_inv_renaming_cancel_cls_mset[simp]: "is_renaming \ \ CC \cm \ \cm inv_renaming \ = CC" by (metis inv_renaming_cancel_r subst_cls_mset_comp_subst subst_cls_mset_id_subst) subsubsection \Monotonicity\ lemma subst_cls_mono: "set_mset C \ set_mset D \ set_mset (C \ \) \ set_mset (D \ \)" by force lemma subst_cls_mono_mset: "C \# D \ C \ \ \# D \ \" unfolding subst_clss_def by (metis mset_subset_eq_exists_conv subst_cls_union) lemma subst_subset_mono: "D \# C \ D \ \ \# C \ \" unfolding subst_cls_def by (simp add: image_mset_subset_mono) subsubsection \Size after Substitution\ lemma size_subst[simp]: "size (D \ \) = size D" unfolding subst_cls_def by auto lemma subst_atm_list_length[simp]: "length (As \al \) = length As" unfolding subst_atm_list_def by auto lemma length_subst_atm_mset_list[simp]: "length (AAs \aml \) = length AAs" unfolding subst_atm_mset_list_def by auto lemma subst_atm_mset_lists_length[simp]: "length (AAs \\aml \s) = min (length AAs) (length \s)" unfolding subst_atm_mset_lists_def by auto lemma subst_cls_list_length[simp]: "length (Cs \cl \) = length Cs" unfolding subst_cls_list_def by auto lemma comp_substs_length[simp]: "length (\s \s \s) = min (length \s) (length \s)" unfolding comp_substs_def by auto lemma subst_cls_lists_length[simp]: "length (Cs \\cl \s) = min (length Cs) (length \s)" unfolding subst_cls_lists_def by auto subsubsection \Variable Disjointness\ lemma var_disjoint_clauses: assumes "var_disjoint Cs" shows "\\s. length \s = length Cs \ (\\. Cs \\cl \s = Cs \cl \)" proof clarify fix \s :: "'s list" assume a: "length \s = length Cs" then obtain \ where "\i < length Cs. \S. S \# Cs ! i \ S \ \s ! i = S \ \" using assms unfolding var_disjoint_def by blast then have "\i < length Cs. (Cs ! i) \ \s ! i = (Cs ! i) \ \" by auto then have "Cs \\cl \s = Cs \cl \" using a by (auto intro: nth_equalityI) then show "\\. Cs \\cl \s = Cs \cl \" by auto qed subsubsection \Ground Expressions and Substitutions\ lemma ex_ground_subst: "\\. is_ground_subst \" using make_ground_subst[of "{#}"] by (simp add: is_ground_cls_def) lemma is_ground_cls_list_Cons[simp]: "is_ground_cls_list (C # Cs) = (is_ground_cls C \ is_ground_cls_list Cs)" unfolding is_ground_cls_list_def by auto paragraph \Ground union\ lemma is_ground_atms_union[simp]: "is_ground_atms (AA \ BB) \ is_ground_atms AA \ is_ground_atms BB" unfolding is_ground_atms_def by auto lemma is_ground_atm_mset_union[simp]: "is_ground_atm_mset (AA + BB) \ is_ground_atm_mset AA \ is_ground_atm_mset BB" unfolding is_ground_atm_mset_def by auto lemma is_ground_cls_union[simp]: "is_ground_cls (C + D) \ is_ground_cls C \ is_ground_cls D" unfolding is_ground_cls_def by auto lemma is_ground_clss_union[simp]: "is_ground_clss (CC \ DD) \ is_ground_clss CC \ is_ground_clss DD" unfolding is_ground_clss_def by auto lemma is_ground_cls_list_is_ground_cls_sum_list[simp]: "is_ground_cls_list Cs \ is_ground_cls (sum_list Cs)" by (meson in_mset_sum_list2 is_ground_cls_def is_ground_cls_list_def) paragraph \Ground mono\ lemma is_ground_cls_mono: "C \# D \ is_ground_cls D \ is_ground_cls C" unfolding is_ground_cls_def by (metis set_mset_mono subsetD) lemma is_ground_clss_mono: "CC \ DD \ is_ground_clss DD \ is_ground_clss CC" unfolding is_ground_clss_def by blast lemma grounding_of_clss_mono: "CC \ DD \ grounding_of_clss CC \ grounding_of_clss DD" using grounding_of_clss_def by auto lemma sum_list_subseteq_mset_is_ground_cls_list[simp]: "sum_list Cs \# sum_list Ds \ is_ground_cls_list Ds \ is_ground_cls_list Cs" by (meson in_mset_sum_list is_ground_cls_def is_ground_cls_list_is_ground_cls_sum_list is_ground_cls_mono is_ground_cls_list_def) paragraph \Substituting on ground expression preserves ground\ lemma is_ground_comp_subst[simp]: "is_ground_subst \ \ is_ground_subst (\ \ \)" unfolding is_ground_subst_def is_ground_atm_def by auto lemma ground_subst_ground_atm[simp]: "is_ground_subst \ \ is_ground_atm (A \a \)" by (simp add: is_ground_subst_def) lemma ground_subst_ground_lit[simp]: "is_ground_subst \ \ is_ground_lit (L \l \)" unfolding is_ground_lit_def subst_lit_def by (cases L) auto lemma ground_subst_ground_cls[simp]: "is_ground_subst \ \ is_ground_cls (C \ \)" unfolding is_ground_cls_def by auto lemma ground_subst_ground_clss[simp]: "is_ground_subst \ \ is_ground_clss (CC \cs \)" unfolding is_ground_clss_def subst_clss_def by auto lemma ground_subst_ground_cls_list[simp]: "is_ground_subst \ \ is_ground_cls_list (Cs \cl \)" unfolding is_ground_cls_list_def subst_cls_list_def by auto lemma ground_subst_ground_cls_lists[simp]: "\\ \ set \s. is_ground_subst \ \ is_ground_cls_list (Cs \\cl \s)" unfolding is_ground_cls_list_def subst_cls_lists_def by (auto simp: set_zip) paragraph \Substituting on ground expression has no effect\ lemma is_ground_subst_atm[simp]: "is_ground_atm A \ A \a \ = A" unfolding is_ground_atm_def by simp lemma is_ground_subst_atms[simp]: "is_ground_atms AA \ AA \as \ = AA" unfolding is_ground_atms_def subst_atms_def image_def by auto lemma is_ground_subst_atm_mset[simp]: "is_ground_atm_mset AA \ AA \am \ = AA" unfolding is_ground_atm_mset_def subst_atm_mset_def by auto lemma is_ground_subst_atm_list[simp]: "is_ground_atm_list As \ As \al \ = As" unfolding is_ground_atm_list_def subst_atm_list_def by (auto intro: nth_equalityI) lemma is_ground_subst_atm_list_member[simp]: "is_ground_atm_list As \ i < length As \ As ! i \a \ = As ! i" unfolding is_ground_atm_list_def by auto lemma is_ground_subst_lit[simp]: "is_ground_lit L \ L \l \ = L" unfolding is_ground_lit_def subst_lit_def by (cases L) simp_all lemma is_ground_subst_cls[simp]: "is_ground_cls C \ C \ \ = C" unfolding is_ground_cls_def subst_cls_def by simp lemma is_ground_subst_clss[simp]: "is_ground_clss CC \ CC \cs \ = CC" unfolding is_ground_clss_def subst_clss_def image_def by auto lemma is_ground_subst_cls_lists[simp]: assumes "length P = length Cs" and "is_ground_cls_list Cs" shows "Cs \\cl P = Cs" using assms by (metis is_ground_cls_list_def is_ground_subst_cls min.idem nth_equalityI nth_mem subst_cls_lists_nth subst_cls_lists_length) lemma is_ground_subst_lit_iff: "is_ground_lit L \ (\\. L = L \l \)" using is_ground_atm_def is_ground_lit_def subst_lit_def by (cases L) auto lemma is_ground_subst_cls_iff: "is_ground_cls C \ (\\. C = C \ \)" by (metis ex_ground_subst ground_subst_ground_cls is_ground_subst_cls) paragraph \Members of ground expressions are ground\ lemma is_ground_cls_as_atms: "is_ground_cls C \ (\A \ atms_of C. is_ground_atm A)" by (auto simp: atms_of_def is_ground_cls_def is_ground_lit_def) lemma is_ground_cls_imp_is_ground_lit: "L \# C \ is_ground_cls C \ is_ground_lit L" by (simp add: is_ground_cls_def) lemma is_ground_cls_imp_is_ground_atm: "A \ atms_of C \ is_ground_cls C \ is_ground_atm A" by (simp add: is_ground_cls_as_atms) lemma is_ground_cls_is_ground_atms_atms_of[simp]: "is_ground_cls C \ is_ground_atms (atms_of C)" by (simp add: is_ground_cls_imp_is_ground_atm is_ground_atms_def) lemma grounding_ground: "C \ grounding_of_clss M \ is_ground_cls C" unfolding grounding_of_clss_def grounding_of_cls_def by auto lemma in_subset_eq_grounding_of_clss_is_ground_cls[simp]: "C \ CC \ CC \ grounding_of_clss DD \ is_ground_cls C" unfolding grounding_of_clss_def grounding_of_cls_def by auto lemma is_ground_cls_empty[simp]: "is_ground_cls {#}" unfolding is_ground_cls_def by simp lemma grounding_of_cls_ground: "is_ground_cls C \ grounding_of_cls C = {C}" unfolding grounding_of_cls_def by (simp add: ex_ground_subst) lemma grounding_of_cls_empty[simp]: "grounding_of_cls {#} = {{#}}" by (simp add: grounding_of_cls_ground) subsubsection \Subsumption\ lemma subsumes_empty_left[simp]: "subsumes {#} C" unfolding subsumes_def subst_cls_def by simp lemma strictly_subsumes_empty_left[simp]: "strictly_subsumes {#} C \ C \ {#}" unfolding strictly_subsumes_def subsumes_def subst_cls_def by simp subsubsection \Unifiers\ lemma card_le_one_alt: "finite X \ card X \ 1 \ X = {} \ (\x. X = {x})" by (induct rule: finite_induct) auto lemma is_unifier_subst_atm_eqI: assumes "finite AA" shows "is_unifier \ AA \ A \ AA \ B \ AA \ A \a \ = B \a \" unfolding is_unifier_def subst_atms_def card_le_one_alt[OF finite_imageI[OF assms]] by (metis equals0D imageI insert_iff) lemma is_unifier_alt: assumes "finite AA" shows "is_unifier \ AA \ (\A \ AA. \B \ AA. A \a \ = B \a \)" unfolding is_unifier_def subst_atms_def card_le_one_alt[OF finite_imageI[OF assms(1)]] by (rule iffI, metis empty_iff insert_iff insert_image, blast) lemma is_unifiers_subst_atm_eqI: assumes "finite AA" "is_unifiers \ AAA" "AA \ AAA" "A \ AA" "B \ AA" shows "A \a \ = B \a \" by (metis assms is_unifiers_def is_unifier_subst_atm_eqI) theorem is_unifiers_comp: "is_unifiers \ (set_mset ` set (map2 add_mset As Bs) \ass \) \ is_unifiers (\ \ \) (set_mset ` set (map2 add_mset As Bs))" unfolding is_unifiers_def is_unifier_def subst_atmss_def by auto subsubsection \Most General Unifier\ lemma is_mgu_is_unifiers: "is_mgu \ AAA \ is_unifiers \ AAA" using is_mgu_def by blast lemma is_mgu_is_most_general: "is_mgu \ AAA \ is_unifiers \ AAA \ \\. \ = \ \ \" using is_mgu_def by blast lemma is_unifiers_is_unifier: "is_unifiers \ AAA \ AA \ AAA \ is_unifier \ AA" using is_unifiers_def by simp subsubsection \Generalization and Subsumption\ lemma variants_iff_subsumes: "variants C D \ subsumes C D \ subsumes D C" proof assume "variants C D" then show "subsumes C D \ subsumes D C" unfolding variants_def generalizes_cls_def subsumes_def by (metis subset_mset.order.refl) next assume sub: "subsumes C D \ subsumes D C" then have "size C = size D" unfolding subsumes_def by (metis antisym size_mset_mono size_subst) then show "variants C D" using sub unfolding subsumes_def variants_def generalizes_cls_def by (metis leD mset_subset_size size_mset_mono size_subst subset_mset.order.not_eq_order_implies_strict) qed lemma wf_strictly_generalizes_cls: "wfP strictly_generalizes_cls" proof - { assume "\C_at. \i. strictly_generalizes_cls (C_at (Suc i)) (C_at i)" then obtain C_at :: "nat \ 'a clause" where sg_C: "\i. strictly_generalizes_cls (C_at (Suc i)) (C_at i)" by blast define n :: nat where "n = size (C_at 0)" have sz_C: "size (C_at i) = n" for i proof (induct i) case (Suc i) then show ?case using sg_C[of i] unfolding strictly_generalizes_cls_def generalizes_cls_def subst_cls_def by (metis size_image_mset) qed (simp add: n_def) obtain \_at :: "nat \ 's" where C_\: "\i. image_mset (\L. L \l \_at i) (C_at (Suc i)) = C_at i" using sg_C[unfolded strictly_generalizes_cls_def generalizes_cls_def subst_cls_def] by metis define Ls_at :: "nat \ 'a literal list" where "Ls_at = rec_nat (SOME Ls. mset Ls = C_at 0) (\i Lsi. SOME Ls. mset Ls = C_at (Suc i) \ map (\L. L \l \_at i) Ls = Lsi)" have Ls_at_0: "Ls_at 0 = (SOME Ls. mset Ls = C_at 0)" and Ls_at_Suc: "\i. Ls_at (Suc i) = (SOME Ls. mset Ls = C_at (Suc i) \ map (\L. L \l \_at i) Ls = Ls_at i)" unfolding Ls_at_def by simp+ have mset_Lt_at_0: "mset (Ls_at 0) = C_at 0" unfolding Ls_at_0 by (rule someI_ex) (metis list_of_mset_exi) have "mset (Ls_at (Suc i)) = C_at (Suc i) \ map (\L. L \l \_at i) (Ls_at (Suc i)) = Ls_at i" for i proof (induct i) case 0 then show ?case by (simp add: Ls_at_Suc, rule someI_ex, metis C_\ image_mset_of_subset_list mset_Lt_at_0) next case Suc then show ?case by (subst (1 2) Ls_at_Suc) (rule someI_ex, metis C_\ image_mset_of_subset_list) qed note mset_Ls = this[THEN conjunct1] and Ls_\ = this[THEN conjunct2] have len_Ls: "\i. length (Ls_at i) = n" by (metis mset_Ls mset_Lt_at_0 not0_implies_Suc size_mset sz_C) have is_pos_Ls: "\i j. j < n \ is_pos (Ls_at (Suc i) ! j) \ is_pos (Ls_at i ! j)" using Ls_\ len_Ls by (metis literal.map_disc_iff nth_map subst_lit_def) have Ls_\_strict_lit: "\i \. map (\L. L \l \) (Ls_at i) \ Ls_at (Suc i)" by (metis C_\ mset_Ls Ls_\ mset_map sg_C generalizes_cls_def strictly_generalizes_cls_def subst_cls_def) have Ls_\_strict_tm: "map ((\t. t \a \) \ atm_of) (Ls_at i) \ map atm_of (Ls_at (Suc i))" for i \ proof - obtain j :: nat where j_lt: "j < n" and j_\: "Ls_at i ! j \l \ \ Ls_at (Suc i) ! j" using Ls_\_strict_lit[of \ i] len_Ls by (metis (no_types, lifting) length_map list_eq_iff_nth_eq nth_map) have "atm_of (Ls_at i ! j) \a \ \ atm_of (Ls_at (Suc i) ! j)" using j_\ is_pos_Ls[OF j_lt] by (metis (mono_guards) literal.expand literal.map_disc_iff literal.map_sel subst_lit_def) then show ?thesis using j_lt len_Ls by (metis nth_map o_apply) qed define tm_at :: "nat \ 'a" where "\i. tm_at i = atm_of_atms (map atm_of (Ls_at i))" have "\i. generalizes_atm (tm_at (Suc i)) (tm_at i)" unfolding tm_at_def generalizes_atm_def atm_of_atms_subst using Ls_\[THEN arg_cong, of "map atm_of"] by (auto simp: comp_def) moreover have "\i. \ generalizes_atm (tm_at i) (tm_at (Suc i))" unfolding tm_at_def generalizes_atm_def atm_of_atms_subst by (simp add: Ls_\_strict_tm) ultimately have "\i. strictly_generalizes_atm (tm_at (Suc i)) (tm_at i)" unfolding strictly_generalizes_atm_def by blast then have False using wf_strictly_generalizes_atm[unfolded wfP_def wf_iff_no_infinite_down_chain] by blast } then show "wfP (strictly_generalizes_cls :: 'a clause \ _ \ _)" unfolding wfP_def by (blast intro: wf_iff_no_infinite_down_chain[THEN iffD2]) qed lemma strict_subset_subst_strictly_subsumes: assumes c\_sub: "C \ \ \# D" shows "strictly_subsumes C D" by (metis c\_sub leD mset_subset_size size_mset_mono size_subst strictly_subsumes_def subset_mset.dual_order.strict_implies_order substitution_ops.subsumes_def) lemma subsumes_trans: "subsumes C D \ subsumes D E \ subsumes C E" unfolding subsumes_def by (metis (no_types) subset_mset.order.trans subst_cls_comp_subst subst_cls_mono_mset) lemma subset_strictly_subsumes: "C \# D \ strictly_subsumes C D" using strict_subset_subst_strictly_subsumes[of C id_subst] by auto lemma strictly_subsumes_neq: "strictly_subsumes D' D \ D' \ D \ \" unfolding strictly_subsumes_def subsumes_def by blast lemma strictly_subsumes_has_minimum: assumes "CC \ {}" shows "\C \ CC. \D \ CC. \ strictly_subsumes D C" proof (rule ccontr) assume "\ (\C \ CC. \D\CC. \ strictly_subsumes D C)" then have "\C \ CC. \D \ CC. strictly_subsumes D C" by blast then obtain f where f_p: "\C \ CC. f C \ CC \ strictly_subsumes (f C) C" by metis from assms obtain C where C_p: "C \ CC" by auto define c :: "nat \ 'a clause" where "\n. c n = (f ^^ n) C" have incc: "c i \ CC" for i by (induction i) (auto simp: c_def f_p C_p) have ps: "\i. strictly_subsumes (c (Suc i)) (c i)" using incc f_p unfolding c_def by auto have "\i. size (c i) \ size (c (Suc i))" using ps unfolding strictly_subsumes_def subsumes_def by (metis size_mset_mono size_subst) then have lte: "\i. (size \ c) i \ (size \ c) (Suc i)" unfolding comp_def . then have "\l. \l' \ l. size (c l') = size (c (Suc l'))" using f_Suc_decr_eventually_const comp_def by auto then obtain l where l_p: "\l' \ l. size (c l') = size (c (Suc l'))" by metis then have "\l' \ l. strictly_generalizes_cls (c (Suc l')) (c l')" using ps unfolding strictly_generalizes_cls_def generalizes_cls_def by (metis size_subst less_irrefl strictly_subsumes_def mset_subset_size subset_mset_def subsumes_def strictly_subsumes_neq) then have "\i. strictly_generalizes_cls (c (Suc i + l)) (c (i + l))" unfolding strictly_generalizes_cls_def generalizes_cls_def by auto then have "\f. \i. strictly_generalizes_cls (f (Suc i)) (f i)" by (rule exI[of _ "\x. c (x + l)"]) then show False using wf_strictly_generalizes_cls wf_iff_no_infinite_down_chain[of "{(x, y). strictly_generalizes_cls x y}"] unfolding wfP_def by auto qed end subsection \Most General Unifiers\ locale mgu = substitution subst_atm id_subst comp_subst renamings_apart atm_of_atms for subst_atm :: "'a \ 's \ 'a" and id_subst :: 's and comp_subst :: "'s \ 's \ 's" and atm_of_atms :: "'a list \ 'a" and renamings_apart :: "'a literal multiset list \ 's list" + fixes mgu :: "'a set set \ 's option" assumes mgu_sound: "finite AAA \ (\AA \ AAA. finite AA) \ mgu AAA = Some \ \ is_mgu \ AAA" and mgu_complete: "finite AAA \ (\AA \ AAA. finite AA) \ is_unifiers \ AAA \ \\. mgu AAA = Some \" begin lemmas is_unifiers_mgu = mgu_sound[unfolded is_mgu_def, THEN conjunct1] lemmas is_mgu_most_general = mgu_sound[unfolded is_mgu_def, THEN conjunct2] lemma mgu_unifier: assumes aslen: "length As = n" and aaslen: "length AAs = n" and mgu: "Some \ = mgu (set_mset ` set (map2 add_mset As AAs))" and i_lt: "i < n" and a_in: "A \# AAs ! i" shows "A \a \ = As ! i \a \" proof - from mgu have "is_mgu \ (set_mset ` set (map2 add_mset As AAs))" using mgu_sound by auto then have "is_unifiers \ (set_mset ` set (map2 add_mset As AAs))" using is_mgu_is_unifiers by auto then have "is_unifier \ (set_mset (add_mset (As ! i) (AAs ! i)))" using i_lt aslen aaslen unfolding is_unifiers_def is_unifier_def by simp (metis length_zip min.idem nth_mem nth_zip prod.case set_mset_add_mset_insert) then show ?thesis using aslen aaslen a_in is_unifier_subst_atm_eqI by (metis finite_set_mset insertCI set_mset_add_mset_insert) qed end end diff --git a/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution.thy b/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution.thy --- a/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution.thy +++ b/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution.thy @@ -1,1309 +1,1296 @@ (* Title: First-Order Ordered Resolution Calculus with Selection Author: Anders Schlichtkrull , 2016, 2017 Author: Jasmin Blanchette , 2014, 2017 Author: Dmitriy Traytel , 2014 Author: Sophie Tourret , 2020 Maintainer: Anders Schlichtkrull *) section \First-Order Ordered Resolution Calculus with Selection\ theory FO_Ordered_Resolution imports Abstract_Substitution Ordered_Ground_Resolution Standard_Redundancy begin text \ This material is based on Section 4.3 (``A Simple Resolution Prover for First-Order Clauses'') of Bachmair and Ganzinger's chapter. Specifically, it formalizes the ordered resolution calculus for first-order standard clauses presented in Figure 4 and its related lemmas and theorems, including soundness and Lemma 4.12 (the lifting lemma). The following corresponds to pages 41--42 of Section 4.3, until Figure 5 and its explanation. \ locale FO_resolution = mgu subst_atm id_subst comp_subst atm_of_atms renamings_apart mgu for subst_atm :: "'a :: wellorder \ 's \ 'a" and id_subst :: "'s" and comp_subst :: "'s \ 's \ 's" and renamings_apart :: "'a literal multiset list \ 's list" and atm_of_atms :: "'a list \ 'a" and mgu :: "'a set set \ 's option" + fixes less_atm :: "'a \ 'a \ bool" assumes less_atm_stable: "less_atm A B \ less_atm (A \a \) (B \a \)" begin subsection \Library\ lemma Bex_cartesian_product: "(\xy \ A \ B. P xy) \ (\x \ A. \y \ B. P (x, y))" by simp -(* FIXME: Move to "Multiset.thy" *) -lemma length_sorted_list_of_multiset[simp]: "length (sorted_list_of_multiset A) = size A" - by (metis mset_sorted_list_of_multiset size_mset) - -(* FIXME: move? or avoid? *) lemma eql_map_neg_lit_eql_atm: assumes "map (\L. L \l \) (map Neg As') = map Neg As" shows "As' \al \ = As" using assms by (induction As' arbitrary: As) auto lemma instance_list: assumes "negs (mset As) = SDA' \ \" shows "\As'. negs (mset As') = SDA' \ As' \al \ = As" proof - from assms have negL: "\L \# SDA'. is_neg L" using Melem_subst_cls subst_lit_in_negs_is_neg by metis from assms have "{#L \l \. L \# SDA'#} = mset (map Neg As)" using subst_cls_def by auto then have "\NAs'. map (\L. L \l \) NAs' = map Neg As \ mset NAs' = SDA'" using image_mset_of_subset_list[of "\L. L \l \" SDA' "map Neg As"] by auto then obtain As' where As'_p: "map (\L. L \l \) (map Neg As') = map Neg As \ mset (map Neg As') = SDA'" by (metis (no_types, lifting) Neg_atm_of_iff negL ex_map_conv set_mset_mset) have "negs (mset As') = SDA'" using As'_p by auto moreover have "map (\L. L \l \) (map Neg As') = map Neg As" using As'_p by auto then have "As' \al \ = As" using eql_map_neg_lit_eql_atm by auto ultimately show ?thesis by blast qed context fixes S :: "'a clause \ 'a clause" begin subsection \Calculus\ text \ The following corresponds to Figure 4. \ definition maximal_wrt :: "'a \ 'a literal multiset \ bool" where "maximal_wrt A C \ (\B \ atms_of C. \ less_atm A B)" definition strictly_maximal_wrt :: "'a \ 'a literal multiset \ bool" where "strictly_maximal_wrt A C \ \B \ atms_of C. A \ B \ \ less_atm A B" lemma strictly_maximal_wrt_maximal_wrt: "strictly_maximal_wrt A C \ maximal_wrt A C" unfolding maximal_wrt_def strictly_maximal_wrt_def by auto -(* FIXME: use hd instead of ! 0 *) inductive eligible :: "'s \ 'a list \ 'a clause \ bool" where eligible: "S DA = negs (mset As) \ S DA = {#} \ length As = 1 \ maximal_wrt (As ! 0 \a \) (DA \ \) \ eligible \ As DA" inductive ord_resolve :: "'a clause list \ 'a clause \ 'a multiset list \ 'a list \ 's \ 'a clause \ bool" where ord_resolve: "length CAs = n \ length Cs = n \ length AAs = n \ length As = n \ n \ 0 \ (\i < n. CAs ! i = Cs ! i + poss (AAs ! i)) \ (\i < n. AAs ! i \ {#}) \ Some \ = mgu (set_mset ` set (map2 add_mset As AAs)) \ eligible \ As (D + negs (mset As)) \ (\i < n. strictly_maximal_wrt (As ! i \a \) (Cs ! i \ \)) \ (\i < n. S (CAs ! i) = {#}) \ ord_resolve CAs (D + negs (mset As)) AAs As \ ((\# (mset Cs) + D) \ \)" inductive ord_resolve_rename :: "'a clause list \ 'a clause \ 'a multiset list \ 'a list \ 's \ 'a clause \ bool" where ord_resolve_rename: "length CAs = n \ length AAs = n \ length As = n \ (\i < n. poss (AAs ! i) \# CAs ! i) \ negs (mset As) \# DA \ \ = hd (renamings_apart (DA # CAs)) \ \s = tl (renamings_apart (DA # CAs)) \ ord_resolve (CAs \\cl \s) (DA \ \) (AAs \\aml \s) (As \al \) \ E \ ord_resolve_rename CAs DA AAs As \ E" lemma ord_resolve_empty_main_prem: "\ ord_resolve Cs {#} AAs As \ E" by (simp add: ord_resolve.simps) lemma ord_resolve_rename_empty_main_prem: "\ ord_resolve_rename Cs {#} AAs As \ E" by (simp add: ord_resolve_empty_main_prem ord_resolve_rename.simps) subsection \Soundness\ text \ Soundness is not discussed in the chapter, but it is an important property. \ lemma ord_resolve_ground_inst_sound: assumes res_e: "ord_resolve CAs DA AAs As \ E" and cc_inst_true: "I \m mset CAs \cm \ \cm \" and d_inst_true: "I \ DA \ \ \ \" and ground_subst_\: "is_ground_subst \" shows "I \ E \ \" using res_e proof (cases rule: ord_resolve.cases) case (ord_resolve n Cs D) note da = this(1) and e = this(2) and cas_len = this(3) and cs_len = this(4) and aas_len = this(5) and as_len = this(6) and cas = this(8) and mgu = this(10) and len = this(1) have len: "length CAs = length As" using as_len cas_len by auto have "is_ground_subst (\ \ \)" using ground_subst_\ by (rule is_ground_comp_subst) then have cc_true: "I \m mset CAs \cm \ \cm \" and d_true: "I \ DA \ \ \ \" using cc_inst_true d_inst_true by auto from mgu have unif: "\i < n. \A\#AAs ! i. A \a \ = As ! i \a \" using mgu_unifier as_len aas_len by blast show "I \ E \ \" proof (cases "\A \ set As. A \a \ \a \ \ I") case True then have "\ I \ negs (mset As) \ \ \ \" unfolding true_cls_def[of I] by auto then have "I \ D \ \ \ \" using d_true da by auto then show ?thesis unfolding e by auto next case False then obtain i where a_in_aa: "i < length CAs" and a_false: "(As ! i) \a \ \a \ \ I" using da len by (metis in_set_conv_nth) define C where "C \ Cs ! i" define BB where "BB \ AAs ! i" have c_cf': "C \# \# (mset CAs)" unfolding C_def using a_in_aa cas cas_len by (metis less_subset_eq_Union_mset mset_subset_eq_add_left subset_mset.order.trans) have c_in_cc: "C + poss BB \# mset CAs" using C_def BB_def a_in_aa cas_len in_set_conv_nth cas by fastforce { fix B assume "B \# BB" then have "B \a \ = (As ! i) \a \" using unif a_in_aa cas_len unfolding BB_def by auto } then have "\ I \ poss BB \ \ \ \" using a_false by (auto simp: true_cls_def) moreover have "I \ (C + poss BB) \ \ \ \" using c_in_cc cc_true true_cls_mset_true_cls[of I "mset CAs \cm \ \cm \"] by force ultimately have "I \ C \ \ \ \" by simp then show ?thesis unfolding e subst_cls_union using c_cf' C_def a_in_aa cas_len cs_len by (metis (no_types, lifting) mset_subset_eq_add_left nth_mem_mset set_mset_mono sum_mset.remove true_cls_mono subst_cls_mono) qed qed text \ The previous lemma is not only used to prove soundness, but also the following lemma which is used to prove Lemma 4.10. \ lemma ord_resolve_rename_ground_inst_sound: assumes "ord_resolve_rename CAs DA AAs As \ E" and "\s = tl (renamings_apart (DA # CAs))" and "\ = hd (renamings_apart (DA # CAs))" and "I \m (mset (CAs \\cl \s)) \cm \ \cm \" and "I \ DA \ \ \ \ \ \" and "is_ground_subst \" shows "I \ E \ \" using assms by (cases rule: ord_resolve_rename.cases) (fast intro: ord_resolve_ground_inst_sound) text \ Here follows the soundness theorem for the resolution rule. \ theorem ord_resolve_sound: assumes res_e: "ord_resolve CAs DA AAs As \ E" and cc_d_true: "\\. is_ground_subst \ \ I \m (mset CAs + {#DA#}) \cm \" and ground_subst_\: "is_ground_subst \" shows "I \ E \ \" proof (use res_e in \cases rule: ord_resolve.cases\) case (ord_resolve n Cs D) note da = this(1) and e = this(2) and cas_len = this(3) and cs_len = this(4) and aas_len = this(5) and as_len = this(6) and cas = this(8) and mgu = this(10) have ground_subst_\_\: "is_ground_subst (\ \ \)" using ground_subst_\ by (rule is_ground_comp_subst) have cas_true: "I \m mset CAs \cm \ \cm \" using cc_d_true ground_subst_\_\ by fastforce have da_true: "I \ DA \ \ \ \" using cc_d_true ground_subst_\_\ by fastforce show "I \ E \ \" using ord_resolve_ground_inst_sound[OF res_e cas_true da_true] ground_subst_\ by auto qed lemma subst_sound: assumes "\\. is_ground_subst \ \ I \ (C \ \)" and "is_ground_subst \" shows "I \ (C \ \) \ \" using assms is_ground_comp_subst subst_cls_comp_subst by metis lemma subst_sound_scl: assumes len: "length P = length CAs" and true_cas: "\\. is_ground_subst \ \ I \m (mset CAs) \cm \" and ground_subst_\: "is_ground_subst \" shows "I \m mset (CAs \\cl P) \cm \" proof - from true_cas have "\CA. CA\# mset CAs \ (\\. is_ground_subst \ \ I \ CA \ \)" unfolding true_cls_mset_def by force then have "\i < length CAs. \\. is_ground_subst \ \ (I \ CAs ! i \ \)" using in_set_conv_nth by auto then have true_cp: "\i < length CAs. \\. is_ground_subst \ \ I \ CAs ! i \ P ! i \ \" using subst_sound len by auto { fix CA assume "CA \# mset (CAs \\cl P)" then obtain i where i_x: "i < length (CAs \\cl P)" "CA = (CAs \\cl P) ! i" by (metis in_mset_conv_nth) then have "\\. is_ground_subst \ \ I \ CA \ \" using true_cp unfolding subst_cls_lists_def by (simp add: len) } then show ?thesis using assms unfolding true_cls_mset_def by auto qed text \ Here follows the soundness theorem for the resolution rule with renaming. \ lemma ord_resolve_rename_sound: assumes res_e: "ord_resolve_rename CAs DA AAs As \ E" and cc_d_true: "\\. is_ground_subst \ \ I \m ((mset CAs) + {#DA#}) \cm \" and ground_subst_\: "is_ground_subst \" shows "I \ E \ \" using res_e proof (cases rule: ord_resolve_rename.cases) case (ord_resolve_rename n \ \s) note \s = this(7) and res = this(8) have len: "length \s = length CAs" using \s renamings_apart_length by auto have "\\. is_ground_subst \ \ I \m (mset (CAs \\cl \s) + {#DA \ \#}) \cm \" using subst_sound_scl[OF len, of I] subst_sound cc_d_true by auto then show "I \ E \ \" using ground_subst_\ ord_resolve_sound[OF res] by simp qed subsection \Other Basic Properties\ lemma ord_resolve_unique: assumes "ord_resolve CAs DA AAs As \ E" and "ord_resolve CAs DA AAs As \' E'" shows "\ = \' \ E = E'" using assms proof (cases rule: ord_resolve.cases[case_product ord_resolve.cases], intro conjI) case (ord_resolve_ord_resolve CAs n Cs AAs As \'' DA CAs' n' Cs' AAs' As' \''' DA') note res = this(1-17) and res' = this(18-34) show \: "\ = \'" using res(3-5,14) res'(3-5,14) by (metis option.inject) have "Cs = Cs'" using res(1,3,7,8,12) res'(1,3,7,8,12) by (metis add_right_imp_eq nth_equalityI) moreover have "DA = DA'" using res(2,4) res'(2,4) by fastforce ultimately show "E = E'" using res(5,6) res'(5,6) \ by blast qed lemma ord_resolve_rename_unique: assumes "ord_resolve_rename CAs DA AAs As \ E" and "ord_resolve_rename CAs DA AAs As \' E'" shows "\ = \' \ E = E'" using assms unfolding ord_resolve_rename.simps using ord_resolve_unique by meson lemma ord_resolve_max_side_prems: "ord_resolve CAs DA AAs As \ E \ length CAs \ size DA" by (auto elim!: ord_resolve.cases) lemma ord_resolve_rename_max_side_prems: "ord_resolve_rename CAs DA AAs As \ E \ length CAs \ size DA" by (elim ord_resolve_rename.cases, drule ord_resolve_max_side_prems, simp add: renamings_apart_length) subsection \Inference System\ definition ord_FO_\ :: "'a inference set" where "ord_FO_\ = {Infer (mset CAs) DA E | CAs DA AAs As \ E. ord_resolve_rename CAs DA AAs As \ E}" interpretation ord_FO_resolution: inference_system ord_FO_\ . lemma exists_compose: "\x. P (f x) \ \y. P y" by meson lemma finite_ord_FO_resolution_inferences_between: assumes fin_cc: "finite CC" shows "finite (ord_FO_resolution.inferences_between CC C)" proof - let ?CCC = "CC \ {C}" define all_AA where "all_AA = (\D \ ?CCC. atms_of D)" define max_ary where "max_ary = Max (size ` ?CCC)" define CAS where "CAS = {CAs. CAs \ lists ?CCC \ length CAs \ max_ary}" define AS where "AS = {As. As \ lists all_AA \ length As \ max_ary}" define AAS where "AAS = {AAs. AAs \ lists (mset ` AS) \ length AAs \ max_ary}" note defs = all_AA_def max_ary_def CAS_def AS_def AAS_def let ?infer_of = "\CAs DA AAs As. Infer (mset CAs) DA (THE E. \\. ord_resolve_rename CAs DA AAs As \ E)" let ?Z = "{\ | CAs DA AAs As \ E \. \ = Infer (mset CAs) DA E \ ord_resolve_rename CAs DA AAs As \ E \ infer_from ?CCC \ \ C \# prems_of \}" let ?Y = "{Infer (mset CAs) DA E | CAs DA AAs As \ E. ord_resolve_rename CAs DA AAs As \ E \ set CAs \ {DA} \ ?CCC}" let ?X = "{?infer_of CAs DA AAs As | CAs DA AAs As. CAs \ CAS \ DA \ ?CCC \ AAs \ AAS \ As \ AS}" let ?W = "CAS \ ?CCC \ AAS \ AS" have fin_w: "finite ?W" unfolding defs using fin_cc by (simp add: finite_lists_length_le lists_eq_set) have "?Z \ ?Y" by (force simp: infer_from_def) also have "\ \ ?X" proof - { fix CAs DA AAs As \ E assume res_e: "ord_resolve_rename CAs DA AAs As \ E" and da_in: "DA \ ?CCC" and cas_sub: "set CAs \ ?CCC" have "E = (THE E. \\. ord_resolve_rename CAs DA AAs As \ E) \ CAs \ CAS \ AAs \ AAS \ As \ AS" (is "?e \ ?cas \ ?aas \ ?as") proof (intro conjI) show ?e using res_e ord_resolve_rename_unique by (blast intro: the_equality[symmetric]) next show ?cas unfolding CAS_def max_ary_def using cas_sub ord_resolve_rename_max_side_prems[OF res_e] da_in fin_cc by (auto simp add: Max_ge_iff) next show ?aas using res_e proof (cases rule: ord_resolve_rename.cases) case (ord_resolve_rename n \ \s) note len_cas = this(1) and len_aas = this(2) and len_as = this(3) and aas_sub = this(4) and as_sub = this(5) and res_e' = this(8) show ?thesis unfolding AAS_def proof (clarify, intro conjI) show "AAs \ lists (mset ` AS)" unfolding AS_def image_def proof clarsimp fix AA assume "AA \ set AAs" then obtain i where i_lt: "i < n" and aa: "AA = AAs ! i" by (metis in_set_conv_nth len_aas) have casi_in: "CAs ! i \ ?CCC" using i_lt len_cas cas_sub nth_mem by blast have pos_aa_sub: "poss AA \# CAs ! i" using aa aas_sub i_lt by blast then have "set_mset AA \ atms_of (CAs ! i)" by (metis atms_of_poss lits_subseteq_imp_atms_subseteq set_mset_mono) also have aa_sub: "\ \ all_AA" unfolding all_AA_def using casi_in by force finally have aa_sub: "set_mset AA \ all_AA" . have "size AA = size (poss AA)" by simp also have "\ \ size (CAs ! i)" by (rule size_mset_mono[OF pos_aa_sub]) also have "\ \ max_ary" unfolding max_ary_def using fin_cc casi_in by auto finally have sz_aa: "size AA \ max_ary" . let ?As' = "sorted_list_of_multiset AA" have "?As' \ lists all_AA" using aa_sub by auto moreover have "length ?As' \ max_ary" using sz_aa by simp moreover have "AA = mset ?As'" by simp ultimately show "\xa. xa \ lists all_AA \ length xa \ max_ary \ AA = mset xa" by blast qed next have "length AAs = length As" unfolding len_aas len_as .. also have "\ \ size DA" using as_sub size_mset_mono by fastforce also have "\ \ max_ary" unfolding max_ary_def using fin_cc da_in by auto finally show "length AAs \ max_ary" . qed qed next show ?as unfolding AS_def proof (clarify, intro conjI) have "set As \ atms_of DA" using res_e[simplified ord_resolve_rename.simps] by (metis atms_of_negs lits_subseteq_imp_atms_subseteq set_mset_mono set_mset_mset) also have "\ \ all_AA" unfolding all_AA_def using da_in by blast finally show "As \ lists all_AA" unfolding lists_eq_set by simp next have "length As \ size DA" using res_e[simplified ord_resolve_rename.simps] ord_resolve_rename_max_side_prems[OF res_e] by auto also have "size DA \ max_ary" unfolding max_ary_def using fin_cc da_in by auto finally show "length As \ max_ary" . qed qed } then show ?thesis by simp fast qed also have "\ \ (\(CAs, DA, AAs, As). ?infer_of CAs DA AAs As) ` ?W" unfolding image_def Bex_cartesian_product by fast finally show ?thesis unfolding inference_system.inferences_between_def ord_FO_\_def mem_Collect_eq by (fast intro: rev_finite_subset[OF finite_imageI[OF fin_w]]) qed lemma ord_FO_resolution_inferences_between_empty_empty: "ord_FO_resolution.inferences_between {} {#} = {}" unfolding ord_FO_resolution.inferences_between_def inference_system.inferences_between_def infer_from_def ord_FO_\_def using ord_resolve_rename_empty_main_prem by auto subsection \Lifting\ text \ The following corresponds to the passage between Lemmas 4.11 and 4.12. \ context fixes M :: "'a clause set" assumes select: "selection S" begin interpretation selection by (rule select) definition S_M :: "'a literal multiset \ 'a literal multiset" where "S_M C = (if C \ grounding_of_clss M then (SOME C'. \D \. D \ M \ C = D \ \ \ C' = S D \ \ \ is_ground_subst \) else S C)" lemma S_M_grounding_of_clss: assumes "C \ grounding_of_clss M" obtains D \ where "D \ M \ C = D \ \ \ S_M C = S D \ \ \ is_ground_subst \" proof (atomize_elim, unfold S_M_def eqTrueI[OF assms] if_True, rule someI_ex) from assms show "\C' D \. D \ M \ C = D \ \ \ C' = S D \ \ \ is_ground_subst \" by (auto simp: grounding_of_clss_def grounding_of_cls_def) qed lemma S_M_not_grounding_of_clss: "C \ grounding_of_clss M \ S_M C = S C" unfolding S_M_def by simp lemma S_M_selects_subseteq: "S_M C \# C" by (metis S_M_grounding_of_clss S_M_not_grounding_of_clss S_selects_subseteq subst_cls_mono_mset) lemma S_M_selects_neg_lits: "L \# S_M C \ is_neg L" by (metis Melem_subst_cls S_M_grounding_of_clss S_M_not_grounding_of_clss S_selects_neg_lits subst_lit_is_neg) end end text \ The following corresponds to Lemma 4.12: \ lemma map2_add_mset_map: assumes "length AAs' = n" and "length As' = n" shows "map2 add_mset (As' \al \) (AAs' \aml \) = map2 add_mset As' AAs' \aml \" using assms proof (induction n arbitrary: AAs' As') case (Suc n) then have "map2 add_mset (tl (As' \al \)) (tl (AAs' \aml \)) = map2 add_mset (tl As') (tl AAs') \aml \" by simp - moreover - have Succ: "length (As' \al \) = Suc n" "length (AAs' \aml \) = Suc n" + moreover have Succ: "length (As' \al \) = Suc n" "length (AAs' \aml \) = Suc n" using Suc(3) Suc(2) by auto then have "length (tl (As' \al \)) = n" "length (tl (AAs' \aml \)) = n" by auto then have "length (map2 add_mset (tl (As' \al \)) (tl (AAs' \aml \))) = n" "length (map2 add_mset (tl As') (tl AAs') \aml \) = n" using Suc(2,3) by auto ultimately have "\i < n. tl (map2 add_mset ( (As' \al \)) ((AAs' \aml \))) ! i = tl (map2 add_mset (As') (AAs') \aml \) ! i" using Suc(2,3) Succ by (simp add: map2_tl map_tl subst_atm_mset_list_def del: subst_atm_list_tl) moreover have nn: "length (map2 add_mset ((As' \al \)) ((AAs' \aml \))) = Suc n" "length (map2 add_mset (As') (AAs') \aml \) = Suc n" using Succ Suc by auto ultimately have "\i. i < Suc n \ i > 0 \ map2 add_mset (As' \al \) (AAs' \aml \) ! i = (map2 add_mset As' AAs' \aml \) ! i" by (auto simp: subst_atm_mset_list_def gr0_conv_Suc subst_atm_mset_def) moreover have "add_mset (hd As' \a \) (hd AAs' \am \) = add_mset (hd As') (hd AAs') \am \" unfolding subst_atm_mset_def by auto then have "(map2 add_mset (As' \al \) (AAs' \aml \)) ! 0 = (map2 add_mset (As') (AAs') \aml \) ! 0" using Suc by (simp add: Succ(2) subst_atm_mset_def) ultimately have "\i < Suc n. (map2 add_mset (As' \al \) (AAs' \aml \)) ! i = (map2 add_mset (As') (AAs') \aml \) ! i" using Suc by auto then show ?case using nn list_eq_iff_nth_eq by metis qed auto lemma maximal_wrt_subst: "maximal_wrt (A \a \) (C \ \) \ maximal_wrt A C" unfolding maximal_wrt_def using in_atms_of_subst less_atm_stable by blast lemma strictly_maximal_wrt_subst: "strictly_maximal_wrt (A \a \) (C \ \) \ strictly_maximal_wrt A C" unfolding strictly_maximal_wrt_def using in_atms_of_subst less_atm_stable by blast lemma ground_resolvent_subset: assumes gr_cas: "is_ground_cls_list CAs" and gr_da: "is_ground_cls DA" and res_e: "ord_resolve S CAs DA AAs As \ E" shows "E \# \# (mset CAs) + DA" using res_e proof (cases rule: ord_resolve.cases) case (ord_resolve n Cs D) note da = this(1) and e = this(2) and cas_len = this(3) and cs_len = this(4) and aas_len = this(5) and as_len = this(6) and cas = this(8) and mgu = this(10) then have cs_sub_cas: "\# (mset Cs) \# \# (mset CAs)" using subseteq_list_Union_mset cas_len cs_len by force then have cs_sub_cas: "\# (mset Cs) \# \# (mset CAs)" using subseteq_list_Union_mset cas_len cs_len by force then have gr_cs: "is_ground_cls_list Cs" using gr_cas by simp have d_sub_da: "D \# DA" by (simp add: da) then have gr_d: "is_ground_cls D" using gr_da is_ground_cls_mono by auto have "is_ground_cls (\# (mset Cs) + D)" using gr_cs gr_d by auto with e have "E = \# (mset Cs) + D" by auto then show ?thesis using cs_sub_cas d_sub_da by (auto simp: subset_mset.add_mono) qed lemma ord_resolve_obtain_clauses: assumes res_e: "ord_resolve (S_M S M) CAs DA AAs As \ E" and select: "selection S" and grounding: "{DA} \ set CAs \ grounding_of_clss M" and n: "length CAs = n" and d: "DA = D + negs (mset As)" and c: "(\i < n. CAs ! i = Cs ! i + poss (AAs ! i))" "length Cs = n" "length AAs = n" obtains DA0 \0 CAs0 \s0 As0 AAs0 D0 Cs0 where "length CAs0 = n" "length \s0 = n" "DA0 \ M" "DA0 \ \0 = DA" "S DA0 \ \0 = S_M S M DA" "\CA0 \ set CAs0. CA0 \ M" "CAs0 \\cl \s0 = CAs" "map S CAs0 \\cl \s0 = map (S_M S M) CAs" "is_ground_subst \0" "is_ground_subst_list \s0" "As0 \al \0 = As" "AAs0 \\aml \s0 = AAs" "length As0 = n" "D0 \ \0 = D" "DA0 = D0 + (negs (mset As0))" "S_M S M (D + negs (mset As)) \ {#} \ negs (mset As0) = S DA0" "length Cs0 = n" "Cs0 \\cl \s0 = Cs" "\i < n. CAs0 ! i = Cs0 ! i + poss (AAs0 ! i)" "length AAs0 = n" using res_e proof (cases rule: ord_resolve.cases) case (ord_resolve n_twin Cs_twins D_twin) note da = this(1) and e = this(2) and cas = this(8) and mgu = this(10) and eligible = this(11) from ord_resolve have "n_twin = n" "D_twin = D" using n d by auto moreover have "Cs_twins = Cs" using c cas n calculation(1) \length Cs_twins = n_twin\ by (auto simp add: nth_equalityI) ultimately have nz: "n \ 0" and cs_len: "length Cs = n" and aas_len: "length AAs = n" and as_len: "length As = n" and da: "DA = D + negs (mset As)" and eligible: "eligible (S_M S M) \ As (D + negs (mset As))" and cas: "\in \ 0\ \length CAs = n\ \length Cs = n\ \length AAs = n\ \length As = n\ interpret S: selection S by (rule select) \ \Obtain FO side premises\ have "\CA \ set CAs. \CA0 \c0. CA0 \ M \ CA0 \ \c0 = CA \ S CA0 \ \c0 = S_M S M CA \ is_ground_subst \c0" using grounding S_M_grounding_of_clss select by (metis (no_types) le_supE subset_iff) then have "\i < n. \CA0 \c0. CA0 \ M \ CA0 \ \c0 = (CAs ! i) \ S CA0 \ \c0 = S_M S M (CAs ! i) \ is_ground_subst \c0" using n by force then obtain \s0f CAs0f where f_p: "\i < n. CAs0f i \ M" "\i < n. (CAs0f i) \ (\s0f i) = (CAs ! i)" "\i < n. S (CAs0f i) \ (\s0f i) = S_M S M (CAs ! i)" "\i < n. is_ground_subst (\s0f i)" using n by (metis (no_types)) define \s0 where "\s0 = map \s0f [0 ..s0 = n" "length CAs0 = n" unfolding \s0_def CAs0_def by auto note n = \length \s0 = n\ \length CAs0 = n\ n \ \The properties we need of the FO side premises\ have CAs0_in_M: "\CA0 \ set CAs0. CA0 \ M" unfolding CAs0_def using f_p(1) by auto have CAs0_to_CAs: "CAs0 \\cl \s0 = CAs" unfolding CAs0_def \s0_def using f_p(2) by (auto simp: n intro: nth_equalityI) have SCAs0_to_SMCAs: "(map S CAs0) \\cl \s0 = map (S_M S M) CAs" unfolding CAs0_def \s0_def using f_p(3) n by (force intro: nth_equalityI) have sub_ground: "\\c0 \ set \s0. is_ground_subst \c0" unfolding \s0_def using f_p n by force then have "is_ground_subst_list \s0" using n unfolding is_ground_subst_list_def by auto \ \Split side premises CAs0 into Cs0 and AAs0\ obtain AAs0 Cs0 where AAs0_Cs0_p: "AAs0 \\aml \s0 = AAs" "length Cs0 = n" "Cs0 \\cl \s0 = Cs" "\i < n. CAs0 ! i = Cs0 ! i + poss (AAs0 ! i)" "length AAs0 = n" proof - have "\i < n. \AA0. AA0 \am \s0 ! i = AAs ! i \ poss AA0 \# CAs0 ! i" proof (rule, rule) fix i assume "i < n" have "CAs0 ! i \ \s0 ! i = CAs ! i" using \i < n\ \CAs0 \\cl \s0 = CAs\ n by force moreover have "poss (AAs ! i) \# CAs !i" using \i < n\ cas by auto ultimately obtain poss_AA0 where nn: "poss_AA0 \ \s0 ! i = poss (AAs ! i) \ poss_AA0 \# CAs0 ! i" using cas image_mset_of_subset unfolding subst_cls_def by metis then have l: "\L \# poss_AA0. is_pos L" unfolding subst_cls_def by (metis Melem_subst_cls imageE literal.disc(1) literal.map_disc_iff set_image_mset subst_cls_def subst_lit_def) define AA0 where "AA0 = image_mset atm_of poss_AA0" have na: "poss AA0 = poss_AA0" using l unfolding AA0_def by auto then have "AA0 \am \s0 ! i = AAs ! i" using nn by (metis (mono_tags) literal.inject(1) multiset.inj_map_strong subst_cls_poss) moreover have "poss AA0 \# CAs0 ! i" using na nn by auto ultimately show "\AA0. AA0 \am \s0 ! i = AAs ! i \ poss AA0 \# CAs0 ! i" by blast qed then obtain AAs0f where AAs0f_p: "\i < n. AAs0f i \am \s0 ! i = AAs ! i \ (poss (AAs0f i)) \# CAs0 ! i" by metis define AAs0 where "AAs0 = map AAs0f [0 ..length AAs0 = n\ from AAs0_def have "\i < n. AAs0 ! i \am \s0 ! i = AAs ! i" using AAs0f_p by auto then have AAs0_AAs: "AAs0 \\aml \s0 = AAs" using n by (auto intro: nth_equalityI) from AAs0_def have AAs0_in_CAs0: "\i < n. poss (AAs0 ! i) \# CAs0 ! i" using AAs0f_p by auto define Cs0 where "Cs0 = map2 (-) CAs0 (map poss AAs0)" have "length Cs0 = n" using Cs0_def n by auto note n = n \length Cs0 = n\ have "\i < n. CAs0 ! i = Cs0 ! i + poss (AAs0 ! i)" using AAs0_in_CAs0 Cs0_def n by auto then have "Cs0 \\cl \s0 = Cs" using \CAs0 \\cl \s0 = CAs\ AAs0_AAs cas n by (auto intro: nth_equalityI) show ?thesis using that \AAs0 \\aml \s0 = AAs\ \Cs0 \\cl \s0 = Cs\ \\i < n. CAs0 ! i = Cs0 ! i + poss (AAs0 ! i)\ \length AAs0 = n\ \length Cs0 = n\ by blast qed \ \Obtain FO main premise\ have "\DA0 \0. DA0 \ M \ DA = DA0 \ \0 \ S DA0 \ \0 = S_M S M DA \ is_ground_subst \0" using grounding S_M_grounding_of_clss select by (metis le_supE singletonI subsetCE) then obtain DA0 \0 where DA0_\0_p: "DA0 \ M \ DA = DA0 \ \0 \ S DA0 \ \0 = S_M S M DA \ is_ground_subst \0" by auto \ \The properties we need of the FO main premise\ have DA0_in_M: "DA0 \ M" using DA0_\0_p by auto have DA0_to_DA: "DA0 \ \0 = DA" using DA0_\0_p by auto have SDA0_to_SMDA: "S DA0 \ \0 = S_M S M DA" using DA0_\0_p by auto have "is_ground_subst \0" using DA0_\0_p by auto \ \Split main premise DA0 into D0 and As0\ obtain D0 As0 where D0As0_p: "As0 \al \0 = As" "length As0 = n" "D0 \ \0 = D" "DA0 = D0 + (negs (mset As0))" "S_M S M (D + negs (mset As)) \ {#} \ negs (mset As0) = S DA0" proof - { assume a: "S_M S M (D + negs (mset As)) = {#} \ length As = (Suc 0) \ maximal_wrt (As ! 0 \a \) ((D + negs (mset As)) \ \)" then have as: "mset As = {#As ! 0#}" by (auto intro: nth_equalityI) then have "negs (mset As) = {#Neg (As ! 0)#}" by (simp add: \mset As = {#As ! 0#}\) then have "DA = D + {#Neg (As ! 0)#}" using da by auto then obtain L where "L \# DA0 \ L \l \0 = Neg (As ! 0)" using DA0_to_DA by (metis Melem_subst_cls mset_subset_eq_add_right single_subset_iff) then have "Neg (atm_of L) \# DA0 \ Neg (atm_of L) \l \0 = Neg (As ! 0)" by (metis Neg_atm_of_iff literal.sel(2) subst_lit_is_pos) then have "[atm_of L] \al \0 = As \ negs (mset [atm_of L]) \# DA0" using as subst_lit_def by auto then have "\As0. As0 \al \0 = As \ negs (mset As0) \# DA0 \ (S_M S M (D + negs (mset As)) \ {#} \ negs (mset As0) = S DA0)" using a by blast } moreover { assume "S_M S M (D + negs (mset As)) = negs (mset As)" then have "negs (mset As) = S DA0 \ \0" using da \S DA0 \ \0 = S_M S M DA\ by auto then have "\As0. negs (mset As0) = S DA0 \ As0 \al \0 = As" using instance_list[of As "S DA0" \0] S.S_selects_neg_lits by auto then have "\As0. As0 \al \0 = As \ negs (mset As0) \# DA0 \ (S_M S M (D + negs (mset As)) \ {#} \ negs (mset As0) = S DA0)" using S.S_selects_subseteq by auto } ultimately have "\As0. As0 \al \0 = As \ (negs (mset As0)) \# DA0 \ (S_M S M (D + negs (mset As)) \ {#} \ negs (mset As0) = S DA0)" using eligible unfolding eligible.simps by auto then obtain As0 where As0_p: "As0 \al \0 = As \ negs (mset As0) \# DA0 \ (S_M S M (D + negs (mset As)) \ {#} \ negs (mset As0) = S DA0)" by blast then have "length As0 = n" using as_len by auto note n = n this have "As0 \al \0 = As" using As0_p by auto define D0 where "D0 = DA0 - negs (mset As0)" then have "DA0 = D0 + negs (mset As0)" using As0_p by auto then have "D0 \ \0 = D" using DA0_to_DA da As0_p by auto have "S_M S M (D + negs (mset As)) \ {#} \ negs (mset As0) = S DA0" using As0_p by blast then show ?thesis using that \As0 \al \0 = As\ \D0 \ \0= D\ \DA0 = D0 + (negs (mset As0))\ \length As0 = n\ by metis qed show ?thesis using that[OF n(2,1) DA0_in_M DA0_to_DA SDA0_to_SMDA CAs0_in_M CAs0_to_CAs SCAs0_to_SMCAs \is_ground_subst \0\ \is_ground_subst_list \s0\ \As0 \al \0 = As\ \AAs0 \\aml \s0 = AAs\ \length As0 = n\ \D0 \ \0 = D\ \DA0 = D0 + (negs (mset As0))\ \S_M S M (D + negs (mset As)) \ {#} \ negs (mset As0) = S DA0\ \length Cs0 = n\ \Cs0 \\cl \s0 = Cs\ \\i < n. CAs0 ! i = Cs0 ! i + poss (AAs0 ! i)\ \length AAs0 = n\] by auto qed -lemma - assumes "Pos A \# C" - shows "A \ atms_of C" - using assms - by (simp add: atm_iff_pos_or_neg_lit) - lemma ord_resolve_rename_lifting: assumes sel_stable: "\\ C. is_renaming \ \ S (C \ \) = S C \ \" and res_e: "ord_resolve (S_M S M) CAs DA AAs As \ E" and select: "selection S" and grounding: "{DA} \ set CAs \ grounding_of_clss M" obtains \s \ \2 CAs0 DA0 AAs0 As0 E0 \ where "is_ground_subst \" "is_ground_subst_list \s" "is_ground_subst \2" "ord_resolve_rename S CAs0 DA0 AAs0 As0 \ E0" "CAs0 \\cl \s = CAs" "DA0 \ \ = DA" "E0 \ \2 = E" "{DA0} \ set CAs0 \ M" "length CAs0 = length CAs" using res_e proof (cases rule: ord_resolve.cases) case (ord_resolve n Cs D) note da = this(1) and e = this(2) and cas_len = this(3) and cs_len = this(4) and aas_len = this(5) and as_len = this(6) and nz = this(7) and cas = this(8) and aas_not_empt = this(9) and mgu = this(10) and eligible = this(11) and str_max = this(12) and sel_empt = this(13) have sel_ren_list_inv: "\\s Cs. length \s = length Cs \ is_renaming_list \s \ map S (Cs \\cl \s) = map S Cs \\cl \s" using sel_stable unfolding is_renaming_list_def by (auto intro: nth_equalityI) note n = \n \ 0\ \length CAs = n\ \length Cs = n\ \length AAs = n\ \length As = n\ interpret S: selection S by (rule select) obtain DA0 \0 CAs0 \s0 As0 AAs0 D0 Cs0 where as0: "length CAs0 = n" "length \s0 = n" "DA0 \ M" "DA0 \ \0 = DA" "S DA0 \ \0 = S_M S M DA" "\CA0 \ set CAs0. CA0 \ M" "CAs0 \\cl \s0 = CAs" "map S CAs0 \\cl \s0 = map (S_M S M) CAs" "is_ground_subst \0" "is_ground_subst_list \s0" "As0 \al \0 = As" "AAs0 \\aml \s0 = AAs" "length As0 = n" "D0 \ \0 = D" "DA0 = D0 + (negs (mset As0))" "S_M S M (D + negs (mset As)) \ {#} \ negs (mset As0) = S DA0" "length Cs0 = n" "Cs0 \\cl \s0 = Cs" "\i < n. CAs0 ! i = Cs0 ! i + poss (AAs0 ! i)" "length AAs0 = n" using ord_resolve_obtain_clauses[of S M CAs DA, OF res_e select grounding n(2) \DA = D + negs (mset As)\ \\i \length Cs = n\ \length AAs = n\, of thesis] by blast note n = \length CAs0 = n\ \length \s0 = n\ \length As0 = n\ \length AAs0 = n\ \length Cs0 = n\ n have "length (renamings_apart (DA0 # CAs0)) = Suc n" using n renamings_apart_length by auto note n = this n define \ where "\ = hd (renamings_apart (DA0 # CAs0))" define \s where "\s = tl (renamings_apart (DA0 # CAs0))" define DA0' where "DA0' = DA0 \ \" define D0' where "D0' = D0 \ \" define As0' where "As0' = As0 \al \" define CAs0' where "CAs0' = CAs0 \\cl \s" define Cs0' where "Cs0' = Cs0 \\cl \s" define AAs0' where "AAs0' = AAs0 \\aml \s" define \0' where "\0' = inv_renaming \ \ \0" define \s0' where "\s0' = map inv_renaming \s \s \s0" have renames_DA0: "is_renaming \" using renamings_apart_length renamings_apart_renaming unfolding \_def by (metis length_greater_0_conv list.exhaust_sel list.set_intros(1) list.simps(3)) have renames_CAs0: "is_renaming_list \s" using renamings_apart_length renamings_apart_renaming unfolding \s_def by (metis is_renaming_list_def length_greater_0_conv list.set_sel(2) list.simps(3)) have "length \s = n" unfolding \s_def using n by auto note n = n \length \s = n\ have "length As0' = n" unfolding As0'_def using n by auto have "length CAs0' = n" using as0(1) n unfolding CAs0'_def by auto have "length Cs0' = n" unfolding Cs0'_def using n by auto have "length AAs0' = n" unfolding AAs0'_def using n by auto have "length \s0' = n" using as0(2) n unfolding \s0'_def by auto note n = \length CAs0' = n\ \length \s0' = n\ \length As0' = n\ \length AAs0' = n\ \length Cs0' = n\ n have DA0'_DA: "DA0' \ \0' = DA" using as0(4) unfolding \0'_def DA0'_def using renames_DA0 by simp have D0'_D: "D0' \ \0' = D" using as0(14) unfolding \0'_def D0'_def using renames_DA0 by simp have As0'_As: "As0' \al \0' = As" using as0(11) unfolding \0'_def As0'_def using renames_DA0 by auto have "S DA0' \ \0' = S_M S M DA" using as0(5) unfolding \0'_def DA0'_def using renames_DA0 sel_stable by auto have CAs0'_CAs: "CAs0' \\cl \s0' = CAs" using as0(7) unfolding CAs0'_def \s0'_def using renames_CAs0 n by auto have Cs0'_Cs: "Cs0' \\cl \s0' = Cs" using as0(18) unfolding Cs0'_def \s0'_def using renames_CAs0 n by auto have AAs0'_AAs: "AAs0' \\aml \s0' = AAs" using as0(12) unfolding \s0'_def AAs0'_def using renames_CAs0 using n by auto have "map S CAs0' \\cl \s0' = map (S_M S M) CAs" unfolding CAs0'_def \s0'_def using as0(8) n renames_CAs0 sel_ren_list_inv by auto have DA0'_split: "DA0' = D0' + negs (mset As0')" using as0(15) DA0'_def D0'_def As0'_def by auto then have D0'_subset_DA0': "D0' \# DA0'" by auto from DA0'_split have negs_As0'_subset_DA0': "negs (mset As0') \# DA0'" by auto have CAs0'_split: "\ii# CAs0' ! i" by auto from CAs0'_split have poss_AAs0'_subset_CAs0': "\i# CAs0' ! i" by auto then have AAs0'_in_atms_of_CAs0': "\i < n. \A\#AAs0' ! i. A \ atms_of (CAs0' ! i)" by (auto simp add: atm_iff_pos_or_neg_lit) have as0': "S_M S M (D + negs (mset As)) \ {#} \ negs (mset As0') = S DA0'" proof - assume a: "S_M S M (D + negs (mset As)) \ {#}" then have "negs (mset As0) \ \ = S DA0 \ \" using as0(16) unfolding \_def by metis then show "negs (mset As0') = S DA0'" using As0'_def DA0'_def using sel_stable[of \ DA0] renames_DA0 by auto qed have vd: "var_disjoint (DA0' # CAs0')" unfolding DA0'_def CAs0'_def using renamings_apart_var_disjoint unfolding \_def \s_def by (metis length_greater_0_conv list.exhaust_sel n(6) substitution.subst_cls_lists_Cons substitution_axioms zero_less_Suc) \ \Introduce ground substitution\ from vd DA0'_DA CAs0'_CAs have "\\. \i < Suc n. \S. S \# (DA0' # CAs0') ! i \ S \ (\0'#\s0') ! i = S \ \" unfolding var_disjoint_def using n by auto then obtain \ where \_p: "\i < Suc n. \S. S \# (DA0' # CAs0') ! i \ S \ (\0'#\s0') ! i = S \ \" by auto have \_p_lit: "\i < Suc n. \L. L \# (DA0' # CAs0') ! i \ L \l (\0'#\s0') ! i = L \l \" proof (rule, rule, rule, rule) fix i :: "nat" and L :: "'a literal" assume a: "i < Suc n" "L \# (DA0' # CAs0') ! i" then have "\S. S \# (DA0' # CAs0') ! i \ S \ (\0' # \s0') ! i = S \ \" using \_p by auto then have "{# L #} \ (\0' # \s0') ! i = {# L #} \ \" using a by (meson single_subset_iff) then show "L \l (\0' # \s0') ! i = L \l \" by auto qed have \_p_atm: "\i < Suc n. \A. A \ atms_of ((DA0' # CAs0') ! i) \ A \a (\0'#\s0') ! i = A \a \" proof (rule, rule, rule, rule) fix i :: "nat" and A :: "'a" assume a: "i < Suc n" "A \ atms_of ((DA0' # CAs0') ! i)" then obtain L where L_p: "atm_of L = A \ L \# (DA0' # CAs0') ! i" unfolding atms_of_def by auto then have "L \l (\0'#\s0') ! i = L \l \" using \_p_lit a by auto then show "A \a (\0' # \s0') ! i = A \a \" using L_p unfolding subst_lit_def by (cases L) auto qed have DA0'_DA: "DA0' \ \ = DA" using DA0'_DA \_p by auto have "D0' \ \ = D" using \_p D0'_D n D0'_subset_DA0' by auto have "As0' \al \ = As" proof (rule nth_equalityI) show "length (As0' \al \) = length As" using n by auto next fix i show "ial \) \ (As0' \al \) ! i = As ! i" proof - assume a: "i < length (As0' \al \)" have A_eq: "\A. A \ atms_of DA0' \ A \a \0' = A \a \" using \_p_atm n by force have "As0' ! i \ atms_of DA0'" using negs_As0'_subset_DA0' unfolding atms_of_def using a n by force then have "As0' ! i \a \0' = As0' ! i \a \" using A_eq by simp then show "(As0' \al \) ! i = As ! i" using As0'_As \length As0' = n\ a by auto qed qed interpret selection by (rule select) have "S DA0' \ \ = S_M S M DA" using \S DA0' \ \0' = S_M S M DA\ \_p S.S_selects_subseteq by auto from \_p have \_p_CAs0': "\i < n. (CAs0' ! i) \ (\s0' ! i) = (CAs0'! i) \ \" using n by auto then have "CAs0' \\cl \s0' = CAs0' \cl \" using n by (auto intro: nth_equalityI) then have CAs0'_\_fo_CAs: "CAs0' \cl \ = CAs" using CAs0'_CAs \_p n by auto from \_p have "\i < n. S (CAs0' ! i) \ \s0' ! i = S (CAs0' ! i) \ \" using S.S_selects_subseteq n by auto then have "map S CAs0' \\cl \s0' = map S CAs0' \cl \" using n by (auto intro: nth_equalityI) then have SCAs0'_\_fo_SMCAs: "map S CAs0' \cl \ = map (S_M S M) CAs" using \map S CAs0' \\cl \s0' = map (S_M S M) CAs\ by auto have "Cs0' \cl \ = Cs" proof (rule nth_equalityI) show "length (Cs0' \cl \) = length Cs" using n by auto next fix i show "icl \) \ (Cs0' \cl \) ! i = Cs ! i" proof - assume "i < length (Cs0' \cl \)" then have a: "i < n" using n by force have "(Cs0' \\cl \s0') ! i = Cs ! i" using Cs0'_Cs a n by force moreover have \_p_CAs0': "\S. S \# CAs0' ! i \ S \ \s0' ! i = S \ \" using \_p a by force have "Cs0' ! i \ \s0' ! i = (Cs0' \cl \) ! i" using \_p_CAs0' \\i# CAs0' ! i\ a n by force then have "(Cs0' \\cl \s0') ! i = (Cs0' \cl \) ! i " using a n by force ultimately show "(Cs0' \cl \) ! i = Cs ! i" by auto qed qed have AAs0'_AAs: "AAs0' \aml \ = AAs" proof (rule nth_equalityI) show "length (AAs0' \aml \) = length AAs" using n by auto next fix i show "iaml \) \ (AAs0' \aml \) ! i = AAs ! i" proof - assume a: "i < length (AAs0' \aml \)" then have "i < n" using n by force then have "\A. A \ atms_of ((DA0' # CAs0') ! Suc i) \ A \a (\0' # \s0') ! Suc i = A \a \" using \_p_atm n by force then have A_eq: "\A. A \ atms_of (CAs0' ! i) \ A \a \s0' ! i = A \a \" by auto have AAs_CAs0': "\A \# AAs0' ! i. A \ atms_of (CAs0' ! i)" using AAs0'_in_atms_of_CAs0' unfolding atms_of_def using a n by force then have "AAs0' ! i \am \s0' ! i = AAs0' ! i \am \" unfolding subst_atm_mset_def using A_eq unfolding subst_atm_mset_def by auto then show "(AAs0' \aml \) ! i = AAs ! i" using AAs0'_AAs \length AAs0' = n\ \length \s0' = n\ a by auto qed qed \ \Obtain MGU and substitution\ obtain \ \ where \\: "Some \ = mgu (set_mset ` set (map2 add_mset As0' AAs0'))" "\ \ \ = \ \ \" proof - have uu: "is_unifiers \ (set_mset ` set (map2 add_mset (As0' \al \) (AAs0' \aml \)))" using mgu mgu_sound is_mgu_def unfolding \AAs0' \aml \ = AAs\ using \As0' \al \ = As\ by auto have \\uni: "is_unifiers (\ \ \) (set_mset ` set (map2 add_mset As0' AAs0'))" proof - have "set_mset ` set (map2 add_mset As0' AAs0' \aml \) = set_mset ` set (map2 add_mset As0' AAs0') \ass \" unfolding subst_atmss_def subst_atm_mset_list_def using subst_atm_mset_def subst_atms_def by (simp add: image_image subst_atm_mset_def subst_atms_def) then have "is_unifiers \ (set_mset ` set (map2 add_mset As0' AAs0') \ass \)" using uu by (auto simp: n map2_add_mset_map) then show ?thesis using is_unifiers_comp by auto qed then obtain \ where \_p: "Some \ = mgu (set_mset ` set (map2 add_mset As0' AAs0'))" using mgu_complete by (metis (mono_tags, hide_lams) List.finite_set finite_imageI finite_set_mset image_iff) moreover then obtain \ where \_p: "\ \ \ = \ \ \" by (metis (mono_tags, hide_lams) finite_set \\uni finite_imageI finite_set_mset image_iff mgu_sound set_mset_mset substitution_ops.is_mgu_def) (* should be simpler *) ultimately show thesis using that by auto qed \ \Lifting eligibility\ have eligible0': "eligible S \ As0' (D0' + negs (mset As0'))" proof - have "S_M S M (D + negs (mset As)) = negs (mset As) \ S_M S M (D + negs (mset As)) = {#} \ length As = 1 \ maximal_wrt (As ! 0 \a \) ((D + negs (mset As)) \ \)" using eligible unfolding eligible.simps by auto then show ?thesis proof assume "S_M S M (D + negs (mset As)) = negs (mset As)" then have "S_M S M (D + negs (mset As)) \ {#}" using n by force then have "S (D0' + negs (mset As0')) = negs (mset As0')" using as0' DA0'_split by auto then show ?thesis unfolding eligible.simps[simplified] by auto next assume asm: "S_M S M (D + negs (mset As)) = {#} \ length As = 1 \ maximal_wrt (As ! 0 \a \) ((D + negs (mset As)) \ \)" then have "S (D0' + negs (mset As0')) = {#}" using \D0' \ \ = D\[symmetric] \As0' \al \ = As\[symmetric] \S (DA0') \ \ = S_M S M (DA)\ da DA0'_split subst_cls_empty_iff by metis moreover from asm have l: "length As0' = 1" using \As0' \al \ = As\ by auto moreover from asm have "maximal_wrt (As0' ! 0 \a (\ \ \)) ((D0' + negs (mset As0')) \ (\ \ \))" using \As0' \al \ = As\ \D0' \ \ = D\ using l \\ by auto then have "maximal_wrt (As0' ! 0 \a \ \a \) ((D0' + negs (mset As0')) \ \ \ \)" by auto then have "maximal_wrt (As0' ! 0 \a \) ((D0' + negs (mset As0')) \ \)" using maximal_wrt_subst by blast ultimately show ?thesis unfolding eligible.simps[simplified] by auto qed qed \ \Lifting maximality\ have maximality: "\i < n. strictly_maximal_wrt (As0' ! i \a \) (Cs0' ! i \ \)" (* Reformulate in list notation? *) proof - from str_max have "\i < n. strictly_maximal_wrt ((As0' \al \) ! i \a \) ((Cs0' \cl \) ! i \ \)" using \As0' \al \ = As\ \Cs0' \cl \ = Cs\ by simp then have "\i < n. strictly_maximal_wrt (As0' ! i \a (\ \ \)) (Cs0' ! i \ (\ \ \))" using n \\ by simp then have "\i < n. strictly_maximal_wrt (As0' ! i \a \ \a \) (Cs0' ! i \ \ \ \)" by auto then show "\i < n. strictly_maximal_wrt (As0' ! i \a \) (Cs0' ! i \ \)" using strictly_maximal_wrt_subst \\ by blast qed \ \Lifting nothing being selected\ have nothing_selected: "\i < n. S (CAs0' ! i) = {#}" proof - have "\i < n. (map S CAs0' \cl \) ! i = map (S_M S M) CAs ! i" by (simp add: \map S CAs0' \cl \ = map (S_M S M) CAs\) then have "\i < n. S (CAs0' ! i) \ \ = S_M S M (CAs ! i)" using n by auto then have "\i < n. S (CAs0' ! i) \ \ = {#}" using sel_empt \\i < n. S (CAs0' ! i) \ \ = S_M S M (CAs ! i)\ by auto then show "\i < n. S (CAs0' ! i) = {#}" using subst_cls_empty_iff by blast qed \ \Lifting AAs0's non-emptiness\ have "\i < n. AAs0' ! i \ {#}" using n aas_not_empt \AAs0' \aml \ = AAs\ by auto \ \Resolve the lifted clauses\ define E0' where "E0' = ((\# (mset Cs0')) + D0') \ \" have res_e0': "ord_resolve S CAs0' DA0' AAs0' As0' \ E0'" using ord_resolve.intros[of CAs0' n Cs0' AAs0' As0' \ S D0', OF _ _ _ _ _ _ \\i < n. AAs0' ! i \ {#}\ \\(1) eligible0' \\i < n. strictly_maximal_wrt (As0' ! i \a \) (Cs0' ! i \ \)\ \\i < n. S (CAs0' ! i) = {#}\] unfolding E0'_def using DA0'_split n \\i by blast \ \Prove resolvent instantiates to ground resolvent\ have e0'\e: "E0' \ \ = E" proof - have "E0' \ \ = ((\# (mset Cs0')) + D0') \ (\ \ \)" unfolding E0'_def by auto also have "\ = (\# (mset Cs0') + D0') \ (\ \ \)" using \\ by auto also have "\ = (\# (mset Cs) + D) \ \" using \Cs0' \cl \ = Cs\ \D0' \ \ = D\ by auto also have "\ = E" using e by auto finally show e0'\e: "E0' \ \ = E" . qed \ \Replace @{term \} with a true ground substitution\ obtain \2 where ground_\2: "is_ground_subst \2" "E0' \ \2 = E" proof - have "is_ground_cls_list CAs" "is_ground_cls DA" using grounding grounding_ground unfolding is_ground_cls_list_def by auto then have "is_ground_cls E" using res_e ground_resolvent_subset by (force intro: is_ground_cls_mono) then show thesis using that e0'\e make_ground_subst by auto qed have \length CAs0 = length CAs\ using n by simp \ \Wrap up the proof\ have "ord_resolve S (CAs0 \\cl \s) (DA0 \ \) (AAs0 \\aml \s) (As0 \al \) \ E0'" using res_e0' As0'_def \_def AAs0'_def \s_def DA0'_def \_def CAs0'_def \s_def by simp moreover have "\i# CAs0 ! i" using as0(19) by auto moreover have "negs (mset As0) \# DA0" using local.as0(15) by auto ultimately have "ord_resolve_rename S CAs0 DA0 AAs0 As0 \ E0'" using ord_resolve_rename[of CAs0 n AAs0 As0 DA0 \ \s S \ E0'] \_def \s_def n by auto then show thesis using that[of \0 \s0 \2 CAs0 DA0] \is_ground_subst \0\ \is_ground_subst_list \s0\ \is_ground_subst \2\ \CAs0 \\cl \s0 = CAs\ \DA0 \ \0 = DA\ \E0' \ \2 = E\ \DA0 \ M\ \\CA \ set CAs0. CA \ M\ \length CAs0 = length CAs\ by blast qed end end diff --git a/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution_Prover.thy b/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution_Prover.thy --- a/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution_Prover.thy +++ b/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution_Prover.thy @@ -1,1544 +1,1544 @@ (* Title: An Ordered Resolution Prover for First-Order Clauses Author: Anders Schlichtkrull , 2016, 2017 Author: Jasmin Blanchette , 2014, 2017 Author: Dmitriy Traytel , 2014 Maintainer: Anders Schlichtkrull *) section \An Ordered Resolution Prover for First-Order Clauses\ theory FO_Ordered_Resolution_Prover imports FO_Ordered_Resolution begin text \ This material is based on Section 4.3 (``A Simple Resolution Prover for First-Order Clauses'') of Bachmair and Ganzinger's chapter. Specifically, it formalizes the RP prover defined in Figure 5 and its related lemmas and theorems, including Lemmas 4.10 and 4.11 and Theorem 4.13 (completeness). \ definition is_least :: "(nat \ bool) \ nat \ bool" where "is_least P n \ P n \ (\n' < n. \ P n')" lemma least_exists: "P n \ \n. is_least P n" using exists_least_iff unfolding is_least_def by auto text \ The following corresponds to page 42 and 43 of Section 4.3, from the explanation of RP to Lemma 4.10. \ type_synonym 'a state = "'a clause set \ 'a clause set \ 'a clause set" locale FO_resolution_prover = FO_resolution subst_atm id_subst comp_subst renamings_apart atm_of_atms mgu less_atm + selection S for S :: "('a :: wellorder) clause \ 'a clause" and subst_atm :: "'a \ 's \ 'a" and id_subst :: "'s" and comp_subst :: "'s \ 's \ 's" and renamings_apart :: "'a clause list \ 's list" and atm_of_atms :: "'a list \ 'a" and mgu :: "'a set set \ 's option" and less_atm :: "'a \ 'a \ bool" + assumes sel_stable: "\\ C. is_renaming \ \ S (C \ \) = S C \ \" and less_atm_ground: "is_ground_atm A \ is_ground_atm B \ less_atm A B \ A < B" begin fun N_of_state :: "'a state \ 'a clause set" where "N_of_state (N, P, Q) = N" fun P_of_state :: "'a state \ 'a clause set" where "P_of_state (N, P, Q) = P" text \ \O\ denotes relation composition in Isabelle, so the formalization uses \Q\ instead. \ fun Q_of_state :: "'a state \ 'a clause set" where "Q_of_state (N, P, Q) = Q" -definition clss_of_state :: "'a state \ 'a clause set" where - "clss_of_state St = N_of_state St \ P_of_state St \ Q_of_state St" +abbreviation clss_of_state :: "'a state \ 'a clause set" where + "clss_of_state St \ N_of_state St \ P_of_state St \ Q_of_state St" abbreviation grounding_of_state :: "'a state \ 'a clause set" where "grounding_of_state St \ grounding_of_clss (clss_of_state St)" interpretation ord_FO_resolution: inference_system "ord_FO_\ S" . text \ The following inductive predicate formalizes the resolution prover in Figure 5. \ inductive RP :: "'a state \ 'a state \ bool" (infix "\" 50) where tautology_deletion: "Neg A \# C \ Pos A \# C \ (N \ {C}, P, Q) \ (N, P, Q)" | forward_subsumption: "D \ P \ Q \ subsumes D C \ (N \ {C}, P, Q) \ (N, P, Q)" | backward_subsumption_P: "D \ N \ strictly_subsumes D C \ (N, P \ {C}, Q) \ (N, P, Q)" | backward_subsumption_Q: "D \ N \ strictly_subsumes D C \ (N, P, Q \ {C}) \ (N, P, Q)" | forward_reduction: "D + {#L'#} \ P \ Q \ - L = L' \l \ \ D \ \ \# C \ (N \ {C + {#L#}}, P, Q) \ (N \ {C}, P, Q)" | backward_reduction_P: "D + {#L'#} \ N \ - L = L' \l \ \ D \ \ \# C \ (N, P \ {C + {#L#}}, Q) \ (N, P \ {C}, Q)" | backward_reduction_Q: "D + {#L'#} \ N \ - L = L' \l \ \ D \ \ \# C \ (N, P, Q \ {C + {#L#}}) \ (N, P \ {C}, Q)" | clause_processing: "(N \ {C}, P, Q) \ (N, P \ {C}, Q)" | inference_computation: "N = concls_of (ord_FO_resolution.inferences_between Q C) \ ({}, P \ {C}, Q) \ (N, P, Q \ {C})" lemma final_RP: "\ ({}, {}, Q) \ St" by (auto elim: RP.cases) definition Sup_state :: "'a state llist \ 'a state" where "Sup_state Sts = (Sup_llist (lmap N_of_state Sts), Sup_llist (lmap P_of_state Sts), Sup_llist (lmap Q_of_state Sts))" definition Liminf_state :: "'a state llist \ 'a state" where "Liminf_state Sts = (Liminf_llist (lmap N_of_state Sts), Liminf_llist (lmap P_of_state Sts), Liminf_llist (lmap Q_of_state Sts))" context fixes Sts Sts' :: "'a state llist" assumes Sts: "lfinite Sts" "lfinite Sts'" "\ lnull Sts" "\ lnull Sts'" "llast Sts' = llast Sts" begin lemma N_of_Liminf_state_fin: "N_of_state (Liminf_state Sts') = N_of_state (Liminf_state Sts)" and P_of_Liminf_state_fin: "P_of_state (Liminf_state Sts') = P_of_state (Liminf_state Sts)" and Q_of_Liminf_state_fin: "Q_of_state (Liminf_state Sts') = Q_of_state (Liminf_state Sts)" using Sts by (simp_all add: Liminf_state_def lfinite_Liminf_llist llast_lmap) lemma Liminf_state_fin: "Liminf_state Sts' = Liminf_state Sts" using N_of_Liminf_state_fin P_of_Liminf_state_fin Q_of_Liminf_state_fin by (simp add: Liminf_state_def) end context fixes Sts Sts' :: "'a state llist" assumes Sts: "\ lfinite Sts" "emb Sts Sts'" begin lemma N_of_Liminf_state_inf: "N_of_state (Liminf_state Sts') \ N_of_state (Liminf_state Sts)" and P_of_Liminf_state_inf: "P_of_state (Liminf_state Sts') \ P_of_state (Liminf_state Sts)" and Q_of_Liminf_state_inf: "Q_of_state (Liminf_state Sts') \ Q_of_state (Liminf_state Sts)" using Sts by (simp_all add: Liminf_state_def emb_Liminf_llist_infinite emb_lmap) lemma clss_of_Liminf_state_inf: "clss_of_state (Liminf_state Sts') \ clss_of_state (Liminf_state Sts)" - unfolding clss_of_state_def using N_of_Liminf_state_inf P_of_Liminf_state_inf Q_of_Liminf_state_inf by blast end definition fair_state_seq :: "'a state llist \ bool" where "fair_state_seq Sts \ N_of_state (Liminf_state Sts) = {} \ P_of_state (Liminf_state Sts) = {}" text \ The following formalizes Lemma 4.10. \ context - fixes - Sts :: "'a state llist" - assumes - deriv: "chain (\) Sts" and - empty_Q0: "Q_of_state (lhd Sts) = {}" - (* FIXME: make hypothesis empty_Q0 more local---and is it really needed? *) + fixes Sts :: "'a state llist" + assumes deriv: "chain (\) Sts" begin lemmas lhd_lmap_Sts = llist.map_sel(1)[OF chain_not_lnull[OF deriv]] definition S_Q :: "'a clause \ 'a clause" where "S_Q = S_M S (Q_of_state (Liminf_state Sts))" interpretation sq: selection S_Q unfolding S_Q_def using S_M_selects_subseteq S_M_selects_neg_lits selection_axioms by unfold_locales auto interpretation gr: ground_resolution_with_selection S_Q by unfold_locales interpretation sr: standard_redundancy_criterion_reductive gr.ord_\ by unfold_locales interpretation sr: standard_redundancy_criterion_counterex_reducing gr.ord_\ "ground_resolution_with_selection.INTERP S_Q" by unfold_locales text \ The extension of ordered resolution mentioned in 4.10. We let it consist of all sound rules. \ definition ground_sound_\:: "'a inference set" where "ground_sound_\ = {Infer CC D E | CC D E. (\I. I \m CC \ I \ D \ I \ E)}" text \ We prove that we indeed defined an extension. \ lemma gd_ord_\_ngd_ord_\: "gr.ord_\ \ ground_sound_\" unfolding ground_sound_\_def using gr.ord_\_def gr.ord_resolve_sound by fastforce lemma sound_ground_sound_\: "sound_inference_system ground_sound_\" unfolding sound_inference_system_def ground_sound_\_def by auto lemma sat_preserving_ground_sound_\: "sat_preserving_inference_system ground_sound_\" using sound_ground_sound_\ sat_preserving_inference_system.intro sound_inference_system.\_sat_preserving by blast definition sr_ext_Ri :: "'a clause set \ 'a inference set" where "sr_ext_Ri N = sr.Ri N \ (ground_sound_\ - gr.ord_\)" interpretation sr_ext: sat_preserving_redundancy_criterion "ground_sound_\" "sr.Rf" "sr_ext_Ri" unfolding sat_preserving_redundancy_criterion_def sr_ext_Ri_def using sat_preserving_ground_sound_\ redundancy_criterion_standard_extension gd_ord_\_ngd_ord_\ sr.redundancy_criterion_axioms by auto lemma strict_subset_subsumption_redundant_clause: assumes sub: "D \ \ \# C" and ground_\: "is_ground_subst \" shows "C \ sr.Rf (grounding_of_cls D)" proof - from sub have "\I. I \ D \ \ \ I \ C" unfolding true_cls_def by blast moreover have "C > D \ \" using sub by (simp add: subset_imp_less_mset) moreover have "D \ \ \ grounding_of_cls D" - using ground_\ by (metis (mono_tags, lifting) mem_Collect_eq substitution_ops.grounding_of_cls_def) + using ground_\ by (metis (mono_tags) mem_Collect_eq substitution_ops.grounding_of_cls_def) ultimately have "set_mset {#D \ \#} \ grounding_of_cls D" "(\I. I \m {#D \ \#} \ I \ C)" "(\D'. D' \# {#D \ \#} \ D' < C)" by auto then show ?thesis using sr.Rf_def by blast qed lemma strict_subset_subsumption_redundant_clss: assumes "D \ \ \# C" and "is_ground_subst \" and "D \ CC" shows "C \ sr.Rf (grounding_of_clss CC)" using assms proof - have "C \ sr.Rf (grounding_of_cls D)" using strict_subset_subsumption_redundant_clause assms by auto then show ?thesis - using assms unfolding clss_of_state_def grounding_of_clss_def + using assms unfolding grounding_of_clss_def by (metis (no_types) sr.Rf_mono sup_ge1 SUP_absorb contra_subsetD) qed lemma strict_subset_subsumption_grounding_redundant_clss: assumes D\_subset_C: "D \ \ \# C" and D_in_St: "D \ CC" shows "grounding_of_cls C \ sr.Rf (grounding_of_clss CC)" proof fix C\ assume "C\ \ grounding_of_cls C" then obtain \ where \_p: "C\ = C \ \ \ is_ground_subst \" unfolding grounding_of_cls_def by auto have D\\C\: "D \ \ \ \ \# C \ \" using D\_subset_C subst_subset_mono by auto then show "C\ \ sr.Rf (grounding_of_clss CC)" - using \_p strict_subset_subsumption_redundant_clss[of D "\ \ \" "C \ \"] D_in_St - unfolding clss_of_state_def by auto + using \_p strict_subset_subsumption_redundant_clss[of D "\ \ \" "C \ \"] D_in_St by auto qed lemma subst_cls_eq_grounding_of_cls_subset_eq: assumes "D \ \ = C" shows "grounding_of_cls C \ grounding_of_cls D" proof fix C\' assume "C\' \ grounding_of_cls C" then obtain \' where C\': "C \ \' = C\'" "is_ground_subst \'" unfolding grounding_of_cls_def by auto then have "C \ \' = D \ \ \ \' \ is_ground_subst (\ \ \')" using assms by auto then show "C\' \ grounding_of_cls D" unfolding grounding_of_cls_def using C\'(1) by force qed lemma derive_if_remove_subsumed: assumes "D \ clss_of_state St" and "subsumes D C" shows "sr_ext.derive (grounding_of_state St \ grounding_of_cls C) (grounding_of_state St)" proof - from assms obtain \ where "D \ \ = C \ D \ \ \# C" by (auto simp: subsumes_def subset_mset_def) then have "D \ \ = C \ D \ \ \# C" by (simp add: subset_mset_def) then show ?thesis proof assume "D \ \ = C" then have "grounding_of_cls C \ grounding_of_cls D" using subst_cls_eq_grounding_of_cls_subset_eq by simp then have "(grounding_of_state St \ grounding_of_cls C) = grounding_of_state St" - using assms unfolding clss_of_state_def grounding_of_clss_def by auto + using assms unfolding grounding_of_clss_def by auto then show ?thesis by (auto intro: sr_ext.derive.intros) next assume a: "D \ \ \# C" then have "grounding_of_cls C \ sr.Rf (grounding_of_state St)" using strict_subset_subsumption_grounding_redundant_clss assms by auto then show ?thesis - unfolding clss_of_state_def grounding_of_clss_def by (force intro: sr_ext.derive.intros) + unfolding grounding_of_clss_def by (force intro: sr_ext.derive.intros) qed qed -(* FIXME: better name *) lemma reduction_in_concls_of: assumes "C\ \ grounding_of_cls C" and "D + {#L'#} \ CC" and "- L = L' \l \" and "D \ \ \# C" shows "C\ \ concls_of (sr_ext.inferences_from (grounding_of_clss (CC \ {C + {#L#}})))" proof - from assms obtain \ where \_p: "C\ = C \ \ \ is_ground_subst \" unfolding grounding_of_cls_def by auto define \ where "\ = Infer {#(C + {#L#})\ \#} ((D + {#L'#}) \ \ \ \) (C \ \)" have "(D + {#L'#}) \ \ \ \ \ grounding_of_clss (CC \ {C + {#L#}})" unfolding grounding_of_clss_def grounding_of_cls_def - by (rule UN_I[of "D + {#L'#}"], use assms(2) clss_of_state_def in simp, + by (rule UN_I[of "D + {#L'#}"], use assms(2) in simp, metis (mono_tags, lifting) \_p is_ground_comp_subst mem_Collect_eq subst_cls_comp_subst) moreover have "(C + {#L#}) \ \ \ grounding_of_clss (CC \ {C + {#L#}})" using \_p unfolding grounding_of_clss_def grounding_of_cls_def by auto - moreover have "\I. I \ D \ \ \ \ + {#- (L \l \)#} \ I \ C \ \ + {#L \l \#} \ I \ D \ \ \ \ + C \ \" + moreover have + "\I. I \ D \ \ \ \ + {#- (L \l \)#} \ I \ C \ \ + {#L \l \#} \ I \ D \ \ \ \ + C \ \" by auto then have "\I. I \ (D + {#L'#}) \ \ \ \ \ I \ (C + {#L#}) \ \ \ I \ D \ \ \ \ + C \ \" using assms by (metis add_mset_add_single subst_cls_add_mset subst_cls_union subst_minus) then have "\I. I \ (D + {#L'#}) \ \ \ \ \ I \ (C + {#L#}) \ \ \ I \ C \ \" using assms by (metis (no_types, lifting) subset_mset.le_iff_add subst_cls_union true_cls_union) then have "\I. I \m {#(D + {#L'#}) \ \ \ \#} \ I \ (C + {#L#}) \ \ \ I \ C \ \" by (meson true_cls_mset_singleton) ultimately have "\ \ sr_ext.inferences_from (grounding_of_clss (CC \ {C + {#L#}}))" unfolding sr_ext.inferences_from_def unfolding ground_sound_\_def infer_from_def \_def by auto then have "C \ \ \ concls_of (sr_ext.inferences_from (grounding_of_clss (CC \ {C + {#L#}})))" using image_iff unfolding \_def by fastforce then show "C\ \ concls_of (sr_ext.inferences_from (grounding_of_clss (CC \ {C + {#L#}})))" using \_p by auto qed lemma reduction_derivable: assumes "D + {#L'#} \ CC" and "- L = L' \l \" and "D \ \ \# C" shows "sr_ext.derive (grounding_of_clss (CC \ {C + {#L#}})) (grounding_of_clss (CC \ {C}))" proof - from assms have "grounding_of_clss (CC \ {C}) - grounding_of_clss (CC \ {C + {#L#}}) \ concls_of (sr_ext.inferences_from (grounding_of_clss (CC \ {C + {#L#}})))" - using reduction_in_concls_of unfolding grounding_of_clss_def clss_of_state_def - by auto + using reduction_in_concls_of unfolding grounding_of_clss_def by auto moreover have "grounding_of_cls (C + {#L#}) \ sr.Rf (grounding_of_clss (CC \ {C}))" using strict_subset_subsumption_grounding_redundant_clss[of C "id_subst"] by auto then have "grounding_of_clss (CC \ {C + {#L#}}) - grounding_of_clss (CC \ {C}) \ sr.Rf (grounding_of_clss (CC \ {C}))" - unfolding clss_of_state_def grounding_of_clss_def by auto - ultimately show "sr_ext.derive (grounding_of_clss (CC \ {C + {#L#}})) (grounding_of_clss (CC \ {C}))" + unfolding grounding_of_clss_def by auto + ultimately show + "sr_ext.derive (grounding_of_clss (CC \ {C + {#L#}})) (grounding_of_clss (CC \ {C}))" using sr_ext.derive.intros[of "grounding_of_clss (CC \ {C})" "grounding_of_clss (CC \ {C + {#L#}})"] by auto qed text \ The following corresponds the part of Lemma 4.10 that states we have a theorem proving process: \ lemma RP_ground_derive: "St \ St' \ sr_ext.derive (grounding_of_state St) (grounding_of_state St')" proof (induction rule: RP.induct) case (tautology_deletion A C N P Q) { fix C\ assume "C\ \ grounding_of_cls C" then obtain \ where "C\ = C \ \" unfolding grounding_of_cls_def by auto then have "Neg (A \a \) \# C\ \ Pos (A \a \) \# C\" using tautology_deletion Neg_Melem_subst_atm_subst_cls Pos_Melem_subst_atm_subst_cls by auto then have "C\ \ sr.Rf (grounding_of_state (N, P, Q))" using sr.tautology_Rf by auto } then have "grounding_of_state (N \ {C}, P, Q) - grounding_of_state (N, P, Q) \ sr.Rf (grounding_of_state (N, P, Q))" - unfolding clss_of_state_def grounding_of_clss_def by auto + unfolding grounding_of_clss_def by auto moreover have "grounding_of_state (N, P, Q) - grounding_of_state (N \ {C}, P, Q) = {}" - unfolding clss_of_state_def grounding_of_clss_def by auto + unfolding grounding_of_clss_def by auto ultimately show ?case - using sr_ext.derive.intros[of "grounding_of_state (N, P, Q)" "grounding_of_state (N \ {C}, P, Q)"] + using sr_ext.derive.intros[of "grounding_of_state (N, P, Q)" + "grounding_of_state (N \ {C}, P, Q)"] by auto next case (forward_subsumption D P Q C N) then show ?case - using derive_if_remove_subsumed[of D "(N, P, Q)" C] unfolding grounding_of_clss_def clss_of_state_def + using derive_if_remove_subsumed[of D "(N, P, Q)" C] unfolding grounding_of_clss_def by (simp add: sup_commute sup_left_commute) next case (backward_subsumption_P D N C P Q) then show ?case - using derive_if_remove_subsumed[of D "(N, P, Q)" C] strictly_subsumes_def unfolding grounding_of_clss_def clss_of_state_def - by (simp add: sup_commute sup_left_commute) + using derive_if_remove_subsumed[of D "(N, P, Q)" C] strictly_subsumes_def + unfolding grounding_of_clss_def by (simp add: sup_commute sup_left_commute) next case (backward_subsumption_Q D N C P Q) then show ?case - using derive_if_remove_subsumed[of D "(N, P, Q)" C] strictly_subsumes_def unfolding grounding_of_clss_def clss_of_state_def - by (simp add: sup_commute sup_left_commute) + using derive_if_remove_subsumed[of D "(N, P, Q)" C] strictly_subsumes_def + unfolding grounding_of_clss_def by (simp add: sup_commute sup_left_commute) next case (forward_reduction D L' P Q L \ C N) then show ?case - using reduction_derivable[of _ _ "N \ P \ Q"] unfolding clss_of_state_def by force + using reduction_derivable[of _ _ "N \ P \ Q"] by force next case (backward_reduction_P D L' N L \ C P Q) then show ?case - using reduction_derivable[of _ _ "N \ P \ Q"] unfolding clss_of_state_def by force + using reduction_derivable[of _ _ "N \ P \ Q"] by force next case (backward_reduction_Q D L' N L \ C P Q) then show ?case - using reduction_derivable[of _ _ "N \ P \ Q"] unfolding clss_of_state_def by force + using reduction_derivable[of _ _ "N \ P \ Q"] by force next case (clause_processing N C P Q) then show ?case - unfolding clss_of_state_def using sr_ext.derive.intros by auto + using sr_ext.derive.intros by auto next case (inference_computation N Q C P) { fix E\ assume "E\ \ grounding_of_clss N" then obtain \ E where E_\_p: "E\ = E \ \ \ E \ N \ is_ground_subst \" unfolding grounding_of_clss_def grounding_of_cls_def by auto then have E_concl: "E \ concls_of (ord_FO_resolution.inferences_between Q C)" using inference_computation by auto then obtain \ where \_p: "\ \ ord_FO_\ S \ infer_from (Q \ {C}) \ \ C \# prems_of \ \ concl_of \ = E" unfolding ord_FO_resolution.inferences_between_def by auto then obtain CC CAs D AAs As \ where \_p2: "\ = Infer CC D E \ ord_resolve_rename S CAs D AAs As \ E \ mset CAs = CC" unfolding ord_FO_\_def by auto define \ where "\ = hd (renamings_apart (D # CAs))" define \s where "\s = tl (renamings_apart (D # CAs))" define \_ground where "\_ground = Infer (mset (CAs \\cl \s) \cm \ \cm \) (D \ \ \ \ \ \) (E \ \)" have "\I. I \m mset (CAs \\cl \s) \cm \ \cm \ \ I \ D \ \ \ \ \ \ \ I \ E \ \" using ord_resolve_rename_ground_inst_sound[of _ _ _ _ _ _ _ _ _ _ \] \_def \s_def E_\_p \_p2 by auto then have "\_ground \ {Infer cc d e | cc d e. \I. I \m cc \ I \ d \ I \ e}" unfolding \_ground_def by auto moreover have "set_mset (prems_of \_ground) \ grounding_of_state ({}, P \ {C}, Q)" proof - have "D = C \ D \ Q" unfolding \_ground_def using E_\_p \_p2 \_p unfolding infer_from_def - unfolding clss_of_state_def grounding_of_clss_def - unfolding grounding_of_cls_def - by simp + unfolding grounding_of_clss_def grounding_of_cls_def by simp then have "D \ \ \ \ \ \ \ grounding_of_cls C \ (\x \ Q. D \ \ \ \ \ \ \ grounding_of_cls x)" using E_\_p unfolding grounding_of_cls_def by (metis (mono_tags, lifting) is_ground_comp_subst mem_Collect_eq subst_cls_comp_subst) then have "(D \ \ \ \ \ \ \ grounding_of_cls C \ (\x \ P. D \ \ \ \ \ \ \ grounding_of_cls x) \ (\x \ Q. D \ \ \ \ \ \ \ grounding_of_cls x))" by metis moreover have "\i < length (CAs \\cl \s \cl \ \cl \). ((CAs \\cl \s \cl \ \cl \) ! i) \ {C \ \ |\. is_ground_subst \} \ ((\C \ P. {C \ \ | \. is_ground_subst \}) \ (\C\Q. {C \ \ | \. is_ground_subst \}))" proof (rule, rule) fix i assume "i < length (CAs \\cl \s \cl \ \cl \)" then have a: "i < length CAs \ i < length \s" by simp moreover from a have "CAs ! i \ {C} \ Q" using \_p2 \_p unfolding infer_from_def by (metis (no_types, lifting) Un_subset_iff inference.sel(1) set_mset_union sup_commute nth_mem_mset subsetCE) ultimately have "(CAs \\cl \s \cl \ \cl \) ! i \ {C \ \ |\. is_ground_subst \} \ ((CAs \\cl \s \cl \ \cl \) ! i \ (\C\P. {C \ \ |\. is_ground_subst \}) \ (CAs \\cl \s \cl \ \cl \) ! i \ (\C \ Q. {C \ \ | \. is_ground_subst \}))" - unfolding \_ground_def using E_\_p \_p2 \_p unfolding infer_from_def - unfolding clss_of_state_def grounding_of_clss_def - unfolding grounding_of_cls_def + using E_\_p \_p2 \_p + unfolding \_ground_def infer_from_def grounding_of_clss_def grounding_of_cls_def apply - apply (cases "CAs ! i = C") subgoal apply (rule disjI1) apply (rule Set.CollectI) - apply (rule_tac x="(\s ! i) \ \ \ \" in exI) - using \s_def using renamings_apart_length apply (auto;fail) - done + apply (rule_tac x = "(\s ! i) \ \ \ \" in exI) + using \s_def using renamings_apart_length by (auto; fail) subgoal apply (rule disjI2) apply (rule disjI2) - apply (rule_tac a="CAs ! i" in UN_I) - subgoal - apply blast - done + apply (rule_tac a = "CAs ! i" in UN_I) + subgoal by blast subgoal apply (rule Set.CollectI) - apply (rule_tac x="(\s ! i) \ \ \ \" in exI) - using \s_def using renamings_apart_length apply (auto;fail) - done + apply (rule_tac x = "(\s ! i) \ \ \ \" in exI) + using \s_def using renamings_apart_length by (auto; fail) done done then show "(CAs \\cl \s \cl \ \cl \) ! i \ {C \ \ |\. is_ground_subst \} \ ((\C \ P. {C \ \ |\. is_ground_subst \}) \ (\C \ Q. {C \ \ |\. is_ground_subst \}))" by blast qed then have "\x \# mset (CAs \\cl \s \cl \ \cl \). x \ {C \ \ |\. is_ground_subst \} \ ((\C \ P. {C \ \ |\. is_ground_subst \}) \ (\C \ Q. {C \ \ |\. is_ground_subst \}))" by (metis (lifting) in_set_conv_nth set_mset_mset) then have "set_mset (mset (CAs \\cl \s) \cm \ \cm \) \ grounding_of_cls C \ grounding_of_clss P \ grounding_of_clss Q" unfolding grounding_of_cls_def grounding_of_clss_def using mset_subst_cls_list_subst_cls_mset by auto ultimately show ?thesis - unfolding \_ground_def clss_of_state_def grounding_of_clss_def by auto + unfolding \_ground_def grounding_of_clss_def by auto qed - ultimately have "E \ \ \ concls_of (sr_ext.inferences_from (grounding_of_state ({}, P \ {C}, Q)))" - unfolding sr_ext.inferences_from_def inference_system.inferences_from_def ground_sound_\_def infer_from_def + ultimately have + "E \ \ \ concls_of (sr_ext.inferences_from (grounding_of_state ({}, P \ {C}, Q)))" + unfolding sr_ext.inferences_from_def inference_system.inferences_from_def ground_sound_\_def + infer_from_def using \_ground_def by (metis (no_types, lifting) imageI inference.sel(3) mem_Collect_eq) then have "E\ \ concls_of (sr_ext.inferences_from (grounding_of_state ({}, P \ {C}, Q)))" using E_\_p by auto } then have "grounding_of_state (N, P, Q \ {C}) - grounding_of_state ({}, P \ {C}, Q) \ concls_of (sr_ext.inferences_from (grounding_of_state ({}, P \ {C}, Q)))" - unfolding clss_of_state_def grounding_of_clss_def by auto + unfolding grounding_of_clss_def by auto moreover have "grounding_of_state ({}, P \ {C}, Q) - grounding_of_state (N, P, Q \ {C}) = {}" - unfolding clss_of_state_def grounding_of_clss_def by auto + unfolding grounding_of_clss_def by auto ultimately show ?case using sr_ext.derive.intros[of "(grounding_of_state (N, P, Q \ {C}))" "(grounding_of_state ({}, P \ {C}, Q))"] by auto qed text \ A useful consequence: \ theorem RP_model: "St \ St' \ I \s grounding_of_state St' \ I \s grounding_of_state St" proof (drule RP_ground_derive, erule sr_ext.derive.cases, hypsubst) let ?gSt = "grounding_of_state St" and ?gSt' = "grounding_of_state St'" assume deduct: "?gSt' - ?gSt \ concls_of (sr_ext.inferences_from ?gSt)" (is "_ \ ?concls") and delete: "?gSt - ?gSt' \ sr.Rf ?gSt'" show "I \s ?gSt' \ I \s ?gSt" proof assume bef: "I \s ?gSt" then have "I \s ?concls" - unfolding ground_sound_\_def inference_system.inferences_from_def true_clss_def true_cls_mset_def + unfolding ground_sound_\_def inference_system.inferences_from_def true_clss_def + true_cls_mset_def by (auto simp add: image_def infer_from_def dest!: spec[of _ I]) then have diff: "I \s ?gSt' - ?gSt" using deduct by (blast intro: true_clss_mono) then show "I \s ?gSt'" using bef unfolding true_clss_def by blast next assume aft: "I \s ?gSt'" have "I \s ?gSt' \ sr.Rf ?gSt'" - by (rule sr.Rf_model) (metis aft sr.Rf_mono[OF Un_upper1] Diff_eq_empty_iff Diff_subset - Un_Diff true_clss_mono true_clss_union) + by (rule sr.Rf_model) (smt Diff_eq_empty_iff Diff_subset Un_Diff aft + standard_redundancy_criterion.Rf_mono sup_bot.right_neutral sup_ge1 true_clss_mono) then have "I \s sr.Rf ?gSt'" using true_clss_union by blast then have diff: "I \s ?gSt - ?gSt'" using delete by (blast intro: true_clss_mono) then show "I \s ?gSt" using aft unfolding true_clss_def by blast qed qed text \ Another formulation of the part of Lemma 4.10 that states we have a theorem proving process: \ -(* FIXME: rename *) -lemma RP_ground_derive_chain: - "chain sr_ext.derive (lmap grounding_of_state Sts)" +lemma ground_derive_chain: "chain sr_ext.derive (lmap grounding_of_state Sts)" using deriv RP_ground_derive by (simp add: chain_lmap[of "(\)"]) text \ The following is used prove to Lemma 4.11: \ lemma in_Sup_llist_in_nth: "C \ Sup_llist Gs \ \j. enat j < llength Gs \ C \ lnth Gs j" unfolding Sup_llist_def by auto \ \Note: Gs is called Ns in the chapter\ lemma Sup_llist_grounding_of_state_ground: assumes "C \ Sup_llist (lmap grounding_of_state Sts)" shows "is_ground_cls C" proof - - have "\j. enat j < llength (lmap grounding_of_state Sts) \ C \ lnth (lmap grounding_of_state Sts) j" - using assms in_Sup_llist_in_nth by metis + have "\j. enat j < llength (lmap grounding_of_state Sts) + \ C \ lnth (lmap grounding_of_state Sts) j" + using assms in_Sup_llist_in_nth by fast then obtain j where "enat j < llength (lmap grounding_of_state Sts)" "C \ lnth (lmap grounding_of_state Sts) j" by blast then show ?thesis unfolding grounding_of_clss_def grounding_of_cls_def by auto qed lemma Liminf_grounding_of_state_ground: "C \ Liminf_llist (lmap grounding_of_state Sts) \ is_ground_cls C" using Liminf_llist_subset_Sup_llist[of "lmap grounding_of_state Sts"] Sup_llist_grounding_of_state_ground by blast lemma in_Sup_llist_in_Sup_state: assumes "C \ Sup_llist (lmap grounding_of_state Sts)" shows "\D \. D \ clss_of_state (Sup_state Sts) \ D \ \ = C \ is_ground_subst \" proof - from assms obtain i where i_p: "enat i < llength Sts \ C \ lnth (lmap grounding_of_state Sts) i" using in_Sup_llist_in_nth by fastforce then obtain D \ where "D \ clss_of_state (lnth Sts i) \ D \ \ = C \ is_ground_subst \" using assms unfolding grounding_of_clss_def grounding_of_cls_def by fastforce then have "D \ clss_of_state (Sup_state Sts) \ D \ \ = C \ is_ground_subst \" - using i_p unfolding Sup_state_def clss_of_state_def + using i_p unfolding Sup_state_def by (metis (no_types, lifting) UnCI UnE contra_subsetD N_of_state.simps P_of_state.simps Q_of_state.simps llength_lmap lnth_lmap lnth_subset_Sup_llist) then show ?thesis by auto qed lemma N_of_state_Liminf: "N_of_state (Liminf_state Sts) = Liminf_llist (lmap N_of_state Sts)" and P_of_state_Liminf: "P_of_state (Liminf_state Sts) = Liminf_llist (lmap P_of_state Sts)" unfolding Liminf_state_def by auto lemma eventually_removed_from_N: assumes d_in: "D \ N_of_state (lnth Sts i)" and fair: "fair_state_seq Sts" and i_Sts: "enat i < llength Sts" - shows "\l. D \ N_of_state (lnth Sts l) \ D \ N_of_state (lnth Sts (Suc l)) \ i \ l \ enat (Suc l) < llength Sts" + shows "\l. D \ N_of_state (lnth Sts l) \ D \ N_of_state (lnth Sts (Suc l)) \ i \ l + \ enat (Suc l) < llength Sts" proof (rule ccontr) assume a: "\ ?thesis" have "i \ l \ enat l < llength Sts \ D \ N_of_state (lnth Sts l)" for l using d_in by (induction l, blast, metis a Suc_ile_eq le_SucE less_imp_le) then have "D \ Liminf_llist (lmap N_of_state Sts)" unfolding Liminf_llist_def using i_Sts by auto then show False using fair unfolding fair_state_seq_def by (simp add: N_of_state_Liminf) qed lemma eventually_removed_from_P: assumes d_in: "D \ P_of_state (lnth Sts i)" and fair: "fair_state_seq Sts" and i_Sts: "enat i < llength Sts" - shows "\l. D \ P_of_state (lnth Sts l) \ D \ P_of_state (lnth Sts (Suc l)) \ i \ l \ enat (Suc l) < llength Sts" + shows "\l. D \ P_of_state (lnth Sts l) \ D \ P_of_state (lnth Sts (Suc l)) \ i \ l + \ enat (Suc l) < llength Sts" proof (rule ccontr) assume a: "\ ?thesis" have "i \ l \ enat l < llength Sts \ D \ P_of_state (lnth Sts l)" for l using d_in by (induction l, blast, metis a Suc_ile_eq le_SucE less_imp_le) then have "D \ Liminf_llist (lmap P_of_state Sts)" unfolding Liminf_llist_def using i_Sts by auto then show False using fair unfolding fair_state_seq_def by (simp add: P_of_state_Liminf) qed -(* FIXME: come up with a better name *) lemma instance_if_subsumed_and_in_limit: assumes ns: "Gs = lmap grounding_of_state Sts" and c: "C \ Liminf_llist Gs - sr.Rf (Liminf_llist Gs)" and - (* FIXME: use clss_of_state below *) - d: "D \ N_of_state (lnth Sts i) \ P_of_state (lnth Sts i) \ Q_of_state (lnth Sts i)" - "enat i < llength Sts" "subsumes D C" + d: "D \ clss_of_state (lnth Sts i)" "enat i < llength Sts" "subsumes D C" shows "\\. D \ \ = C \ is_ground_subst \" proof - let ?Ps = "\i. P_of_state (lnth Sts i)" let ?Qs = "\i. Q_of_state (lnth Sts i)" have ground_C: "is_ground_cls C" using c using Liminf_grounding_of_state_ground ns by auto have derivns: "chain sr_ext.derive Gs" - using RP_ground_derive_chain deriv ns by auto + using ground_derive_chain deriv ns by auto have "\\. D \ \ = C" proof (rule ccontr) assume "\\. D \ \ = C" moreover from d(3) obtain \_proto where "D \ \_proto \# C" unfolding subsumes_def by blast then obtain \ where \_p: "D \ \ \# C \ is_ground_subst \" using ground_C by (metis is_ground_cls_mono make_ground_subst subset_mset.order_refl) ultimately have subsub: "D \ \ \# C" using subset_mset.le_imp_less_or_eq by auto moreover have "is_ground_subst \" using \_p by auto moreover have "D \ clss_of_state (lnth Sts i)" - using d unfolding clss_of_state_def by auto + using d by auto ultimately have "C \ sr.Rf (grounding_of_state (lnth Sts i))" using strict_subset_subsumption_redundant_clss by auto then have "C \ sr.Rf (Sup_llist Gs)" - using d ns by (metis contra_subsetD llength_lmap lnth_lmap lnth_subset_Sup_llist sr.Rf_mono) + using d ns by (smt contra_subsetD llength_lmap lnth_lmap lnth_subset_Sup_llist sr.Rf_mono) then have "C \ sr.Rf (Liminf_llist Gs)" unfolding ns using local.sr_ext.Rf_limit_Sup derivns ns by auto then show False using c by auto qed then obtain \ where "D \ \ = C \ is_ground_subst \" using ground_C by (metis make_ground_subst) then show ?thesis by auto qed lemma from_Q_to_Q_inf: assumes fair: "fair_state_seq Sts" and ns: "Gs = lmap grounding_of_state Sts" and c: "C \ Liminf_llist Gs - sr.Rf (Liminf_llist Gs)" and d: "D \ Q_of_state (lnth Sts i)" "enat i < llength Sts" "subsumes D C" and d_least: "\E \ {E. E \ (clss_of_state (Sup_state Sts)) \ subsumes E C}. \ strictly_subsumes E D" shows "D \ Q_of_state (Liminf_state Sts)" proof - let ?Ps = "\i. P_of_state (lnth Sts i)" let ?Qs = "\i. Q_of_state (lnth Sts i)" have ground_C: "is_ground_cls C" using c using Liminf_grounding_of_state_ground ns by auto have derivns: "chain sr_ext.derive Gs" - using RP_ground_derive_chain deriv ns by auto + using ground_derive_chain deriv ns by auto have "\\. D \ \ = C \ is_ground_subst \" - using instance_if_subsumed_and_in_limit ns c d by blast + using instance_if_subsumed_and_in_limit c d unfolding ns by blast then obtain \ where \: "D \ \ = C" "is_ground_subst \" by auto have in_Sts_in_Sts_Suc: - "\l \ i. enat (Suc l) < llength Sts \ D \ Q_of_state (lnth Sts l) \ D \ Q_of_state (lnth Sts (Suc l))" + "\l \ i. enat (Suc l) < llength Sts \ D \ Q_of_state (lnth Sts l) \ + D \ Q_of_state (lnth Sts (Suc l))" proof (rule, rule, rule, rule) fix l assume len: "i \ l" and llen: "enat (Suc l) < llength Sts" and d_in_q: "D \ Q_of_state (lnth Sts l)" have "lnth Sts l \ lnth Sts (Suc l)" using llen deriv chain_lnth_rel by blast then show "D \ Q_of_state (lnth Sts (Suc l))" proof (cases rule: RP.cases) case (backward_subsumption_Q D' N D_removed P Q) moreover { assume "D_removed = D" then obtain D_subsumes where D_subsumes_p: "D_subsumes \ N \ strictly_subsumes D_subsumes D" using backward_subsumption_Q by auto moreover from D_subsumes_p have "subsumes D_subsumes C" using d subsumes_trans unfolding strictly_subsumes_def by blast moreover from backward_subsumption_Q have "D_subsumes \ clss_of_state (Sup_state Sts)" using D_subsumes_p llen - by (metis (no_types) UnI1 clss_of_state_def N_of_state.simps llength_lmap lnth_lmap - lnth_subset_Sup_llist rev_subsetD Sup_state_def) + by (metis (no_types) UnI1 N_of_state.simps llength_lmap lnth_lmap lnth_subset_Sup_llist + rev_subsetD Sup_state_def) ultimately have False using d_least unfolding subsumes_def by auto } ultimately show ?thesis using d_in_q by auto next case (backward_reduction_Q E L' N L \ D' P Q) { assume "D' + {#L#} = D" then have D'_p: "strictly_subsumes D' D \ D' \ ?Ps (Suc l)" using subset_strictly_subsumes[of D' D] backward_reduction_Q by auto then have subc: "subsumes D' C" using d(3) subsumes_trans unfolding strictly_subsumes_def by auto from D'_p have "D' \ clss_of_state (Sup_state Sts)" - using llen by (metis (no_types) UnI1 clss_of_state_def P_of_state.simps llength_lmap - lnth_lmap lnth_subset_Sup_llist subsetCE sup_ge2 Sup_state_def) + using llen by (metis (no_types) UnI1 P_of_state.simps llength_lmap lnth_lmap + lnth_subset_Sup_llist subsetCE sup_ge2 Sup_state_def) then have False using d_least D'_p subc by auto } then show ?thesis using backward_reduction_Q d_in_q by auto qed (use d_in_q in auto) qed have D_in_Sts: "D \ Q_of_state (lnth Sts l)" and D_in_Sts_Suc: "D \ Q_of_state (lnth Sts (Suc l))" if l_i: "l \ i" and enat: "enat (Suc l) < llength Sts" for l proof - show "D \ Q_of_state (lnth Sts l)" using l_i enat apply (induction "l - i" arbitrary: l) subgoal using d by auto subgoal using d(1) in_Sts_in_Sts_Suc by (metis (no_types, lifting) Suc_ile_eq add_Suc_right add_diff_cancel_left' le_SucE le_Suc_ex less_imp_le) done then show "D \ Q_of_state (lnth Sts (Suc l))" using l_i enat in_Sts_in_Sts_Suc by blast qed have "i \ x \ enat x < llength Sts \ D \ Q_of_state (lnth Sts x)" for x apply (cases x) subgoal using d(1) by (auto intro!: exI[of _ i] simp: less_Suc_eq) subgoal for x' using d(1) D_in_Sts_Suc[of x'] by (cases \i \ x'\) (auto simp: not_less_eq_eq) done then have "D \ Liminf_llist (lmap Q_of_state Sts)" unfolding Liminf_llist_def by (auto intro!: exI[of _ i] simp: d) then show ?thesis unfolding Liminf_state_def by auto qed lemma from_P_to_Q: assumes fair: "fair_state_seq Sts" and ns: "Gs = lmap grounding_of_state Sts" and c: "C \ Liminf_llist Gs - sr.Rf (Liminf_llist Gs)" and d: "D \ P_of_state (lnth Sts i)" "enat i < llength Sts" "subsumes D C" and - d_least: "\E \ {E. E \ (clss_of_state (Sup_state Sts)) \ subsumes E C}. \ strictly_subsumes E D" + d_least: "\E \ {E. E \ (clss_of_state (Sup_state Sts)) \ subsumes E C}. + \ strictly_subsumes E D" shows "\l. D \ Q_of_state (lnth Sts l) \ enat l < llength Sts" proof - let ?Ns = "\i. N_of_state (lnth Sts i)" let ?Ps = "\i. P_of_state (lnth Sts i)" let ?Qs = "\i. Q_of_state (lnth Sts i)" have ground_C: "is_ground_cls C" using c using Liminf_grounding_of_state_ground ns by auto have derivns: "chain sr_ext.derive Gs" - using RP_ground_derive_chain deriv ns by auto + using ground_derive_chain deriv ns by auto have "\\. D \ \ = C \ is_ground_subst \" using instance_if_subsumed_and_in_limit ns c d by blast then obtain \ where \: "D \ \ = C" "is_ground_subst \" by auto obtain l where - l_p: "D \ P_of_state (lnth Sts l) \ D \ P_of_state (lnth Sts (Suc l)) \ i \ l \ enat (Suc l) < llength Sts" + l_p: "D \ P_of_state (lnth Sts l) \ D \ P_of_state (lnth Sts (Suc l)) \ i \ l + \ enat (Suc l) < llength Sts" using fair using eventually_removed_from_P d unfolding ns by auto then have l_Gs: "enat (Suc l) < llength Gs" using ns by auto from l_p have "lnth Sts l \ lnth Sts (Suc l)" using deriv using chain_lnth_rel by auto then show ?thesis proof (cases rule: RP.cases) case (backward_subsumption_P D' N D_twin P Q) note lrhs = this(1,2) and D'_p = this(3,4) then have twins: "D_twin = D" "?Ns (Suc l) = N" "?Ns l = N" "?Ps (Suc l) = P" "?Ps l = P \ {D_twin}" "?Qs (Suc l) = Q" "?Qs l = Q" using l_p by auto note D'_p = D'_p[unfolded twins(1)] then have subc: "subsumes D' C" unfolding strictly_subsumes_def subsumes_def using \ by (metis subst_cls_comp_subst subst_cls_mono_mset) from D'_p have "D' \ clss_of_state (Sup_state Sts)" unfolding twins(2)[symmetric] using l_p - by (metis (no_types) UnI1 clss_of_state_def N_of_state.simps llength_lmap lnth_lmap - lnth_subset_Sup_llist subsetCE Sup_state_def) + by (metis (no_types) UnI1 N_of_state.simps llength_lmap lnth_lmap lnth_subset_Sup_llist + subsetCE Sup_state_def) then have False using d_least D'_p subc by auto then show ?thesis by auto next case (backward_reduction_P E L' N L \ D' P Q) then have twins: "D' + {#L#} = D" "?Ns (Suc l) = N" "?Ns l = N" "?Ps (Suc l) = P \ {D'}" "?Ps l = P \ {D' + {#L#}}" "?Qs (Suc l) = Q" "?Qs l = Q" using l_p by auto then have D'_p: "strictly_subsumes D' D \ D' \ ?Ps (Suc l)" using subset_strictly_subsumes[of D' D] by auto then have subc: "subsumes D' C" using d(3) subsumes_trans unfolding strictly_subsumes_def by auto from D'_p have "D' \ clss_of_state (Sup_state Sts)" - using l_p by (metis (no_types) UnI1 clss_of_state_def P_of_state.simps llength_lmap lnth_lmap + using l_p by (metis (no_types) UnI1 P_of_state.simps llength_lmap lnth_lmap lnth_subset_Sup_llist subsetCE sup_ge2 Sup_state_def) then have False using d_least D'_p subc by auto then show ?thesis by auto next case (inference_computation N Q D_twin P) then have twins: "D_twin = D" "?Ps (Suc l) = P" "?Ps l = P \ {D_twin}" "?Qs (Suc l) = Q \ {D_twin}" "?Qs l = Q" using l_p by auto then show ?thesis using d \ l_p by auto qed (use l_p in auto) qed lemma variants_sym: "variants D D' \ variants D' D" unfolding variants_def by auto lemma variants_imp_exists_subtitution: "variants D D' \ \\. D \ \ = D'" unfolding variants_iff_subsumes subsumes_def by (meson strictly_subsumes_def subset_mset_def strict_subset_subst_strictly_subsumes subsumes_def) lemma properly_subsume_variants: assumes "strictly_subsumes E D" and "variants D D'" shows "strictly_subsumes E D'" proof - from assms obtain \ \' where \_\'_p: "D \ \ = D' \ D' \ \' = D" using variants_imp_exists_subtitution variants_sym by metis from assms obtain \'' where "E \ \'' \# D" unfolding strictly_subsumes_def subsumes_def by auto then have "E \ \'' \ \ \# D \ \" using subst_cls_mono_mset by blast then have "E \ (\'' \ \) \# D'" using \_\'_p by auto moreover from assms have n: "(\\. D \ \ \# E)" unfolding strictly_subsumes_def subsumes_def by auto have "\\. D' \ \ \# E" proof assume "\\'''. D' \ \''' \# E" then obtain \''' where "D' \ \''' \# E" by auto then have "D \ (\ \ \''') \# E" using \_\'_p by auto then show False using n by metis qed ultimately show ?thesis unfolding strictly_subsumes_def subsumes_def by metis qed lemma neg_properly_subsume_variants: assumes "\ strictly_subsumes E D" and "variants D D'" shows "\ strictly_subsumes E D'" using assms properly_subsume_variants variants_sym by auto lemma from_N_to_P_or_Q: assumes fair: "fair_state_seq Sts" and ns: "Gs = lmap grounding_of_state Sts" and c: "C \ Liminf_llist Gs - sr.Rf (Liminf_llist Gs)" and d: "D \ N_of_state (lnth Sts i)" "enat i < llength Sts" "subsumes D C" and d_least: "\E \ {E. E \ (clss_of_state (Sup_state Sts)) \ subsumes E C}. \ strictly_subsumes E D" shows "\l D' \'. D' \ P_of_state (lnth Sts l) \ Q_of_state (lnth Sts l) \ enat l < llength Sts \ (\E \ {E. E \ (clss_of_state (Sup_state Sts)) \ subsumes E C}. \ strictly_subsumes E D') \ D' \ \' = C \ is_ground_subst \' \ subsumes D' C" proof - let ?Ns = "\i. N_of_state (lnth Sts i)" let ?Ps = "\i. P_of_state (lnth Sts i)" let ?Qs = "\i. Q_of_state (lnth Sts i)" have ground_C: "is_ground_cls C" using c using Liminf_grounding_of_state_ground ns by auto have derivns: "chain sr_ext.derive Gs" - using RP_ground_derive_chain deriv ns by auto + using ground_derive_chain deriv ns by auto have "\\. D \ \ = C \ is_ground_subst \" using instance_if_subsumed_and_in_limit ns c d by blast then obtain \ where \: "D \ \ = C" "is_ground_subst \" by auto from c have no_taut: "\ (\A. Pos A \# C \ Neg A \# C)" using sr.tautology_Rf by auto - have "\l. D \ N_of_state (lnth Sts l) \ D \ N_of_state (lnth Sts (Suc l)) \ i \ l \ enat (Suc l) < llength Sts" + have "\l. D \ N_of_state (lnth Sts l) + \ D \ N_of_state (lnth Sts (Suc l)) \ i \ l \ enat (Suc l) < llength Sts" using fair using eventually_removed_from_N d unfolding ns by auto then obtain l where - l_p: "D \ N_of_state (lnth Sts l) \ D \ N_of_state (lnth Sts (Suc l)) \ i \ l \ enat (Suc l) < llength Sts" + l_p: "D \ N_of_state (lnth Sts l) \ D \ N_of_state (lnth Sts (Suc l)) \ i \ l + \ enat (Suc l) < llength Sts" by auto then have l_Gs: "enat (Suc l) < llength Gs" using ns by auto from l_p have "lnth Sts l \ lnth Sts (Suc l)" using deriv using chain_lnth_rel by auto then show ?thesis proof (cases rule: RP.cases) case (tautology_deletion A D_twin N P Q) then have "D_twin = D" using l_p by auto then have "Pos (A \a \) \# C \ Neg (A \a \) \# C" using tautology_deletion(3,4) \ by (metis Melem_subst_cls eql_neg_lit_eql_atm eql_pos_lit_eql_atm) then have False using no_taut by metis then show ?thesis by blast next case (forward_subsumption D' P Q D_twin N) note lrhs = this(1,2) and D'_p = this(3,4) then have twins: "D_twin = D" "?Ns (Suc l) = N" "?Ns l = N \ {D_twin}" "?Ps (Suc l) = P " "?Ps l = P" "?Qs (Suc l) = Q" "?Qs l = Q" using l_p by auto note D'_p = D'_p[unfolded twins(1)] from D'_p(2) have subs: "subsumes D' C" using d(3) by (blast intro: subsumes_trans) moreover have "D' \ clss_of_state (Sup_state Sts)" - using twins D'_p l_p unfolding clss_of_state_def Sup_state_def + using twins D'_p l_p unfolding Sup_state_def by simp (metis (no_types) contra_subsetD llength_lmap lnth_lmap lnth_subset_Sup_llist) ultimately have "\ strictly_subsumes D' D" using d_least by auto then have "subsumes D D'" unfolding strictly_subsumes_def using D'_p by auto then have v: "variants D D'" using D'_p unfolding variants_iff_subsumes by auto - then have mini: "\E \ {E \ clss_of_state (Sup_state Sts). subsumes E C}. \ strictly_subsumes E D'" + then have mini: "\E \ {E \ clss_of_state (Sup_state Sts). subsumes E C}. + \ strictly_subsumes E D'" using d_least D'_p neg_properly_subsume_variants[of _ D D'] by auto from v have "\\'. D' \ \' = C" using \ variants_imp_exists_subtitution variants_sym by (metis subst_cls_comp_subst) then have "\\'. D' \ \' = C \ is_ground_subst \'" using ground_C by (meson make_ground_subst refl) then obtain \' where \'_p: "D' \ \' = C \ is_ground_subst \'" by metis show ?thesis using D'_p twins l_p subs mini \'_p by auto next case (forward_reduction E L' P Q L \ D' N) then have twins: "D' + {#L#} = D" "?Ns (Suc l) = N \ {D'}" "?Ns l = N \ {D' + {#L#}}" "?Ps (Suc l) = P " "?Ps l = P" "?Qs (Suc l) = Q" "?Qs l = Q" using l_p by auto then have D'_p: "strictly_subsumes D' D \ D' \ ?Ns (Suc l)" using subset_strictly_subsumes[of D' D] by auto then have subc: "subsumes D' C" using d(3) subsumes_trans unfolding strictly_subsumes_def by blast from D'_p have "D' \ clss_of_state (Sup_state Sts)" - using l_p by (metis (no_types) UnI1 clss_of_state_def N_of_state.simps llength_lmap lnth_lmap + using l_p by (metis (no_types) UnI1 N_of_state.simps llength_lmap lnth_lmap lnth_subset_Sup_llist subsetCE Sup_state_def) then have False using d_least D'_p subc by auto then show ?thesis by auto next case (clause_processing N D_twin P Q) then have twins: "D_twin = D" "?Ns (Suc l) = N" "?Ns l = N \ {D}" "?Ps (Suc l) = P \ {D}" "?Ps l = P" "?Qs (Suc l) = Q" "?Qs l = Q" using l_p by auto then show ?thesis using d \ l_p d_least by blast qed (use l_p in auto) qed lemma eventually_in_Qinf: assumes D_p: "D \ clss_of_state (Sup_state Sts)" - "subsumes D C" "\E \ {E. E \ (clss_of_state (Sup_state Sts)) \ subsumes E C}. \ strictly_subsumes E D" and + "subsumes D C" "\E \ {E. E \ (clss_of_state (Sup_state Sts)) \ subsumes E C}. + \ strictly_subsumes E D" and fair: "fair_state_seq Sts" and - (* We could also, we guess, in this proof obtain a D with property D_p(3) from one with only properties D_p(2,3). *) + (* We could also, we guess, in this proof obtain a D with property D_p(3) from one with only + properties D_p(2,3). *) ns: "Gs = lmap grounding_of_state Sts" and c: "C \ Liminf_llist Gs - sr.Rf (Liminf_llist Gs)" and ground_C: "is_ground_cls C" shows "\D' \'. D' \ Q_of_state (Liminf_state Sts) \ D' \ \' = C \ is_ground_subst \'" proof - let ?Ns = "\i. N_of_state (lnth Sts i)" let ?Ps = "\i. P_of_state (lnth Sts i)" let ?Qs = "\i. Q_of_state (lnth Sts i)" from D_p obtain i where i_p: "i < llength Sts" "D \ ?Ns i \ D \ ?Ps i \ D \ ?Qs i" - unfolding clss_of_state_def Sup_state_def + unfolding Sup_state_def by simp_all (metis (no_types) in_Sup_llist_in_nth llength_lmap lnth_lmap) - have derivns: "chain sr_ext.derive Gs" using RP_ground_derive_chain deriv ns by auto + have derivns: "chain sr_ext.derive Gs" + using ground_derive_chain deriv ns by auto have "\\. D \ \ = C \ is_ground_subst \" using instance_if_subsumed_and_in_limit[OF ns c] D_p i_p by blast then obtain \ where \: "D \ \ = C" "is_ground_subst \" by blast { assume a: "D \ ?Ns i" then obtain D' \' l where D'_p: "D' \ ?Ps l \ ?Qs l" "D' \ \' = C" "enat l < llength Sts" "is_ground_subst \'" "\E \ {E. E \ (clss_of_state (Sup_state Sts)) \ subsumes E C}. \ strictly_subsumes E D'" "subsumes D' C" using from_N_to_P_or_Q deriv fair ns c i_p(1) D_p(2) D_p(3) by blast then obtain l' where l'_p: "D' \ ?Qs l'" "l' < llength Sts" using from_P_to_Q[OF fair ns c _ D'_p(3) D'_p(6) D'_p(5)] by blast then have "D' \ Q_of_state (Liminf_state Sts)" using from_Q_to_Q_inf[OF fair ns c _ l'_p(2)] D'_p by auto then have ?thesis using D'_p by auto } moreover { assume a: "D \ ?Ps i" then obtain l' where l'_p: "D \ ?Qs l'" "l' < llength Sts" using from_P_to_Q[OF fair ns c a i_p(1) D_p(2) D_p(3)] by auto then have "D \ Q_of_state (Liminf_state Sts)" using from_Q_to_Q_inf[OF fair ns c l'_p(1) l'_p(2)] D_p(3) \(1) \(2) D_p(2) by auto then have ?thesis using D_p \ by auto } moreover { assume a: "D \ ?Qs i" then have "D \ Q_of_state (Liminf_state Sts)" using from_Q_to_Q_inf[OF fair ns c a i_p(1)] \ D_p(2,3) by auto then have ?thesis using D_p \ by auto } ultimately show ?thesis using i_p by auto qed text \ The following corresponds to Lemma 4.11: \ lemma fair_imp_Liminf_minus_Rf_subset_ground_Liminf_state: assumes fair: "fair_state_seq Sts" and ns: "Gs = lmap grounding_of_state Sts" shows "Liminf_llist Gs - sr.Rf (Liminf_llist Gs) \ grounding_of_clss (Q_of_state (Liminf_state Sts))" proof let ?Ns = "\i. N_of_state (lnth Sts i)" let ?Ps = "\i. P_of_state (lnth Sts i)" let ?Qs = "\i. Q_of_state (lnth Sts i)" have SQinf: "clss_of_state (Liminf_state Sts) = Liminf_llist (lmap Q_of_state Sts)" - using fair unfolding fair_state_seq_def Liminf_state_def clss_of_state_def by auto + using fair unfolding fair_state_seq_def Liminf_state_def by auto fix C assume C_p: "C \ Liminf_llist Gs - sr.Rf (Liminf_llist Gs)" then have "C \ Sup_llist Gs" using Liminf_llist_subset_Sup_llist[of Gs] by blast then obtain D_proto where "D_proto \ clss_of_state (Sup_state Sts) \ subsumes D_proto C" using in_Sup_llist_in_Sup_state unfolding ns subsumes_def by blast then obtain D where D_p: "D \ clss_of_state (Sup_state Sts)" "subsumes D C" "\E \ {E. E \ clss_of_state (Sup_state Sts) \ subsumes E C}. \ strictly_subsumes E D" using strictly_subsumes_has_minimum[of "{E. E \ clss_of_state (Sup_state Sts) \ subsumes E C}"] by auto have ground_C: "is_ground_cls C" using C_p using Liminf_grounding_of_state_ground ns by auto have "\D' \'. D' \ Q_of_state (Liminf_state Sts) \ D' \ \' = C \ is_ground_subst \'" - using eventually_in_Qinf[of D C Gs] using D_p(1) D_p(2) D_p(3) fair ns C_p ground_C by auto + using eventually_in_Qinf[of D C Gs] using D_p(1-3) fair ns C_p ground_C by auto then obtain D' \' where D'_p: "D' \ Q_of_state (Liminf_state Sts) \ D' \ \' = C \ is_ground_subst \'" by blast then have "D' \ clss_of_state (Liminf_state Sts)" - by (simp add: clss_of_state_def) + by simp then have "C \ grounding_of_state (Liminf_state Sts)" unfolding grounding_of_clss_def grounding_of_cls_def using D'_p by auto then show "C \ grounding_of_clss (Q_of_state (Liminf_state Sts))" - using SQinf clss_of_state_def fair fair_state_seq_def by auto + using SQinf fair fair_state_seq_def by auto qed text \ The following corresponds to (one direction of) Theorem 4.13: \ lemma ground_subclauses: assumes "\i < length CAs. CAs ! i = Cs ! i + poss (AAs ! i)" and "length Cs = length CAs" and "is_ground_cls_list CAs" shows "is_ground_cls_list Cs" unfolding is_ground_cls_list_def by (metis assms in_set_conv_nth is_ground_cls_list_def is_ground_cls_union) lemma subseteq_Liminf_state_eventually_always: fixes CC assumes "finite CC" and "CC \ {}" and "CC \ Q_of_state (Liminf_state Sts)" shows "\j. enat j < llength Sts \ (\j' \ enat j. j' < llength Sts \ CC \ Q_of_state (lnth Sts j'))" proof - from assms(3) have "\C \ CC. \j. enat j < llength Sts \ (\j' \ enat j. j' < llength Sts \ C \ Q_of_state (lnth Sts j'))" unfolding Liminf_state_def Liminf_llist_def by force then obtain f where f_p: "\C \ CC. f C < llength Sts \ (\j' \ enat (f C). j' < llength Sts \ C \ Q_of_state (lnth Sts j'))" by moura define j :: nat where "j = Max (f ` CC)" have "enat j < llength Sts" unfolding j_def using f_p assms(1) by (metis (mono_tags) Max_in assms(2) finite_imageI imageE image_is_empty) moreover have "\C j'. C \ CC \ enat j \ j' \ j' < llength Sts \ C \ Q_of_state (lnth Sts j')" proof (intro allI impI) fix C :: "'a clause" and j' :: nat assume a: "C \ CC" "enat j \ enat j'" "enat j' < llength Sts" then have "f C \ j'" unfolding j_def using assms(1) Max.bounded_iff by auto then show "C \ Q_of_state (lnth Sts j')" using f_p a by auto qed ultimately show ?thesis by auto qed lemma empty_clause_in_Q_of_Liminf_state: assumes empty_in: "{#} \ Liminf_llist (lmap grounding_of_state Sts)" and fair: "fair_state_seq Sts" shows "{#} \ Q_of_state (Liminf_state Sts)" proof - define Gs :: "'a clause set llist" where ns: "Gs = lmap grounding_of_state Sts" from empty_in have in_Liminf_not_Rf: "{#} \ Liminf_llist Gs - sr.Rf (Liminf_llist Gs)" unfolding ns sr.Rf_def by auto then have "{#} \ grounding_of_clss (Q_of_state (Liminf_state Sts))" using fair_imp_Liminf_minus_Rf_subset_ground_Liminf_state[OF fair ns] by auto then show ?thesis unfolding grounding_of_clss_def grounding_of_cls_def by auto qed lemma grounding_of_state_Liminf_state_subseteq: "grounding_of_state (Liminf_state Sts) \ Liminf_llist (lmap grounding_of_state Sts)" proof fix C :: "'a clause" assume "C \ grounding_of_state (Liminf_state Sts)" then obtain D \ where D_\_p: "D \ clss_of_state (Liminf_state Sts)" "D \ \ = C" "is_ground_subst \" - unfolding clss_of_state_def grounding_of_clss_def grounding_of_cls_def by auto - then have ii: "D \ Liminf_llist (lmap N_of_state Sts) \ - D \ Liminf_llist (lmap P_of_state Sts) \ - D \ Liminf_llist (lmap Q_of_state Sts)" - unfolding clss_of_state_def Liminf_state_def by simp - then have "C \ Liminf_llist (lmap grounding_of_clss (lmap N_of_state Sts)) \ - C \ Liminf_llist (lmap grounding_of_clss (lmap P_of_state Sts)) \ - C \ Liminf_llist (lmap grounding_of_clss (lmap Q_of_state Sts))" + unfolding grounding_of_clss_def grounding_of_cls_def by auto + then have ii: "D \ Liminf_llist (lmap N_of_state Sts) + \ D \ Liminf_llist (lmap P_of_state Sts) \ D \ Liminf_llist (lmap Q_of_state Sts)" + unfolding Liminf_state_def by simp + then have "C \ Liminf_llist (lmap grounding_of_clss (lmap N_of_state Sts)) + \ C \ Liminf_llist (lmap grounding_of_clss (lmap P_of_state Sts)) + \ C \ Liminf_llist (lmap grounding_of_clss (lmap Q_of_state Sts))" unfolding Liminf_llist_def grounding_of_clss_def grounding_of_cls_def apply - apply (erule disjE) subgoal apply (rule disjI1) using D_\_p by auto subgoal apply (erule HOL.disjE) subgoal apply (rule disjI2) apply (rule disjI1) using D_\_p by auto subgoal apply (rule disjI2) apply (rule disjI2) using D_\_p by auto done done then show "C \ Liminf_llist (lmap grounding_of_state Sts)" - unfolding Liminf_llist_def clss_of_state_def grounding_of_clss_def by auto + unfolding Liminf_llist_def grounding_of_clss_def by auto qed theorem RP_sound: assumes "{#} \ clss_of_state (Liminf_state Sts)" shows "\ satisfiable (grounding_of_state (lhd Sts))" proof - from assms have "{#} \ grounding_of_state (Liminf_state Sts)" unfolding grounding_of_clss_def by (force intro: ex_ground_subst) then have "{#} \ Liminf_llist (lmap grounding_of_state Sts)" using grounding_of_state_Liminf_state_subseteq by auto then have "\ satisfiable (Liminf_llist (lmap grounding_of_state Sts))" using true_clss_def by auto then have "\ satisfiable (lhd (lmap grounding_of_state Sts))" - using sr_ext.sat_limit_iff RP_ground_derive_chain by metis + using sr_ext.sat_limit_iff ground_derive_chain by blast then show ?thesis unfolding lhd_lmap_Sts . qed lemma ground_ord_resolve_ground: assumes CAs_p: "gr.ord_resolve CAs DA AAs As E" and ground_cas: "is_ground_cls_list CAs" and ground_da: "is_ground_cls DA" shows "is_ground_cls E" proof - have a1: "atms_of E \ (\CA \ set CAs. atms_of CA) \ atms_of DA" using gr.ord_resolve_atms_of_concl_subset[of CAs DA _ _ E] CAs_p by auto { fix L :: "'a literal" assume "L \# E" then have "atm_of L \ atms_of E" by (meson atm_of_lit_in_atms_of) then have "is_ground_atm (atm_of L)" using a1 ground_cas ground_da is_ground_cls_imp_is_ground_atm is_ground_cls_list_def by auto } then show ?thesis unfolding is_ground_cls_def is_ground_lit_def by simp qed theorem RP_saturated_if_fair: - assumes fair: "fair_state_seq Sts" + assumes + fair: "fair_state_seq Sts" and + empty_Q0: "Q_of_state (lhd Sts) = {}" shows "sr.saturated_upto (Liminf_llist (lmap grounding_of_state Sts))" proof - define Gs :: "'a clause set llist" where ns: "Gs = lmap grounding_of_state Sts" let ?N = "\i. grounding_of_state (lnth Sts i)" let ?Ns = "\i. N_of_state (lnth Sts i)" let ?Ps = "\i. P_of_state (lnth Sts i)" let ?Qs = "\i. Q_of_state (lnth Sts i)" have ground_ns_in_ground_limit_st: "Liminf_llist Gs - sr.Rf (Liminf_llist Gs) \ grounding_of_clss (Q_of_state (Liminf_state Sts))" using fair deriv fair_imp_Liminf_minus_Rf_subset_ground_Liminf_state ns by blast have derivns: "chain sr_ext.derive Gs" - using RP_ground_derive_chain deriv ns by auto + using ground_derive_chain deriv ns by auto { fix \ :: "'a inference" assume \_p: "\ \ gr.ord_\" let ?CC = "side_prems_of \" let ?DA = "main_prem_of \" let ?E = "concl_of \" assume a: "set_mset ?CC \ {?DA} - \ Liminf_llist (lmap grounding_of_state Sts) - sr.Rf (Liminf_llist (lmap grounding_of_state Sts))" + \ Liminf_llist (lmap grounding_of_state Sts) + - sr.Rf (Liminf_llist (lmap grounding_of_state Sts))" have ground_ground_Liminf: "is_ground_clss (Liminf_llist (lmap grounding_of_state Sts))" using Liminf_grounding_of_state_ground unfolding is_ground_clss_def by auto have ground_cc: "is_ground_clss (set_mset ?CC)" using a ground_ground_Liminf is_ground_clss_def by auto have ground_da: "is_ground_cls ?DA" using a grounding_ground singletonI ground_ground_Liminf by (simp add: Liminf_grounding_of_state_ground) from \_p obtain CAs AAs As where CAs_p: "gr.ord_resolve CAs ?DA AAs As ?E \ mset CAs = ?CC" unfolding gr.ord_\_def by auto have DA_CAs_in_ground_Liminf: "{?DA} \ set CAs \ grounding_of_clss (Q_of_state (Liminf_state Sts))" - using a CAs_p unfolding clss_of_state_def using fair unfolding fair_state_seq_def - by (metis (no_types, lifting) Un_empty_left ground_ns_in_ground_limit_st a clss_of_state_def - ns set_mset_mset subset_trans sup_commute) + using a CAs_p fair unfolding fair_state_seq_def + by (metis (no_types, lifting) Un_empty_left ground_ns_in_ground_limit_st a ns set_mset_mset + subset_trans sup_commute) then have ground_cas: "is_ground_cls_list CAs" using CAs_p unfolding is_ground_cls_list_def by auto then have ground_e: "is_ground_cls ?E" using ground_ord_resolve_ground CAs_p ground_da by auto have "\AAs As \. ord_resolve (S_M S (Q_of_state (Liminf_state Sts))) CAs ?DA AAs As \ ?E" using CAs_p[THEN conjunct1] proof (cases rule: gr.ord_resolve.cases) case (ord_resolve n Cs D) note DA = this(1) and e = this(2) and cas_len = this(3) and cs_len = this(4) and aas_len = this(5) and as_len = this(6) and nz = this(7) and cas = this(8) and aas_not_empt = this(9) and as_aas = this(10) and eligibility = this(11) and str_max = this(12) and sel_empt = this(13) have len_aas_len_as: "length AAs = length As" using aas_len as_len by auto from as_aas have "\iA \# add_mset (As ! i) (AAs ! i). A = As ! i" using ord_resolve by simp then have "\i < n. card (set_mset (add_mset (As ! i) (AAs ! i))) \ Suc 0" using all_the_same by metis then have "\i < length AAs. card (set_mset (add_mset (As ! i) (AAs ! i))) \ Suc 0" using aas_len by auto then have "\AA \ set (map2 add_mset As AAs). card (set_mset AA) \ Suc 0" using set_map2_ex[of AAs As add_mset, OF len_aas_len_as] by auto then have "is_unifiers id_subst (set_mset ` set (map2 add_mset As AAs))" unfolding is_unifiers_def is_unifier_def by auto moreover have "finite (set_mset ` set (map2 add_mset As AAs))" by auto moreover have "\AA \ set_mset ` set (map2 add_mset As AAs). finite AA" by auto ultimately obtain \ where \_p: "Some \ = mgu (set_mset ` set (map2 add_mset As AAs))" using mgu_complete by metis have ground_elig: "gr.eligible As (D + negs (mset As))" using ord_resolve by simp have ground_cs: "\i < n. is_ground_cls (Cs ! i)" using ord_resolve(8) ord_resolve(3,4) ground_cas using ground_subclauses[of CAs Cs AAs] unfolding is_ground_cls_list_def by auto have ground_set_as: "is_ground_atms (set As)" using ord_resolve(1) ground_da - by (metis atms_of_negs is_ground_cls_union set_mset_mset is_ground_cls_is_ground_atms_atms_of) + by (metis atms_of_negs is_ground_cls_union set_mset_mset + is_ground_cls_is_ground_atms_atms_of) then have ground_mset_as: "is_ground_atm_mset (mset As)" unfolding is_ground_atm_mset_def is_ground_atms_def by auto have ground_as: "is_ground_atm_list As" using ground_set_as is_ground_atm_list_def is_ground_atms_def by auto have ground_d: "is_ground_cls D" using ground_da ord_resolve by simp from as_len nz have "atms_of D \ set As \ {}" "finite (atms_of D \ set As)" by auto then have "Max (atms_of D \ set As) \ atms_of D \ set As" using Max_in by metis then have is_ground_Max: "is_ground_atm (Max (atms_of D \ set As))" using ground_d ground_mset_as is_ground_cls_imp_is_ground_atm unfolding is_ground_atm_mset_def by auto then have Max\_is_Max: "\\. Max (atms_of D \ set As) \a \ = Max (atms_of D \ set As)" by auto have ann1: "maximal_wrt (Max (atms_of D \ set As)) (D + negs (mset As))" unfolding maximal_wrt_def by clarsimp (metis Max_less_iff UnCI \atms_of D \ set As \ {}\ \finite (atms_of D \ set As)\ ground_d ground_set_as infinite_growing is_ground_Max is_ground_atms_def is_ground_cls_imp_is_ground_atm less_atm_ground) from ground_elig have ann2: "Max (atms_of D \ set As) \a \ = Max (atms_of D \ set As)" "D \ \ + negs (mset As \am \) = D + negs (mset As)" using is_ground_Max ground_mset_as ground_d by auto from ground_elig have fo_elig: "eligible (S_M S (Q_of_state (Liminf_state Sts))) \ As (D + negs (mset As))" unfolding gr.eligible.simps eligible.simps gr.maximal_wrt_def using ann1 ann2 by (auto simp: S_Q_def) have l: "\i < n. gr.strictly_maximal_wrt (As ! i) (Cs ! i)" using ord_resolve by simp then have "\i < n. strictly_maximal_wrt (As ! i) (Cs ! i)" unfolding gr.strictly_maximal_wrt_def strictly_maximal_wrt_def using ground_as[unfolded is_ground_atm_list_def] ground_cs as_len less_atm_ground by clarsimp (fastforce simp: is_ground_cls_as_atms)+ then have ll: "\i < n. strictly_maximal_wrt (As ! i \a \) (Cs ! i \ \)" by (simp add: ground_as ground_cs as_len) have m: "\i < n. S_Q (CAs ! i) = {#}" using ord_resolve by simp have ground_e: "is_ground_cls (\# (mset Cs) + D)" using ground_d ground_cs ground_e e by simp show ?thesis - using ord_resolve.intros[OF cas_len cs_len aas_len as_len nz cas aas_not_empt \_p fo_elig ll] m DA e ground_e + using ord_resolve.intros + [OF cas_len cs_len aas_len as_len nz cas aas_not_empt \_p fo_elig ll] m DA e ground_e unfolding S_Q_def by auto qed then obtain AAs As \ where \_p: "ord_resolve (S_M S (Q_of_state (Liminf_state Sts))) CAs ?DA AAs As \ ?E" by auto then obtain \s' \' \2' CAs' DA' AAs' As' \' E' where s_p: "is_ground_subst \'" "is_ground_subst_list \s'" "is_ground_subst \2'" "ord_resolve_rename S CAs' DA' AAs' As' \' E'" "CAs' \\cl \s' = CAs" "DA' \ \' = ?DA" "E' \ \2' = ?E" "{DA'} \ set CAs' \ Q_of_state (Liminf_state Sts)" using ord_resolve_rename_lifting[OF sel_stable, of "Q_of_state (Liminf_state Sts)" CAs ?DA] \_p selection_axioms DA_CAs_in_ground_Liminf by metis from this(8) have "\j. enat j < llength Sts \ (set CAs' \ {DA'} \ ?Qs j)" unfolding Liminf_llist_def using subseteq_Liminf_state_eventually_always[of "{DA'} \ set CAs'"] by auto then obtain j where j_p: "is_least (\j. enat j < llength Sts \ set CAs' \ {DA'} \ ?Qs j) j" using least_exists[of "\j. enat j < llength Sts \ set CAs' \ {DA'} \ ?Qs j"] by force then have j_p': "enat j < llength Sts" "set CAs' \ {DA'} \ ?Qs j" unfolding is_least_def by auto then have jn0: "j \ 0" using empty_Q0 by (metis bot_eq_sup_iff gr_implies_not_zero insert_not_empty llength_lnull lnth_0_conv_lhd sup.orderE) then have j_adds_CAs': "\ set CAs' \ {DA'} \ ?Qs (j - 1)" "set CAs' \ {DA'} \ ?Qs j" using j_p unfolding is_least_def apply (metis (no_types) One_nat_def Suc_diff_Suc Suc_ile_eq diff_diff_cancel diff_zero less_imp_le less_one neq0_conv zero_less_diff) using j_p'(2) by blast have "lnth Sts (j - 1) \ lnth Sts j" - using j_p'(1) jn0 deriv chain_lnth_rel[of _ _ "j - 1"] by force + using j_p'(1) jn0 deriv chain_lnth_rel[of _ _ "j - 1"] by force then obtain C' where C'_p: "?Ns (j - 1) = {}" "?Ps (j - 1) = ?Ps j \ {C'}" "?Qs j = ?Qs (j - 1) \ {C'}" "?Ns j = concls_of (ord_FO_resolution.inferences_between (?Qs (j - 1)) C')" "C' \ set CAs' \ {DA'}" "C' \ ?Qs (j - 1)" using j_adds_CAs' by (induction rule: RP.cases) auto have "E' \ ?Ns j" proof - have "E' \ concls_of (ord_FO_resolution.inferences_between (Q_of_state (lnth Sts (j - 1))) C')" - unfolding infer_from_def ord_FO_\_def unfolding inference_system.inferences_between_def + unfolding infer_from_def ord_FO_\_def inference_system.inferences_between_def apply (rule_tac x = "Infer (mset CAs') DA' E'" in image_eqI) subgoal by auto subgoal - using s_p(4) unfolding infer_from_def - apply (rule ord_resolve_rename.cases) - using s_p(4) - using C'_p(3) C'_p(5) j_p'(2) apply force - done + by (rule ord_resolve_rename.cases[OF s_p(4)]) + (use s_p(4) C'_p(3) C'_p(5) j_p'(2) in force) done then show ?thesis using C'_p(4) by auto qed then have "E' \ clss_of_state (lnth Sts j)" - using j_p' unfolding clss_of_state_def by auto + using j_p' by auto then have "?E \ grounding_of_state (lnth Sts j)" using s_p(7) s_p(3) unfolding grounding_of_clss_def grounding_of_cls_def by force then have "\ \ sr.Ri (grounding_of_state (lnth Sts j))" using sr.Ri_effective \_p by auto then have "\ \ sr_ext_Ri (?N j)" unfolding sr_ext_Ri_def by auto then have "\ \ sr_ext_Ri (Sup_llist (lmap grounding_of_state Sts))" - using j_p' contra_subsetD llength_lmap lnth_lmap lnth_subset_Sup_llist sr_ext.Ri_mono by metis + using j_p' contra_subsetD llength_lmap lnth_lmap lnth_subset_Sup_llist sr_ext.Ri_mono by smt then have "\ \ sr_ext_Ri (Liminf_llist (lmap grounding_of_state Sts))" using sr_ext.Ri_limit_Sup[of Gs] derivns ns by blast } then have "sr_ext.saturated_upto (Liminf_llist (lmap grounding_of_state Sts))" unfolding sr_ext.saturated_upto_def sr_ext.inferences_from_def infer_from_def sr_ext_Ri_def by auto then show ?thesis using gd_ord_\_ngd_ord_\ sr.redundancy_criterion_axioms redundancy_criterion_standard_extension_saturated_upto_iff[of gr.ord_\] unfolding sr_ext_Ri_def by auto qed corollary RP_complete_if_fair: assumes fair: "fair_state_seq Sts" and + empty_Q0: "Q_of_state (lhd Sts) = {}" and unsat: "\ satisfiable (grounding_of_state (lhd Sts))" shows "{#} \ Q_of_state (Liminf_state Sts)" proof - have "\ satisfiable (Liminf_llist (lmap grounding_of_state Sts))" - unfolding sr_ext.sat_limit_iff[OF RP_ground_derive_chain] + unfolding sr_ext.sat_limit_iff[OF ground_derive_chain] by (rule unsat[folded lhd_lmap_Sts[of grounding_of_state]]) moreover have "sr.saturated_upto (Liminf_llist (lmap grounding_of_state Sts))" - by (rule RP_saturated_if_fair[OF fair, simplified]) + by (rule RP_saturated_if_fair[OF fair empty_Q0, simplified]) ultimately have "{#} \ Liminf_llist (lmap grounding_of_state Sts)" using sr.saturated_upto_complete_if by auto then show ?thesis using empty_clause_in_Q_of_Liminf_state fair by auto qed end end end diff --git a/thys/Ordered_Resolution_Prover/Ordered_Ground_Resolution.thy b/thys/Ordered_Resolution_Prover/Ordered_Ground_Resolution.thy --- a/thys/Ordered_Resolution_Prover/Ordered_Ground_Resolution.thy +++ b/thys/Ordered_Resolution_Prover/Ordered_Ground_Resolution.thy @@ -1,467 +1,468 @@ (* Title: Ground Ordered Resolution Calculus with Selection Author: Anders Schlichtkrull , 2016, 2017 Author: Jasmin Blanchette , 2014, 2017 Author: Dmitriy Traytel , 2014 Maintainer: Anders Schlichtkrull *) section \Ground Ordered Resolution Calculus with Selection\ theory Ordered_Ground_Resolution imports Inference_System Ground_Resolution_Model begin text \ Ordered ground resolution with selection is the second inference system studied in Section~3 (``Standard Resolution'') of Bachmair and Ganzinger's chapter. \ subsection \Inference Rule\ text \ Ordered ground resolution consists of a single rule, called \ord_resolve\ below. Like \unord_resolve\, the rule is sound and counterexample-reducing. In addition, it is reductive. \ context ground_resolution_with_selection begin text \ The following inductive definition corresponds to Figure 2. \ definition maximal_wrt :: "'a \ 'a literal multiset \ bool" where - "maximal_wrt A DA \ A = Max (atms_of DA)" (* FIXME: change definition so that it returns true if DA is empty *) + "maximal_wrt A DA \ DA = {#} \ A = Max (atms_of DA)" definition strictly_maximal_wrt :: "'a \ 'a literal multiset \ bool" where "strictly_maximal_wrt A CA \ (\B \ atms_of CA. B < A)" inductive eligible :: "'a list \ 'a clause \ bool" where eligible: "(S DA = negs (mset As)) \ (S DA = {#} \ length As = 1 \ maximal_wrt (As ! 0) DA) \ eligible As DA" lemma "(S DA = negs (mset As) \ S DA = {#} \ length As = 1 \ maximal_wrt (As ! 0) DA) \ eligible As DA" using eligible.intros ground_resolution_with_selection.eligible.cases ground_resolution_with_selection_axioms by blast - inductive ord_resolve :: "'a clause list \ 'a clause \ 'a multiset list \ 'a list \ 'a clause \ bool" where ord_resolve: "length CAs = n \ length Cs = n \ length AAs = n \ length As = n \ n \ 0 \ (\i < n. CAs ! i = Cs ! i + poss (AAs ! i)) \ (\i < n. AAs ! i \ {#}) \ (\i < n. \A \# AAs ! i. A = As ! i) \ eligible As (D + negs (mset As)) \ (\i < n. strictly_maximal_wrt (As ! i) (Cs ! i)) \ (\i < n. S (CAs ! i) = {#}) \ ord_resolve CAs (D + negs (mset As)) AAs As (\# (mset Cs) + D)" lemma ord_resolve_sound: assumes res_e: "ord_resolve CAs DA AAs As E" and cc_true: "I \m mset CAs" and d_true: "I \ DA" shows "I \ E" using res_e proof (cases rule: ord_resolve.cases) case (ord_resolve n Cs D) note DA = this(1) and e = this(2) and cas_len = this(3) and cs_len = this(4) and as_len = this(6) and cas = this(8) and aas_ne = this(9) and a_eq = this(10) show ?thesis proof (cases "\A \ set As. A \ I") case True then have "\ I \ negs (mset As)" unfolding true_cls_def by fastforce then have "I \ D" using d_true DA by fast then show ?thesis unfolding e by blast next case False then obtain i where a_in_aa: "i < n" and a_false: "As ! i \ I" using cas_len as_len by (metis in_set_conv_nth) have "\ I \ poss (AAs ! i)" using a_false a_eq aas_ne a_in_aa unfolding true_cls_def by auto moreover have "I \ CAs ! i" using a_in_aa cc_true unfolding true_cls_mset_def using cas_len by auto ultimately have "I \ Cs ! i" using cas a_in_aa by auto then show ?thesis using a_in_aa cs_len unfolding e true_cls_def by (meson in_Union_mset_iff nth_mem_mset union_iff) qed qed lemma filter_neg_atm_of_S: "{#Neg (atm_of L). L \# S C#} = S C" by (simp add: S_selects_neg_lits) text \ This corresponds to Lemma 3.13: \ lemma ord_resolve_reductive: assumes "ord_resolve CAs DA AAs As E" shows "E < DA" using assms proof (cases rule: ord_resolve.cases) case (ord_resolve n Cs D) note DA = this(1) and e = this(2) and cas_len = this(3) and cs_len = this(4) and ai_len = this(6) and nz = this(7) and cas = this(8) and maxim = this(12) show ?thesis proof (cases "\# (mset Cs) = {#}") case True have "negs (mset As) \ {#}" using nz ai_len by auto then show ?thesis unfolding True e DA by auto next case False - define max_A_of_Cs where "max_A_of_Cs = Max (atms_of (\# (mset Cs)))" + define max_A_of_Cs where + "max_A_of_Cs = Max (atms_of (\# (mset Cs)))" have mc_in: "max_A_of_Cs \ atms_of (\# (mset Cs))" and mc_max: "\B. B \ atms_of (\# (mset Cs)) \ B \ max_A_of_Cs" using max_A_of_Cs_def False by auto then have "\C_max \ set Cs. max_A_of_Cs \ atms_of (C_max)" by (metis atm_imp_pos_or_neg_lit in_Union_mset_iff neg_lit_in_atms_of pos_lit_in_atms_of set_mset_mset) then obtain max_i where cm_in_cas: "max_i < length CAs" and mc_in_cm: "max_A_of_Cs \ atms_of (Cs ! max_i)" using in_set_conv_nth[of _ CAs] by (metis cas_len cs_len in_set_conv_nth) define CA_max where "CA_max = CAs ! max_i" define A_max where "A_max = As ! max_i" define C_max where "C_max = Cs ! max_i" have mc_lt_ma: "max_A_of_Cs < A_max" using maxim cm_in_cas mc_in_cm cas_len unfolding strictly_maximal_wrt_def A_max_def by auto then have ucas_ne_neg_aa: "\# (mset Cs) \ negs (mset As)" using mc_in mc_max mc_lt_ma cm_in_cas cas_len ai_len unfolding A_max_def by (metis atms_of_negs nth_mem set_mset_mset leD) moreover have ucas_lt_ma: "\B \ atms_of (\# (mset Cs)). B < A_max" using mc_max mc_lt_ma by fastforce moreover have "\ Neg A_max \# \# (mset Cs)" using ucas_lt_ma neg_lit_in_atms_of[of A_max "\# (mset Cs)"] by auto moreover have "Neg A_max \# negs (mset As)" using cm_in_cas cas_len ai_len A_max_def by auto ultimately have "\# (mset Cs) < negs (mset As)" unfolding less_multiset\<^sub>H\<^sub>O by (metis (no_types) atms_less_eq_imp_lit_less_eq_neg count_greater_zero_iff count_inI le_imp_less_or_eq less_imp_not_less not_le) then show ?thesis unfolding e DA by auto qed qed text \ This corresponds to Theorem 3.15: \ theorem ord_resolve_counterex_reducing: assumes ec_ni_n: "{#} \ N" and d_in_n: "DA \ N" and d_cex: "\ INTERP N \ DA" and d_min: "\C. C \ N \ \ INTERP N \ C \ DA \ C" obtains CAs AAs As E where "set CAs \ N" "INTERP N \m mset CAs" "\CA. CA \ set CAs \ productive N CA" "ord_resolve CAs DA AAs As E" "\ INTERP N \ E" "E < DA" proof - have d_ne: "DA \ {#}" using d_in_n ec_ni_n by blast have "\As. As \ [] \ negs (mset As) \# DA \ eligible As DA" proof (cases "S DA = {#}") assume s_d_e: "S DA = {#}" define A where "A = Max (atms_of DA)" define As where "As = [A]" define D where "D = DA-{#Neg A #}" have na_in_d: "Neg A \# DA" unfolding A_def using s_d_e d_ne d_in_n d_cex d_min by (metis Max_in_lits Max_lit_eq_pos_or_neg_Max_atm max_pos_imp_Interp Interp_imp_INTERP) - then have das: "DA = D + negs (mset As)" unfolding D_def As_def by auto + then have das: "DA = D + negs (mset As)" + unfolding D_def As_def by auto moreover from na_in_d have "negs (mset As) \# DA" by (simp add: As_def) - moreover have "As ! 0 = Max (atms_of (D + negs (mset As)))" + moreover have hd: "As ! 0 = Max (atms_of (D + negs (mset As)))" using A_def As_def das by auto then have "eligible As DA" using eligible s_d_e As_def das maximal_wrt_def by auto ultimately show ?thesis using As_def by blast next assume s_d_e: "S DA \ {#}" define As :: "'a list" where "As = list_of_mset {#atm_of L. L \# S DA#}" define D :: "'a clause" where "D = DA - negs {#atm_of L. L \# S DA#}" have "As \ []" unfolding As_def using s_d_e by (metis image_mset_is_empty_iff list_of_mset_empty) moreover have da_sub_as: "negs {#atm_of L. L \# S DA#} \# DA" using S_selects_subseteq by (auto simp: filter_neg_atm_of_S) then have "negs (mset As) \# DA" unfolding As_def by auto moreover have das: "DA = D + negs (mset As)" using da_sub_as unfolding D_def As_def by auto moreover have "S DA = negs {#atm_of L. L \# S DA#}" by (auto simp: filter_neg_atm_of_S) then have "S DA = negs (mset As)" unfolding As_def by auto then have "eligible As DA" unfolding das using eligible by auto ultimately show ?thesis by blast qed then obtain As :: "'a list" where as_ne: "As \ []" and negs_as_le_d: "negs (mset As) \# DA" and s_d: "eligible As DA" by blast define D :: "'a clause" where "D = DA - negs (mset As)" have "set As \ INTERP N" using d_cex negs_as_le_d by force then have prod_ex: "\A \ set As. \D. produces N D A" unfolding INTERP_def by (metis (no_types, lifting) INTERP_def subsetCE UN_E not_produces_imp_notin_production) then have "\A. \D. produces N D A \ A \ set As" using ec_ni_n by (auto intro: productive_in_N) then have "\A. \D. produces N D A \ A \ set As" using prod_ex by blast then obtain CA_of where c_of0: "\A. produces N (CA_of A) A \ A \ set As" by metis then have prod_c0: "\A \ set As. produces N (CA_of A) A" by blast define C_of where "\A. C_of A = {#L \# CA_of A. L \ Pos A#}" define Aj_of where "\A. Aj_of A = image_mset atm_of {#L \# CA_of A. L = Pos A#}" have pospos: "\LL A. {#Pos (atm_of x). x \# {#L \# LL. L = Pos A#}#} = {#L \# LL. L = Pos A#}" by (metis (mono_tags, lifting) image_filter_cong literal.sel(1) multiset.map_ident) have ca_of_c_of_aj_of: "\A. CA_of A = C_of A + poss (Aj_of A)" using pospos[of _ "CA_of _"] by (simp add: C_of_def Aj_of_def) define n :: nat where "n = length As" define Cs :: "'a clause list" where "Cs = map C_of As" define AAs :: "'a multiset list" where "AAs = map Aj_of As" define CAs :: "'a literal multiset list" where "CAs = map CA_of As" have m_nz: "\A. A \ set As \ Aj_of A \ {#}" unfolding Aj_of_def using prod_c0 produces_imp_Pos_in_lits by (metis (full_types) filter_mset_empty_conv image_mset_is_empty_iff) have prod_c: "productive N CA" if ca_in: "CA \ set CAs" for CA proof - obtain i where i_p: "i < length CAs" "CAs ! i = CA" using ca_in by (meson in_set_conv_nth) have "production N (CA_of (As ! i)) = {As ! i}" using i_p CAs_def prod_c0 by auto then show "productive N CA" using i_p CAs_def by auto qed then have cs_subs_n: "set CAs \ N" using productive_in_N by auto have cs_true: "INTERP N \m mset CAs" unfolding true_cls_mset_def using prod_c productive_imp_INTERP by auto have "\A. A \ set As \ \ Neg A \# CA_of A" using prod_c0 produces_imp_neg_notin_lits by auto then have a_ni_c': "\A. A \ set As \ A \ atms_of (C_of A)" unfolding C_of_def using atm_imp_pos_or_neg_lit by force have c'_le_c: "\A. C_of A \ CA_of A" unfolding C_of_def by (auto intro: subset_eq_imp_le_multiset) have a_max_c: "\A. A \ set As \ A = Max (atms_of (CA_of A))" using prod_c0 productive_imp_produces_Max_atom[of N] by auto then have "\A. A \ set As \ C_of A \ {#} \ Max (atms_of (C_of A)) \ A" using c'_le_c by (metis less_eq_Max_atms_of) moreover have "\A. A \ set As \ C_of A \ {#} \ Max (atms_of (C_of A)) \ A" using a_ni_c' Max_in by (metis (no_types) atms_empty_iff_empty finite_atms_of) ultimately have max_c'_lt_a: "\A. A \ set As \ C_of A \ {#} \ Max (atms_of (C_of A)) < A" by (metis order.strict_iff_order) have le_cs_as: "length CAs = length As" unfolding CAs_def by simp have "length CAs = n" by (simp add: le_cs_as n_def) moreover have "length Cs = n" by (simp add: Cs_def n_def) moreover have "length AAs = n" by (simp add: AAs_def n_def) moreover have "length As = n" using n_def by auto moreover have "n \ 0" by (simp add: as_ne n_def) moreover have " \i. i < length AAs \ (\A \# AAs ! i. A = As ! i)" using AAs_def Aj_of_def by auto have "\x B. production N (CA_of x) = {x} \ B \# CA_of x \ B \ Pos x \ atm_of B < x" by (metis atm_of_lit_in_atms_of insert_not_empty le_imp_less_or_eq Pos_atm_of_iff Neg_atm_of_iff pos_neg_in_imp_true produces_imp_Pos_in_lits produces_imp_atms_leq productive_imp_not_interp) then have "\B A. A\set As \ B \# CA_of A \ B \ Pos A \ atm_of B < A" using prod_c0 by auto have "\i. i < length AAs \ AAs ! i \ {#}" unfolding AAs_def using m_nz by simp have "\i < n. CAs ! i = Cs ! i + poss (AAs ! i)" unfolding CAs_def Cs_def AAs_def using ca_of_c_of_aj_of by (simp add: n_def) moreover have "\i < n. AAs ! i \ {#}" using \\i < length AAs. AAs ! i \ {#}\ calculation(3) by blast moreover have "\i < n. \A \# AAs ! i. A = As ! i" by (simp add: \\i < length AAs. \A \# AAs ! i. A = As ! i\ calculation(3)) moreover have "eligible As DA" using s_d by auto then have "eligible As (D + negs (mset As))" using D_def negs_as_le_d by auto moreover have "\i. i < length AAs \ strictly_maximal_wrt (As ! i) ((Cs ! i))" by (simp add: C_of_def Cs_def \\x B. \production N (CA_of x) = {x}; B \# CA_of x; B \ Pos x\ \ atm_of B < x\ atms_of_def calculation(3) n_def prod_c0 strictly_maximal_wrt_def) have "\i < n. strictly_maximal_wrt (As ! i) (Cs ! i)" by (simp add: \\i. i < length AAs \ strictly_maximal_wrt (As ! i) (Cs ! i)\ calculation(3)) moreover have "\CA \ set CAs. S CA = {#}" using prod_c producesD productive_imp_produces_Max_literal by blast have "\CA\set CAs. S CA = {#}" using \\CA\set CAs. S CA = {#}\ by simp then have "\i < n. S (CAs ! i) = {#}" using \length CAs = n\ nth_mem by blast ultimately have res_e: "ord_resolve CAs (D + negs (mset As)) AAs As (\# (mset Cs) + D)" using ord_resolve by auto have "\A. A \ set As \ \ interp N (CA_of A) \ CA_of A" by (simp add: prod_c0 producesD) then have "\A. A \ set As \ \ Interp N (CA_of A) \ C_of A" unfolding prod_c0 C_of_def Interp_def true_cls_def using true_lit_def not_gr_zero prod_c0 by auto then have c'_at_n: "\A. A \ set As \ \ INTERP N \ C_of A" using a_max_c c'_le_c max_c'_lt_a not_Interp_imp_not_INTERP unfolding true_cls_def by (metis true_cls_def true_cls_empty) have "\ INTERP N \ \# (mset Cs)" unfolding Cs_def true_cls_def using c'_at_n by fastforce moreover have "\ INTERP N \ D" using d_cex by (metis D_def add_diff_cancel_right' negs_as_le_d subset_mset.add_diff_assoc2 true_cls_def union_iff) ultimately have e_cex: "\ INTERP N \ \# (mset Cs) + D" by simp have "set CAs \ N" by (simp add: cs_subs_n) moreover have "INTERP N \m mset CAs" by (simp add: cs_true) moreover have "\CA. CA \ set CAs \ productive N CA" by (simp add: prod_c) moreover have "ord_resolve CAs DA AAs As (\# (mset Cs) + D)" using D_def negs_as_le_d res_e by auto moreover have "\ INTERP N \ \# (mset Cs) + D" using e_cex by simp moreover have "\# (mset Cs) + D < DA" using calculation(4) ord_resolve_reductive by auto ultimately show thesis .. qed lemma ord_resolve_atms_of_concl_subset: assumes "ord_resolve CAs DA AAs As E" shows "atms_of E \ (\C \ set CAs. atms_of C) \ atms_of DA" using assms proof (cases rule: ord_resolve.cases) case (ord_resolve n Cs D) note DA = this(1) and e = this(2) and cas_len = this(3) and cs_len = this(4) and cas = this(8) have "\i < n. set_mset (Cs ! i) \ set_mset (CAs ! i)" using cas by auto then have "\i < n. Cs ! i \# \# (mset CAs)" by (metis cas cas_len mset_subset_eq_add_left nth_mem_mset sum_mset.remove union_assoc) then have "\C \ set Cs. C \# \# (mset CAs)" using cs_len in_set_conv_nth[of _ Cs] by auto then have "set_mset (\# (mset Cs)) \ set_mset (\# (mset CAs))" by auto (meson in_mset_sum_list2 mset_subset_eqD) then have "atms_of (\# (mset Cs)) \ atms_of (\# (mset CAs))" by (meson lits_subseteq_imp_atms_subseteq mset_subset_eqD subsetI) moreover have "atms_of (\# (mset CAs)) = (\CA \ set CAs. atms_of CA)" by (intro set_eqI iffI, simp_all, metis in_mset_sum_list2 atm_imp_pos_or_neg_lit neg_lit_in_atms_of pos_lit_in_atms_of, metis in_mset_sum_list atm_imp_pos_or_neg_lit neg_lit_in_atms_of pos_lit_in_atms_of) ultimately have "atms_of (\# (mset Cs)) \ (\CA \ set CAs. atms_of CA)" by auto moreover have "atms_of D \ atms_of DA" using DA by auto ultimately show ?thesis unfolding e by auto qed subsection \Inference System\ text \ Theorem 3.16 is subsumed in the counterexample-reducing inference system framework, which is instantiated below. Unlike its unordered cousin, ordered resolution is additionally a reductive inference system. \ definition ord_\ :: "'a inference set" where "ord_\ = {Infer (mset CAs) DA E | CAs DA AAs As E. ord_resolve CAs DA AAs As E}" sublocale ord_\_sound_counterex_reducing?: sound_counterex_reducing_inference_system "ground_resolution_with_selection.ord_\ S" "ground_resolution_with_selection.INTERP S" + reductive_inference_system "ground_resolution_with_selection.ord_\ S" proof unfold_locales fix DA :: "'a clause" and N :: "'a clause set" assume "{#} \ N" and "DA \ N" and "\ INTERP N \ DA" and "\C. C \ N \ \ INTERP N \ C \ DA \ C" then obtain CAs AAs As E where dd_sset_n: "set CAs \ N" and dd_true: "INTERP N \m mset CAs" and res_e: "ord_resolve CAs DA AAs As E" and e_cex: "\ INTERP N \ E" and e_lt_c: "E < DA" using ord_resolve_counterex_reducing[of N DA thesis] by auto have "Infer (mset CAs) DA E \ ord_\" using res_e unfolding ord_\_def by (metis (mono_tags, lifting) mem_Collect_eq) then show "\CC E. set_mset CC \ N \ INTERP N \m CC \ Infer CC DA E \ ord_\ \ \ INTERP N \ E \ E < DA" using dd_sset_n dd_true e_cex e_lt_c by (metis set_mset_mset) qed (auto simp: ord_\_def intro: ord_resolve_sound ord_resolve_reductive) lemmas clausal_logic_compact = ord_\_sound_counterex_reducing.clausal_logic_compact end text \ A second proof of Theorem 3.12, compactness of clausal logic: \ lemmas clausal_logic_compact = ground_resolution_with_selection.clausal_logic_compact end