diff --git a/thys/Ordered_Resolution_Prover/Lazy_List_Chain.thy b/thys/Ordered_Resolution_Prover/Lazy_List_Chain.thy --- a/thys/Ordered_Resolution_Prover/Lazy_List_Chain.thy +++ b/thys/Ordered_Resolution_Prover/Lazy_List_Chain.thy @@ -1,628 +1,630 @@ (* Title: Relational Chains over Lazy Lists Author: Jasmin Blanchette , 2014, 2017 Author: Dmitriy Traytel , 2017 Author: Anders Schlichtkrull , 2017 Maintainer: Anders Schlichtkrull *) section \Relational Chains over Lazy Lists\ theory Lazy_List_Chain - imports "HOL-Library.BNF_Corec" Lazy_List_Liminf + imports + "HOL-Library.BNF_Corec" + Lazy_List_Liminf begin text \ A chain is a lazy lists of elements such that all pairs of consecutive elements are related by a given relation. A full chain is either an infinite chain or a finite chain that cannot be extended. The inspiration for this theory is Section 4.1 (``Theorem Proving Processes'') of Bachmair and Ganzinger's chapter. \ subsection \Chains\ coinductive chain :: "('a \ 'a \ bool) \ 'a llist \ bool" for R :: "'a \ 'a \ bool" where chain_singleton: "chain R (LCons x LNil)" | chain_cons: "chain R xs \ R x (lhd xs) \ chain R (LCons x xs)" lemma chain_LNil[simp]: "\ chain R LNil" and chain_not_lnull: "chain R xs \ \ lnull xs" by (auto elim: chain.cases) lemma chain_lappend: assumes r_xs: "chain R xs" and r_ys: "chain R ys" and mid: "R (llast xs) (lhd ys)" shows "chain R (lappend xs ys)" proof (cases "lfinite xs") case True then show ?thesis using r_xs mid proof (induct rule: lfinite.induct) case (lfinite_LConsI xs x) note fin = this(1) and ih = this(2) and r_xxs = this(3) and mid = this(4) show ?case proof (cases "xs = LNil") case True then show ?thesis using r_ys mid by simp (rule chain_cons) next case xs_nnil: False have r_xs: "chain R xs" by (metis chain.simps ltl_simps(2) r_xxs xs_nnil) have mid': "R (llast xs) (lhd ys)" by (metis llast_LCons lnull_def mid xs_nnil) have start: "R x (lhd (lappend xs ys))" by (metis (no_types) chain.simps lhd_LCons lhd_lappend chain_not_lnull ltl_simps(2) r_xxs xs_nnil) show ?thesis unfolding lappend_code(2) using ih[OF r_xs mid'] start by (rule chain_cons) qed qed simp qed (simp add: r_xs lappend_inf) lemma chain_length_pos: "chain R xs \ llength xs > 0" by (cases xs) simp+ lemma chain_ldropn: assumes "chain R xs" and "enat n < llength xs" shows "chain R (ldropn n xs)" using assms by (induct n arbitrary: xs, simp, metis chain.cases ldrop_eSuc_ltl ldropn_LNil ldropn_eq_LNil ltl_simps(2) not_less) lemma inf_chain_ldropn_chain: "chain R xs \ \ lfinite xs \ chain R (ldropn n xs)" using chain.simps[of R xs] by (simp add: chain_ldropn not_lfinite_llength) lemma inf_chain_ltl_chain: "chain R xs \ \ lfinite xs \ chain R (ltl xs)" by (metis inf_chain_ldropn_chain ldropn_0 ldropn_ltl) lemma chain_lnth_rel: assumes chain: "chain R xs" and len: "enat (Suc j) < llength xs" shows "R (lnth xs j) (lnth xs (Suc j))" proof - define ys where "ys = ldropn j xs" have "llength ys > 1" unfolding ys_def using len by (metis One_nat_def funpow_swap1 ldropn_0 ldropn_def ldropn_eq_LNil ldropn_ltl not_less one_enat_def) obtain y0 y1 ys' where ys: "ys = LCons y0 (LCons y1 ys')" unfolding ys_def by (metis Suc_ile_eq ldropn_Suc_conv_ldropn len less_imp_not_less not_less) have "chain R ys" unfolding ys_def using Suc_ile_eq chain chain_ldropn len less_imp_le by blast then have "R y0 y1" unfolding ys by (auto elim: chain.cases) then show ?thesis using ys_def unfolding ys by (metis ldropn_Suc_conv_ldropn ldropn_eq_LConsD llist.inject) qed lemma infinite_chain_lnth_rel: assumes "\ lfinite c" and "chain r c" shows "r (lnth c i) (lnth c (Suc i))" using assms chain_lnth_rel lfinite_conv_llength_enat by force lemma lnth_rel_chain: assumes "\ lnull xs" and "\j. enat (j + 1) < llength xs \ R (lnth xs j) (lnth xs (j + 1))" shows "chain R xs" using assms proof (coinduction arbitrary: xs rule: chain.coinduct) case chain note nnul = this(1) and nth_chain = this(2) show ?case proof (cases "lnull (ltl xs)") case True have "xs = LCons (lhd xs) LNil" using nnul True by (simp add: llist.expand) then show ?thesis by blast next case nnul': False moreover have "xs = LCons (lhd xs) (ltl xs)" using nnul by simp moreover have "\j. enat (j + 1) < llength (ltl xs) \ R (lnth (ltl xs) j) (lnth (ltl xs) (j + 1))" using nnul nth_chain by (metis Suc_eq_plus1 ldrop_eSuc_ltl ldropn_Suc_conv_ldropn ldropn_eq_LConsD lnth_ltl) moreover have "R (lhd xs) (lhd (ltl xs))" using nnul' nnul nth_chain[rule_format, of 0, simplified] by (metis ldropn_0 ldropn_Suc_conv_ldropn ldropn_eq_LConsD lhd_LCons_ltl lhd_conv_lnth lnth_Suc_LCons ltl_simps(2)) ultimately show ?thesis by blast qed qed lemma chain_lmap: assumes "\x y. R x y \ R' (f x) (f y)" and "chain R xs" shows "chain R' (lmap f xs)" using assms proof (coinduction arbitrary: xs) case chain then have "(\y. xs = LCons y LNil) \ (\ys x. xs = LCons x ys \ chain R ys \ R x (lhd ys))" using chain.simps[of R xs] by auto then show ?case proof assume "\ys x. xs = LCons x ys \ chain R ys \ R x (lhd ys)" then have "\ys x. lmap f xs = LCons x ys \ (\xs. ys = lmap f xs \ (\x y. R x y \ R' (f x) (f y)) \ chain R xs) \ R' x (lhd ys)" using chain by (metis (no_types) lhd_LCons llist.distinct(1) llist.exhaust_sel llist.map_sel(1) lmap_eq_LNil chain_not_lnull ltl_lmap ltl_simps(2)) then show ?thesis by auto qed auto qed lemma chain_mono: assumes "\x y. R x y \ R' x y" and "chain R xs" shows "chain R' xs" using assms by (rule chain_lmap[of _ _ "\x. x", unfolded llist.map_ident]) lemma lfinite_chain_imp_rtranclp_lhd_llast: "lfinite xs \ chain R xs \ R\<^sup>*\<^sup>* (lhd xs) (llast xs)" proof (induct rule: lfinite.induct) case (lfinite_LConsI xs x) note fin_xs = this(1) and ih = this(2) and r_x_xs = this(3) show ?case proof (cases "xs = LNil") case xs_nnil: False then have r_xs: "chain R xs" using r_x_xs by (blast elim: chain.cases) then show ?thesis using ih[OF r_xs] xs_nnil r_x_xs by (metis chain.cases converse_rtranclp_into_rtranclp lhd_LCons llast_LCons chain_not_lnull ltl_simps(2)) qed simp qed simp lemma tranclp_imp_exists_finite_chain_list: "R\<^sup>+\<^sup>+ x y \ \xs. chain R (llist_of (x # xs @ [y]))" proof (induct rule: tranclp.induct) case (r_into_trancl x y) then have "chain R (llist_of (x # [] @ [y]))" by (auto intro: chain.intros) then show ?case by blast next case (trancl_into_trancl x y z) note rstar_xy = this(1) and ih = this(2) and r_yz = this(3) obtain xs where xs: "chain R (llist_of (x # xs @ [y]))" using ih by blast define ys where "ys = xs @ [y]" have "chain R (llist_of (x # ys @ [z]))" unfolding ys_def using r_yz chain_lappend[OF xs chain_singleton, of z] by (auto simp: lappend_llist_of_LCons llast_LCons) then show ?case by blast qed inductive_cases chain_consE: "chain R (LCons x xs)" inductive_cases chain_nontrivE: "chain R (LCons x (LCons y xs))" subsection \A Coinductive Puzzle\ primrec prepend where "prepend [] ys = ys" | "prepend (x # xs) ys = LCons x (prepend xs ys)" lemma lnull_prepend[simp]: "lnull (prepend xs ys) = (xs = [] \ lnull ys)" by (induct xs) auto lemma lhd_prepend[simp]: "lhd (prepend xs ys) = (if xs \ [] then hd xs else lhd ys)" by (induct xs) auto lemma prepend_LNil[simp]: "prepend xs LNil = llist_of xs" by (induct xs) auto lemma lfinite_prepend[simp]: "lfinite (prepend xs ys) \ lfinite ys" by (induct xs) auto lemma llength_prepend[simp]: "llength (prepend xs ys) = length xs + llength ys" by (induct xs) (auto simp: enat_0 iadd_Suc eSuc_enat[symmetric]) lemma llast_prepend[simp]: "\ lnull ys \ llast (prepend xs ys) = llast ys" by (induct xs) (auto simp: llast_LCons) lemma prepend_prepend: "prepend xs (prepend ys zs) = prepend (xs @ ys) zs" by (induct xs) auto lemma chain_prepend: "chain R (llist_of zs) \ last zs = lhd xs \ chain R xs \ chain R (prepend zs (ltl xs))" by (induct zs; cases xs) (auto split: if_splits simp: lnull_def[symmetric] intro!: chain_cons elim!: chain_consE) lemma lmap_prepend[simp]: "lmap f (prepend xs ys) = prepend (map f xs) (lmap f ys)" by (induct xs) auto lemma lset_prepend[simp]: "lset (prepend xs ys) = set xs \ lset ys" by (induct xs) auto lemma prepend_LCons: "prepend xs (LCons y ys) = prepend (xs @ [y]) ys" by (induct xs) auto lemma lnth_prepend: "lnth (prepend xs ys) i = (if i < length xs then nth xs i else lnth ys (i - length xs))" by (induct xs arbitrary: i) (auto simp: lnth_LCons' nth_Cons') theorem lfinite_less_induct[consumes 1, case_names less]: assumes fin: "lfinite xs" and step: "\xs. lfinite xs \ (\zs. llength zs < llength xs \ P zs) \ P xs" shows "P xs" using fin proof (induct "the_enat (llength xs)" arbitrary: xs rule: less_induct) case (less xs) show ?case using less(2) by (intro step[OF less(2)] less(1)) (auto dest!: lfinite_llength_enat simp: eSuc_enat elim!: less_enatE llength_eq_enat_lfiniteD) qed theorem lfinite_prepend_induct[consumes 1, case_names LNil prepend]: assumes "lfinite xs" and LNil: "P LNil" and prepend: "\xs. lfinite xs \ (\zs. (\ys. xs = prepend ys zs \ ys \ []) \ P zs) \ P xs" shows "P xs" using assms(1) proof (induct xs rule: lfinite_less_induct) case (less xs) from less(1) show ?case by (cases xs) (force simp: LNil neq_Nil_conv dest: lfinite_llength_enat intro!: prepend[of "LCons _ _"] intro: less)+ qed coinductive emb :: "'a llist \ 'a llist \ bool" where "lfinite xs \ emb LNil xs" | "emb xs ys \ emb (LCons x xs) (prepend zs (LCons x ys))" inductive_cases emb_LConsE: "emb (LCons z zs) ys" inductive_cases emb_LNil1E: "emb LNil ys" inductive_cases emb_LNil2E: "emb xs LNil" lemma emb_lfinite: assumes "emb xs ys" shows "lfinite ys \ lfinite xs" proof assume "lfinite xs" then show "lfinite ys" using assms by (induct xs arbitrary: ys rule: lfinite_induct) (auto simp: lnull_def neq_LNil_conv elim!: emb_LNil1E emb_LConsE) next assume "lfinite ys" then show "lfinite xs" using assms proof (induction ys arbitrary: xs rule: lfinite_less_induct) case (less ys) from less.prems \lfinite ys\ show ?case by (cases xs) (auto simp: eSuc_enat elim!: emb_LNil1E emb_LConsE less.IH[rotated] dest!: lfinite_llength_enat) qed qed inductive prepend_cong1 for X where prepend_cong1_base: "X xs \ prepend_cong1 X xs" | prepend_cong1_prepend: "prepend_cong1 X ys \ prepend_cong1 X (prepend xs ys)" lemma emb_prepend_coinduct[rotated, case_names emb]: assumes "(\x1 x2. X x1 x2 \ (\xs. x1 = LNil \ x2 = xs \ lfinite xs) \ (\xs ys x zs. x1 = LCons x xs \ x2 = prepend zs (LCons x ys) \ (prepend_cong1 (X xs) ys \ emb xs ys)))" (is "\x1 x2. X x1 x2 \ ?bisim x1 x2") shows "X x1 x2 \ emb x1 x2" proof (erule emb.coinduct[OF prepend_cong1_base]) fix xs zs assume "prepend_cong1 (X xs) zs" then show "?bisim xs zs" by (induct zs rule: prepend_cong1.induct) (erule assms, force simp: prepend_prepend) qed context begin private coinductive chain' for R where "chain' R (LCons x LNil)" | "chain R (llist_of (x # zs @ [lhd xs])) \ chain' R xs \ chain' R (LCons x (prepend zs xs))" private lemma chain_imp_chain': "chain R xs \ chain' R xs" proof (coinduction arbitrary: xs rule: chain'.coinduct) case chain' then show ?case proof (cases rule: chain.cases) case (chain_cons zs z) then show ?thesis by (intro disjI2 exI[of _ z] exI[of _ "[]"] exI[of _ "zs"]) (auto intro: chain.intros) qed simp qed private lemma chain'_imp_chain: "chain' R xs \ chain R xs" proof (coinduction arbitrary: xs rule: chain.coinduct) case chain then show ?case proof (cases rule: chain'.cases) case (2 y zs ys) then show ?thesis by (intro disjI2 exI[of _ "prepend zs ys"] exI[of _ y]) (force dest!: neq_Nil_conv[THEN iffD1] elim: chain.cases chain_nontrivE intro: chain'.intros) qed simp qed private lemma chain_chain': "chain = chain'" unfolding fun_eq_iff by (metis chain_imp_chain' chain'_imp_chain) lemma chain_prepend_coinduct[case_names chain]: "X x \ (\x. X x \ (\z. x = LCons z LNil) \ (\y xs zs. x = LCons y (prepend zs xs) \ (X xs \ chain R xs) \ chain R (llist_of (y # zs @ [lhd xs])))) \ chain R x" by (subst chain_chain', erule chain'.coinduct) (force simp: chain_chain') end context fixes R :: "'a \ 'a \ bool" begin private definition pick where "pick x y = (SOME xs. chain R (llist_of (x # xs @ [y])))" private lemma pick[simp]: assumes "R\<^sup>+\<^sup>+ x y" shows "chain R (llist_of (x # pick x y @ [y]))" unfolding pick_def using tranclp_imp_exists_finite_chain_list[THEN someI_ex, OF assms] by auto private friend_of_corec prepend where "prepend xs ys = (case xs of [] \ (case ys of LNil \ LNil | LCons x xs \ LCons x xs) | x # xs' \ LCons x (prepend xs' ys))" by (simp split: list.splits llist.splits) transfer_prover private corec wit where "wit xs = (case xs of LCons x (LCons y xs) \ LCons x (prepend (pick x y) (wit (LCons y xs))) | _ \ xs)" private lemma wit_LNil[simp]: "wit LNil = LNil" and wit_lsingleton[simp]: "wit (LCons x LNil) = LCons x LNil" and wit_LCons2: "wit (LCons x (LCons y xs)) = (LCons x (prepend (pick x y) (wit (LCons y xs))))" by (subst wit.code; auto)+ private lemma lnull_wit[simp]: "lnull (wit xs) \ lnull xs" by (subst wit.code) (auto split: llist.splits simp: Let_def) private lemma lhd_wit[simp]: "chain R\<^sup>+\<^sup>+ xs \ lhd (wit xs) = lhd xs" by (erule chain.cases; subst wit.code) (auto split: llist.splits simp: Let_def) private lemma LNil_eq_iff_lnull: "LNil = xs \ lnull xs" by (cases xs) auto lemma emb_wit[simp]: "chain R\<^sup>+\<^sup>+ xs \ emb xs (wit xs)" proof (coinduction arbitrary: xs rule: emb_prepend_coinduct) case (emb xs) then show ?case proof (cases rule: chain.cases) case (chain_cons zs z) then show ?thesis by (subst (2) wit.code) (auto split: llist.splits intro!: exI[of _ "[]"] exI[of _ "_ :: _ llist"] prepend_cong1_prepend[OF prepend_cong1_base]) qed (auto intro!: exI[of _ LNil] exI[of _ "[]"] emb.intros) qed private lemma lfinite_wit[simp]: assumes "chain R\<^sup>+\<^sup>+ xs" shows "lfinite (wit xs) \ lfinite xs" using emb_wit emb_lfinite assms by blast private lemma llast_wit[simp]: assumes "chain R\<^sup>+\<^sup>+ xs" shows "llast (wit xs) = llast xs" proof (cases "lfinite xs") case True from this assms show ?thesis proof (induct rule: lfinite.induct) case (lfinite_LConsI xs x) then show ?case by (cases xs) (auto simp: wit_LCons2 llast_LCons elim: chain_nontrivE) qed auto qed (auto simp: llast_linfinite assms) lemma chain_tranclp_imp_exists_chain: "chain R\<^sup>+\<^sup>+ xs \ \ys. chain R ys \ emb xs ys \ lhd ys = lhd xs \ llast ys = llast xs" proof (intro exI[of _ "wit xs"] conjI, coinduction arbitrary: xs rule: chain_prepend_coinduct) case chain then show ?case by (subst (1 2) wit.code) (erule chain.cases; force split: llist.splits dest: pick) qed auto lemma emb_lset_mono[rotated]: "x \ lset xs \ emb xs ys \ x \ lset ys" by (induct x xs arbitrary: ys rule: llist.set_induct) (auto elim!: emb_LConsE) lemma emb_Ball_lset_antimono: assumes "emb Xs Ys" shows "\Y \ lset Ys. x \ Y \ \X \ lset Xs. x \ X" using emb_lset_mono[OF assms] by blast lemma emb_lfinite_antimono[rotated]: "lfinite ys \ emb xs ys \ lfinite xs" by (induct ys arbitrary: xs rule: lfinite_prepend_induct) (force elim!: emb_LNil2E simp: LNil_eq_iff_lnull prepend_LCons elim: emb.cases)+ lemma emb_Liminf_llist_mono_aux: assumes "emb Xs Ys" and "\ lfinite Xs" and "\ lfinite Ys" and "\j\i. x \ lnth Ys j" shows "\j\i. x \ lnth Xs j" using assms proof (induct i arbitrary: Xs Ys rule: less_induct) case (less i) then show ?case proof (cases i) case 0 then show ?thesis using emb_Ball_lset_antimono[OF less(2), of x] less(5) unfolding Ball_def in_lset_conv_lnth simp_thms not_lfinite_llength[OF less(3)] not_lfinite_llength[OF less(4)] enat_ord_code subset_eq by blast next case [simp]: (Suc nat) from less(2,3) obtain xs as b bs where [simp]: "Xs = LCons b xs" "Ys = prepend as (LCons b bs)" and "emb xs bs" by (auto elim: emb.cases) have IH: "\k\j. x \ lnth xs k" if "\k\j. x \ lnth bs k" "j < i" for j using that less(1)[OF _ \emb xs bs\] less(3,4) by auto from less(5) have "\k\i - length as - 1. x \ lnth xs k" by (intro IH allI) (drule spec[of _ "_ + length as + 1"], auto simp: lnth_prepend lnth_LCons') then show ?thesis by (auto simp: lnth_LCons') qed qed lemma emb_Liminf_llist_infinite: assumes "emb Xs Ys" and "\ lfinite Xs" shows "Liminf_llist Ys \ Liminf_llist Xs" proof - from assms have "\ lfinite Ys" using emb_lfinite_antimono by blast with assms show ?thesis unfolding Liminf_llist_def by (auto simp: not_lfinite_llength dest: emb_Liminf_llist_mono_aux) qed lemma emb_lmap: "emb xs ys \ emb (lmap f xs) (lmap f ys)" proof (coinduction arbitrary: xs ys rule: emb.coinduct) case emb show ?case proof (cases xs) case xs: (LCons x xs') obtain ysa0 and zs0 where ys: "ys = prepend zs0 (LCons x ysa0)" and emb': "emb xs' ysa0" using emb_LConsE[OF emb[unfolded xs]] by metis let ?xa = "f x" let ?xsa = "lmap f xs'" let ?zs = "map f zs0" let ?ysa = "lmap f ysa0" have "lmap f xs = LCons ?xa ?xsa" unfolding xs by simp moreover have "lmap f ys = prepend ?zs (LCons ?xa ?ysa)" unfolding ys by simp moreover have "\xsa ysa. ?xsa = lmap f xsa \ ?ysa = lmap f ysa \ emb xsa ysa" using emb' by blast ultimately show ?thesis by blast qed (simp add: emb_lfinite[OF emb]) qed end lemma chain_inf_llist_if_infinite_chain_function: assumes "\i. r (f (Suc i)) (f i)" shows "\ lfinite (inf_llist f) \ chain r\\ (inf_llist f)" using assms by (simp add: lnth_rel_chain) lemma infinite_chain_function_iff_infinite_chain_llist: "(\f. \i. r (f (Suc i)) (f i)) \ (\c. \ lfinite c \ chain r\\ c)" using chain_inf_llist_if_infinite_chain_function infinite_chain_lnth_rel by blast lemma wfP_iff_no_infinite_down_chain_llist: "wfP r \ (\c. \ lfinite c \ chain r\\ c)" proof - have "wfP r \ wf {(x, y). r x y}" unfolding wfP_def by auto also have "\ \ (\f. \i. (f (Suc i), f i) \ {(x, y). r x y})" using wf_iff_no_infinite_down_chain by blast also have "\ \ (\f. \i. r (f (Suc i)) (f i))" by auto also have "\ \ (\c. \lfinite c \ chain r\\ c)" using infinite_chain_function_iff_infinite_chain_llist by blast finally show ?thesis by auto qed subsection \Full Chains\ coinductive full_chain :: "('a \ 'a \ bool) \ 'a llist \ bool" for R :: "'a \ 'a \ bool" where full_chain_singleton: "(\y. \ R x y) \ full_chain R (LCons x LNil)" | full_chain_cons: "full_chain R xs \ R x (lhd xs) \ full_chain R (LCons x xs)" lemma full_chain_LNil[simp]: "\ full_chain R LNil" and full_chain_not_lnull: "full_chain R xs \ \ lnull xs" by (auto elim: full_chain.cases) lemma full_chain_ldropn: assumes full: "full_chain R xs" and "enat n < llength xs" shows "full_chain R (ldropn n xs)" using assms by (induct n arbitrary: xs, simp, metis full_chain.cases ldrop_eSuc_ltl ldropn_LNil ldropn_eq_LNil ltl_simps(2) not_less) lemma full_chain_iff_chain: "full_chain R xs \ chain R xs \ (lfinite xs \ (\y. \ R (llast xs) y))" proof (intro iffI conjI impI allI; (elim conjE)?) assume full: "full_chain R xs" show chain: "chain R xs" using full by (coinduction arbitrary: xs) (auto elim: full_chain.cases) { fix y assume "lfinite xs" then obtain n where suc_n: "Suc n = llength xs" by (metis chain chain_length_pos lessE less_enatE lfinite_conv_llength_enat) have "full_chain R (ldropn n xs)" by (rule full_chain_ldropn[OF full]) (use suc_n Suc_ile_eq in force) moreover have "ldropn n xs = LCons (llast xs) LNil" using suc_n by (metis enat_le_plus_same(2) enat_ord_simps(2) gen_llength_def ldropn_Suc_conv_ldropn ldropn_all lessI llast_ldropn llast_singleton llength_code) ultimately show "\ R (llast xs) y" by (auto elim: full_chain.cases) } next assume "chain R xs" and "lfinite xs \ (\y. \ R (llast xs) y)" then show "full_chain R xs" by (coinduction arbitrary: xs) (erule chain.cases, simp, metis lfinite_LConsI llast_LCons) qed lemma full_chain_imp_chain: "full_chain R xs \ chain R xs" using full_chain_iff_chain by blast lemma full_chain_length_pos: "full_chain R xs \ llength xs > 0" by (fact chain_length_pos[OF full_chain_imp_chain]) lemma full_chain_lnth_rel: "full_chain R xs \ enat (Suc j) < llength xs \ R (lnth xs j) (lnth xs (Suc j))" by (fact chain_lnth_rel[OF full_chain_imp_chain]) inductive_cases full_chain_consE: "full_chain R (LCons x xs)" inductive_cases full_chain_nontrivE: "full_chain R (LCons x (LCons y xs))" lemma full_chain_tranclp_imp_exists_full_chain: assumes full: "full_chain R\<^sup>+\<^sup>+ xs" shows "\ys. full_chain R ys \ emb xs ys \ lhd ys = lhd xs \ llast ys = llast xs" proof - obtain ys where ys: "chain R ys" "emb xs ys" "lhd ys = lhd xs" "llast ys = llast xs" using full_chain_imp_chain[OF full] chain_tranclp_imp_exists_chain by blast have "full_chain R ys" using ys(1,4) emb_lfinite[OF ys(2)] full unfolding full_chain_iff_chain by auto then show ?thesis using ys(2-4) by auto qed end diff --git a/thys/Ordered_Resolution_Prover/Lazy_List_Liminf.thy b/thys/Ordered_Resolution_Prover/Lazy_List_Liminf.thy --- a/thys/Ordered_Resolution_Prover/Lazy_List_Liminf.thy +++ b/thys/Ordered_Resolution_Prover/Lazy_List_Liminf.thy @@ -1,381 +1,389 @@ (* Title: Supremum and Liminf of Lazy Lists Author: Jasmin Blanchette , 2014, 2017, 2020 Author: Dmitriy Traytel , 2014 Maintainer: Jasmin Blanchette *) section \Supremum and Liminf of Lazy Lists\ theory Lazy_List_Liminf imports Coinductive.Coinductive_List begin text \ Lazy lists, as defined in the \emph{Archive of Formal Proofs}, provide finite and infinite lists in one type, defined coinductively. The present theory introduces the concept of the union of all elements of a lazy list of sets and the limit of such a lazy list. The definitions are stated more generally in terms of lattices. The basis for this theory is Section 4.1 (``Theorem Proving Processes'') of Bachmair and Ganzinger's chapter. \ subsection \Library\ lemma less_llength_ltake: "i < llength (ltake k Xs) \ i < k \ i < llength Xs" by simp subsection \Supremum\ definition Sup_llist :: "'a set llist \ 'a set" where "Sup_llist Xs = (\i \ {i. enat i < llength Xs}. lnth Xs i)" lemma lnth_subset_Sup_llist: "enat i < llength Xs \ lnth Xs i \ Sup_llist Xs" unfolding Sup_llist_def by auto lemma Sup_llist_imp_exists_index: "x \ Sup_llist Xs \ \i. enat i < llength Xs \ x \ lnth Xs i" unfolding Sup_llist_def by auto lemma exists_index_imp_Sup_llist: "enat i < llength Xs \ x \ lnth Xs i \ x \ Sup_llist Xs" unfolding Sup_llist_def by auto lemma Sup_llist_LNil[simp]: "Sup_llist LNil = {}" unfolding Sup_llist_def by auto lemma Sup_llist_LCons[simp]: "Sup_llist (LCons X Xs) = X \ Sup_llist Xs" unfolding Sup_llist_def proof (intro subset_antisym subsetI) fix x assume "x \ (\i \ {i. enat i < llength (LCons X Xs)}. lnth (LCons X Xs) i)" then obtain i where len: "enat i < llength (LCons X Xs)" and nth: "x \ lnth (LCons X Xs) i" by blast from nth have "x \ X \ i > 0 \ x \ lnth Xs (i - 1)" by (metis lnth_LCons' neq0_conv) then have "x \ X \ (\i. enat i < llength Xs \ x \ lnth Xs i)" by (metis len Suc_pred' eSuc_enat iless_Suc_eq less_irrefl llength_LCons not_less order_trans) then show "x \ X \ (\i \ {i. enat i < llength Xs}. lnth Xs i)" by blast qed ((auto)[], metis i0_lb lnth_0 zero_enat_def, metis Suc_ile_eq lnth_Suc_LCons) lemma lhd_subset_Sup_llist: "\ lnull Xs \ lhd Xs \ Sup_llist Xs" by (cases Xs) simp_all subsection \Supremum up-to\ definition Sup_upto_llist :: "'a set llist \ enat \ 'a set" where "Sup_upto_llist Xs j = (\i \ {i. enat i < llength Xs \ enat i \ j}. lnth Xs i)" lemma Sup_upto_llist_eq_Sup_llist_ltake: "Sup_upto_llist Xs j = Sup_llist (ltake (eSuc j) Xs)" unfolding Sup_upto_llist_def Sup_llist_def by (smt Collect_cong Sup.SUP_cong iless_Suc_eq lnth_ltake less_llength_ltake mem_Collect_eq) lemma Sup_upto_llist_enat_0[simp]: - "Sup_upto_llist Xs (enat 0) = (if 0 < llength Xs then lnth Xs 0 else {})" - unfolding Sup_upto_llist_def image_def by (simp add: enat_0[symmetric]) + "Sup_upto_llist Xs (enat 0) = (if lnull Xs then {} else lhd Xs)" +proof (cases "lnull Xs") + case True + then show ?thesis + unfolding Sup_upto_llist_def by auto +next + case False + show ?thesis + unfolding Sup_upto_llist_def image_def by (simp add: lhd_conv_lnth enat_0 enat_0_iff) +qed lemma Sup_upto_llist_Suc[simp]: "Sup_upto_llist Xs (enat (Suc j)) = Sup_upto_llist Xs (enat j) \ (if enat (Suc j) < llength Xs then lnth Xs (Suc j) else {})" unfolding Sup_upto_llist_def image_def by (auto intro: le_SucI elim: le_SucE) lemma Sup_upto_llist_infinity[simp]: "Sup_upto_llist Xs \ = Sup_llist Xs" unfolding Sup_upto_llist_def Sup_llist_def by simp -lemma Sup_upto_llist_0[simp]: - "Sup_upto_llist Xs 0 = (if 0 < llength Xs then lnth Xs 0 else {})" - unfolding Sup_upto_llist_def image_def by (simp add: enat_0[symmetric]) +lemma Sup_upto_llist_0[simp]: "Sup_upto_llist Xs 0 = (if lnull Xs then {} else lhd Xs)" + unfolding zero_enat_def by (rule Sup_upto_llist_enat_0) lemma Sup_upto_llist_eSuc[simp]: "Sup_upto_llist Xs (eSuc j) = (case j of enat k \ Sup_upto_llist Xs (enat (Suc k)) | \ \ Sup_llist Xs)" by (auto simp: eSuc_enat split: enat.split) lemma Sup_upto_llist_mono[simp]: "j \ k \ Sup_upto_llist Xs j \ Sup_upto_llist Xs k" unfolding Sup_upto_llist_def by auto lemma Sup_upto_llist_subset_Sup_llist: "Sup_upto_llist Xs j \ Sup_llist Xs" unfolding Sup_llist_def Sup_upto_llist_def by auto lemma elem_Sup_llist_imp_Sup_upto_llist: "x \ Sup_llist Xs \ \j < llength Xs. x \ Sup_upto_llist Xs (enat j)" unfolding Sup_llist_def Sup_upto_llist_def by blast lemma lnth_subset_Sup_upto_llist: "j < llength Xs \ lnth Xs j \ Sup_upto_llist Xs j" unfolding Sup_upto_llist_def by auto lemma finite_Sup_llist_imp_Sup_upto_llist: assumes "finite X" and "X \ Sup_llist Xs" shows "\k. X \ Sup_upto_llist Xs (enat k)" using assms proof induct case (insert x X) then have x: "x \ Sup_llist Xs" and X: "X \ Sup_llist Xs" by simp+ from x obtain k where k: "x \ Sup_upto_llist Xs (enat k)" using elem_Sup_llist_imp_Sup_upto_llist by fast from X obtain k' where k': "X \ Sup_upto_llist Xs (enat k')" using insert.hyps(3) by fast have "insert x X \ Sup_upto_llist Xs (max k k')" using k k' by (metis (mono_tags) Sup_upto_llist_mono enat_ord_simps(1) insert_subset max.cobounded1 max.cobounded2 subset_iff) then show ?case by fast qed simp subsection \Liminf\ definition Liminf_llist :: "'a set llist \ 'a set" where "Liminf_llist Xs = (\i \ {i. enat i < llength Xs}. \j \ {j. i \ j \ enat j < llength Xs}. lnth Xs j)" lemma Liminf_llist_LNil[simp]: "Liminf_llist LNil = {}" unfolding Liminf_llist_def by simp lemma Liminf_llist_LCons: "Liminf_llist (LCons X Xs) = (if lnull Xs then X else Liminf_llist Xs)" (is "?lhs = ?rhs") proof (cases "lnull Xs") case nnull: False show ?thesis proof { fix x assume "\i. enat i \ llength Xs \ (\j. i \ j \ enat j \ llength Xs \ x \ lnth (LCons X Xs) j)" then have "\i. enat (Suc i) \ llength Xs \ (\j. Suc i \ j \ enat j \ llength Xs \ x \ lnth (LCons X Xs) j)" by (cases "llength Xs", metis not_lnull_conv[THEN iffD1, OF nnull] Suc_le_D eSuc_enat eSuc_ile_mono llength_LCons not_less_eq_eq zero_enat_def zero_le, metis Suc_leD enat_ord_code(3)) then have "\i. enat i < llength Xs \ (\j. i \ j \ enat j < llength Xs \ x \ lnth Xs j)" by (metis Suc_ile_eq Suc_n_not_le_n lift_Suc_mono_le lnth_Suc_LCons nat_le_linear) } then show "?lhs \ ?rhs" by (simp add: Liminf_llist_def nnull) (rule subsetI, simp) { fix x assume "\i. enat i < llength Xs \ (\j. i \ j \ enat j < llength Xs \ x \ lnth Xs j)" then obtain i where i: "enat i < llength Xs" and j: "\j. i \ j \ enat j < llength Xs \ x \ lnth Xs j" by blast have "enat (Suc i) \ llength Xs" using i by (simp add: Suc_ile_eq) moreover have "\j. Suc i \ j \ enat j \ llength Xs \ x \ lnth (LCons X Xs) j" using Suc_ile_eq Suc_le_D j by force ultimately have "\i. enat i \ llength Xs \ (\j. i \ j \ enat j \ llength Xs \ x \ lnth (LCons X Xs) j)" by blast } then show "?rhs \ ?lhs" by (simp add: Liminf_llist_def nnull) (rule subsetI, simp) qed qed (simp add: Liminf_llist_def enat_0_iff(1)) lemma lfinite_Liminf_llist: "lfinite Xs \ Liminf_llist Xs = (if lnull Xs then {} else llast Xs)" proof (induction rule: lfinite_induct) case (LCons xs) then obtain y ys where xs: "xs = LCons y ys" by (meson not_lnull_conv) show ?case unfolding xs by (simp add: Liminf_llist_LCons LCons.IH[unfolded xs, simplified] llast_LCons) qed (simp add: Liminf_llist_def) lemma Liminf_llist_ltl: "\ lnull (ltl Xs) \ Liminf_llist Xs = Liminf_llist (ltl Xs)" by (metis Liminf_llist_LCons lhd_LCons_ltl lnull_ltlI) lemma Liminf_llist_subset_Sup_llist: "Liminf_llist Xs \ Sup_llist Xs" unfolding Liminf_llist_def Sup_llist_def by fast lemma image_Liminf_llist_subset: "f ` Liminf_llist Ns \ Liminf_llist (lmap ((`) f) Ns)" unfolding Liminf_llist_def by auto lemma Liminf_llist_imp_exists_index: "x \ Liminf_llist Xs \ \i. enat i < llength Xs \ x \ lnth Xs i" unfolding Liminf_llist_def by auto lemma not_Liminf_llist_imp_exists_index: "\ lnull Xs \ x \ Liminf_llist Xs \ enat i < llength Xs \ (\j. i \ j \ enat j < llength Xs \ x \ lnth Xs j)" unfolding Liminf_llist_def by auto lemma finite_subset_Liminf_llist_imp_exists_index: assumes nnil: "\ lnull Xs" and fin: "finite X" and in_lim: "X \ Liminf_llist Xs" shows "\i. enat i < llength Xs \ X \ (\j \ {j. i \ j \ enat j < llength Xs}. lnth Xs j)" proof - show ?thesis proof (cases "X = {}") case True then show ?thesis using nnil by (auto intro: exI[of _ 0] simp: zero_enat_def[symmetric]) next case nemp: False have in_lim': "\x \ X. \i. enat i < llength Xs \ x \ (\j \ {j. i \ j \ enat j < llength Xs}. lnth Xs j)" using in_lim[unfolded Liminf_llist_def] in_mono by fastforce obtain i_of where i_of_lt: "\x \ X. enat (i_of x) < llength Xs" and in_inter: "\x \ X. x \ (\j \ {j. i_of x \ j \ enat j < llength Xs}. lnth Xs j)" using bchoice[OF in_lim'] by blast define i_max where "i_max = Max (i_of ` X)" have "i_max \ i_of ` X" by (simp add: fin i_max_def nemp) then obtain x_max where x_max_in: "x_max \ X" and i_max_is: "i_max = i_of x_max" unfolding i_max_def by blast have le_i_max: "\x \ X. i_of x \ i_max" unfolding i_max_def by (simp add: fin) have "enat i_max < llength Xs" using i_of_lt x_max_in i_max_is by auto moreover have "X \ (\j \ {j. i_max \ j \ enat j < llength Xs}. lnth Xs j)" proof fix x assume x_in: "x \ X" then have x_in_inter: "x \ (\j \ {j. i_of x \ j \ enat j < llength Xs}. lnth Xs j)" using in_inter by auto moreover have "{j. i_max \ j \ enat j < llength Xs} \ {j. i_of x \ j \ enat j < llength Xs}" using x_in le_i_max by auto ultimately show "x \ (\j \ {j. i_max \ j \ enat j < llength Xs}. lnth Xs j)" by auto qed ultimately show ?thesis by auto qed qed lemma Liminf_llist_lmap_image: assumes f_inj: "inj_on f (Sup_llist (lmap g xs))" shows "Liminf_llist (lmap (\x. f ` g x) xs) = f ` Liminf_llist (lmap g xs)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" proof fix x assume "x \ Liminf_llist (lmap (\x. f ` g x) xs)" then obtain i where i_lt: "enat i < llength xs" and x_in_fgj: "\j. i \ j \ enat j < llength xs \ x \ f ` g (lnth xs j)" unfolding Liminf_llist_def by auto have ex_in_gi: "\y. y \ g (lnth xs i) \ x = f y" using f_inj i_lt x_in_fgj unfolding inj_on_def Sup_llist_def by auto have "\y. \j. i \ j \ enat j < llength xs \ y \ g (lnth xs j) \ x = f y" apply (rule exI[of _ "SOME y. y \ g (lnth xs i) \ x = f y"]) using someI_ex[OF ex_in_gi] x_in_fgj f_inj i_lt x_in_fgj unfolding inj_on_def Sup_llist_def by simp (metis (no_types, lifting) imageE) then show "x \ f ` Liminf_llist (lmap g xs)" using i_lt unfolding Liminf_llist_def by auto qed next show "?rhs \ ?lhs" using image_Liminf_llist_subset[of f "lmap g xs", unfolded llist.map_comp] by auto qed lemma Liminf_llist_lmap_union: assumes "\x \ lset xs. \Y \ lset xs. g x \ h Y = {}" shows "Liminf_llist (lmap (\x. g x \ h x) xs) = Liminf_llist (lmap g xs) \ Liminf_llist (lmap h xs)" (is "?lhs = ?rhs") proof (intro equalityI subsetI) fix x assume x_in: "x \ ?lhs" then obtain i where i_lt: "enat i < llength xs" and j: "\j. i \ j \ enat j < llength xs \ x \ g (lnth xs j) \ x \ h (lnth xs j)" using x_in[unfolded Liminf_llist_def, simplified] by blast then have "(\i'. enat i' < llength xs \ (\j. i' \ j \ enat j < llength xs \ x \ g (lnth xs j))) \ (\i'. enat i' < llength xs \ (\j. i' \ j \ enat j < llength xs \ x \ h (lnth xs j)))" using assms[unfolded disjoint_iff_not_equal] by (metis in_lset_conv_lnth) then show "x \ ?rhs" unfolding Liminf_llist_def by simp next fix x show "x \ ?rhs \ x \ ?lhs" using assms unfolding Liminf_llist_def by auto qed lemma Liminf_set_filter_commute: "Liminf_llist (lmap (\X. {x \ X. p x}) Xs) = {x \ Liminf_llist Xs. p x}" unfolding Liminf_llist_def by force subsection \Liminf up-to\ definition Liminf_upto_llist :: "'a set llist \ enat \ 'a set" where "Liminf_upto_llist Xs k = (\i \ {i. enat i < llength Xs \ enat i \ k}. \j \ {j. i \ j \ enat j < llength Xs \ enat j \ k}. lnth Xs j)" lemma Liminf_upto_llist_eq_Liminf_llist_ltake: "Liminf_upto_llist Xs j = Liminf_llist (ltake (eSuc j) Xs)" unfolding Liminf_upto_llist_def Liminf_llist_def by (smt Collect_cong Sup.SUP_cong iless_Suc_eq lnth_ltake less_llength_ltake mem_Collect_eq) lemma Liminf_upto_llist_enat[simp]: "Liminf_upto_llist Xs (enat k) = (if enat k < llength Xs then lnth Xs k else if lnull Xs then {} else llast Xs)" proof (cases "enat k < llength Xs") case True then show ?thesis unfolding Liminf_upto_llist_def by force next case k_ge: False show ?thesis proof (cases "lnull Xs") case nil: True then show ?thesis unfolding Liminf_upto_llist_def by simp next case nnil: False then obtain j where j: "eSuc (enat j) = llength Xs" using k_ge by (metis eSuc_enat_iff enat_ile le_less_linear lhd_LCons_ltl llength_LCons) have fin: "lfinite Xs" using k_ge not_lfinite_llength by fastforce have le_k: "enat i < llength Xs \ i \ k \ enat i < llength Xs" for i using k_ge linear order_le_less_subst2 by fastforce have "Liminf_upto_llist Xs (enat k) = llast Xs" using j nnil lfinite_Liminf_llist[OF fin] nnil unfolding Liminf_upto_llist_def Liminf_llist_def using llast_conv_lnth[OF j[symmetric]] by (simp add: le_k) then show ?thesis using k_ge nnil by simp qed qed lemma Liminf_upto_llist_infinity[simp]: "Liminf_upto_llist Xs \ = Liminf_llist Xs" unfolding Liminf_upto_llist_def Liminf_llist_def by simp lemma Liminf_upto_llist_0[simp]: - "Liminf_upto_llist Xs 0 = (if lnull Xs then {} else lnth Xs 0)" - unfolding Liminf_upto_llist_def image_def by (simp add: enat_0[symmetric]) (simp add: enat_0) + "Liminf_upto_llist Xs 0 = (if lnull Xs then {} else lhd Xs)" + unfolding Liminf_upto_llist_def image_def + by (simp add: enat_0[symmetric]) (simp add: enat_0 lnth_0_conv_lhd) lemma Liminf_upto_llist_eSuc[simp]: "Liminf_upto_llist Xs (eSuc j) = (case j of enat k \ Liminf_upto_llist Xs (enat (Suc k)) | \ \ Liminf_llist Xs)" by (auto simp: eSuc_enat split: enat.split) lemma elem_Liminf_llist_imp_Liminf_upto_llist: "x \ Liminf_llist Xs \ \i < llength Xs. \j. i \ j \ j < llength Xs \ x \ Liminf_upto_llist Xs (enat j)" unfolding Liminf_llist_def Liminf_upto_llist_def using enat_ord_simps(1) by force end diff --git a/thys/Ordered_Resolution_Prover/Proving_Process.thy b/thys/Ordered_Resolution_Prover/Proving_Process.thy --- a/thys/Ordered_Resolution_Prover/Proving_Process.thy +++ b/thys/Ordered_Resolution_Prover/Proving_Process.thy @@ -1,336 +1,336 @@ (* Title: Theorem Proving Processes Author: Jasmin Blanchette , 2014, 2017 Author: Dmitriy Traytel , 2014 Author: Anders Schlichtkrull , 2017 Maintainer: Anders Schlichtkrull *) section \Theorem Proving Processes\ theory Proving_Process imports Unordered_Ground_Resolution Lazy_List_Chain begin text \ This material corresponds to Section 4.1 (``Theorem Proving Processes'') of Bachmair and Ganzinger's chapter. The locale assumptions below capture conditions R1 to R3 of Definition 4.1. \Rf\ denotes $\mathcal{R}_{\mathcal{F}}$; \Ri\ denotes $\mathcal{R}_{\mathcal{I}}$. \ locale redundancy_criterion = inference_system + fixes Rf :: "'a clause set \ 'a clause set" and Ri :: "'a clause set \ 'a inference set" assumes Ri_subset_\: "Ri N \ \" and Rf_mono: "N \ N' \ Rf N \ Rf N'" and Ri_mono: "N \ N' \ Ri N \ Ri N'" and Rf_indep: "N' \ Rf N \ Rf N \ Rf (N - N')" and Ri_indep: "N' \ Rf N \ Ri N \ Ri (N - N')" and Rf_sat: "satisfiable (N - Rf N) \ satisfiable N" begin definition saturated_upto :: "'a clause set \ bool" where "saturated_upto N \ inferences_from (N - Rf N) \ Ri N" inductive "derive" :: "'a clause set \ 'a clause set \ bool" (infix "\" 50) where deduction_deletion: "N - M \ concls_of (inferences_from M) \ M - N \ Rf N \ M \ N" lemma derive_subset: "M \ N \ N \ M \ concls_of (inferences_from M)" by (meson Diff_subset_conv derive.cases) end locale sat_preserving_redundancy_criterion = sat_preserving_inference_system "\ :: ('a :: wellorder) inference set" + redundancy_criterion begin lemma deriv_sat_preserving: assumes deriv: "chain (\) Ns" and sat_n0: "satisfiable (lhd Ns)" shows "satisfiable (Sup_llist Ns)" proof - - have ns0: "lnth Ns 0 = lhd Ns" - using deriv by (metis chain_not_lnull lhd_conv_lnth) have len_ns: "llength Ns > 0" using deriv by (case_tac Ns) simp+ { fix DD assume fin: "finite DD" and sset_lun: "DD \ Sup_llist Ns" then obtain k where dd_sset: "DD \ Sup_upto_llist Ns (enat k)" using finite_Sup_llist_imp_Sup_upto_llist by blast have "satisfiable (Sup_upto_llist Ns k)" proof (induct k) case 0 then show ?case - using len_ns ns0 sat_n0 unfolding Sup_upto_llist_def true_clss_def by auto + using len_ns sat_n0 + unfolding Sup_upto_llist_def true_clss_def lhd_conv_lnth[OF chain_not_lnull[OF deriv]] + by auto next case (Suc k) show ?case proof (cases "enat (Suc k) \ llength Ns") case True then have "Sup_upto_llist Ns (enat k) = Sup_upto_llist Ns (enat (Suc k))" unfolding Sup_upto_llist_def using le_Suc_eq by auto then show ?thesis using Suc by simp next case False then have "lnth Ns k \ lnth Ns (Suc k)" using deriv by (auto simp: chain_lnth_rel) then have "lnth Ns (Suc k) \ lnth Ns k \ concls_of (inferences_from (lnth Ns k))" by (rule derive_subset) moreover have "lnth Ns k \ Sup_upto_llist Ns k" unfolding Sup_upto_llist_def using False Suc_ile_eq linear by blast ultimately have "lnth Ns (Suc k) \ Sup_upto_llist Ns k \ concls_of (inferences_from (Sup_upto_llist Ns k))" by clarsimp (metis UnCI UnE image_Un inferences_from_mono le_iff_sup) moreover have "Sup_upto_llist Ns (Suc k) = Sup_upto_llist Ns k \ lnth Ns (Suc k)" unfolding Sup_upto_llist_def using False by (force elim: le_SucE) moreover have "satisfiable (Sup_upto_llist Ns k \ concls_of (inferences_from (Sup_upto_llist Ns k)))" using Suc \_sat_preserving unfolding sat_preserving_inference_system_def by simp ultimately show ?thesis by (metis le_iff_sup true_clss_union) qed qed then have "satisfiable DD" using dd_sset unfolding Sup_upto_llist_def by (blast intro: true_clss_mono) } then show ?thesis using ground_resolution_without_selection.clausal_logic_compact[THEN iffD1] by metis qed text \ This corresponds to Lemma 4.2: \ lemma assumes deriv: "chain (\) Ns" shows Rf_Sup_subset_Rf_Liminf: "Rf (Sup_llist Ns) \ Rf (Liminf_llist Ns)" and Ri_Sup_subset_Ri_Liminf: "Ri (Sup_llist Ns) \ Ri (Liminf_llist Ns)" and sat_limit_iff: "satisfiable (Liminf_llist Ns) \ satisfiable (lhd Ns)" proof - { fix C i j assume c_in: "C \ lnth Ns i" and c_ni: "C \ Rf (Sup_llist Ns)" and j: "j \ i" and j': "enat j < llength Ns" from c_ni have c_ni': "\i. enat i < llength Ns \ C \ Rf (lnth Ns i)" using Rf_mono lnth_subset_Sup_llist Sup_llist_def by (blast dest: contra_subsetD) have "C \ lnth Ns j" using j j' proof (induct j) case 0 then show ?case using c_in by blast next case (Suc k) then show ?case proof (cases "i < Suc k") case True have "i \ k" using True by linarith moreover have "enat k < llength Ns" using Suc.prems(2) Suc_ile_eq by (blast intro: dual_order.strict_implies_order) ultimately have c_in_k: "C \ lnth Ns k" using Suc.hyps by blast have rel: "lnth Ns k \ lnth Ns (Suc k)" using Suc.prems deriv by (auto simp: chain_lnth_rel) then show ?thesis using c_in_k c_ni' Suc.prems(2) by cases auto next case False then show ?thesis using Suc c_in by auto qed qed } then have lu_ll: "Sup_llist Ns - Rf (Sup_llist Ns) \ Liminf_llist Ns" unfolding Sup_llist_def Liminf_llist_def by blast have rf: "Rf (Sup_llist Ns - Rf (Sup_llist Ns)) \ Rf (Liminf_llist Ns)" using lu_ll Rf_mono by simp have ri: "Ri (Sup_llist Ns - Rf (Sup_llist Ns)) \ Ri (Liminf_llist Ns)" using lu_ll Ri_mono by simp show "Rf (Sup_llist Ns) \ Rf (Liminf_llist Ns)" using rf Rf_indep by blast show "Ri (Sup_llist Ns) \ Ri (Liminf_llist Ns)" using ri Ri_indep by blast show "satisfiable (Liminf_llist Ns) \ satisfiable (lhd Ns)" proof assume "satisfiable (lhd Ns)" then have "satisfiable (Sup_llist Ns)" using deriv deriv_sat_preserving by simp then show "satisfiable (Liminf_llist Ns)" using true_clss_mono[OF Liminf_llist_subset_Sup_llist] by blast next assume "satisfiable (Liminf_llist Ns)" then have "satisfiable (Sup_llist Ns - Rf (Sup_llist Ns))" using true_clss_mono[OF lu_ll] by blast then have "satisfiable (Sup_llist Ns)" using Rf_sat by blast then show "satisfiable (lhd Ns)" using deriv true_clss_mono lhd_subset_Sup_llist chain_not_lnull by metis qed qed lemma assumes "chain (\) Ns" shows Rf_limit_Sup: "Rf (Liminf_llist Ns) = Rf (Sup_llist Ns)" and Ri_limit_Sup: "Ri (Liminf_llist Ns) = Ri (Sup_llist Ns)" using assms by (auto simp: Rf_Sup_subset_Rf_Liminf Rf_mono Ri_Sup_subset_Ri_Liminf Ri_mono Liminf_llist_subset_Sup_llist subset_antisym) end text \ The assumption below corresponds to condition R4 of Definition 4.1. \ locale effective_redundancy_criterion = redundancy_criterion + assumes Ri_effective: "\ \ \ \ concl_of \ \ N \ Rf N \ \ \ Ri N" begin definition fair_clss_seq :: "'a clause set llist \ bool" where "fair_clss_seq Ns \ (let N' = Liminf_llist Ns - Rf (Liminf_llist Ns) in concls_of (inferences_from N' - Ri N') \ Sup_llist Ns \ Rf (Sup_llist Ns))" end locale sat_preserving_effective_redundancy_criterion = sat_preserving_inference_system "\ :: ('a :: wellorder) inference set" + effective_redundancy_criterion begin sublocale sat_preserving_redundancy_criterion .. text \ The result below corresponds to Theorem 4.3. \ theorem fair_derive_saturated_upto: assumes deriv: "chain (\) Ns" and fair: "fair_clss_seq Ns" shows "saturated_upto (Liminf_llist Ns)" unfolding saturated_upto_def proof fix \ let ?N' = "Liminf_llist Ns - Rf (Liminf_llist Ns)" assume \: "\ \ inferences_from ?N'" show "\ \ Ri (Liminf_llist Ns)" proof (cases "\ \ Ri ?N'") case True then show ?thesis using Ri_mono by blast next case False have "concls_of (inferences_from ?N' - Ri ?N') \ Sup_llist Ns \ Rf (Sup_llist Ns)" using fair unfolding fair_clss_seq_def Let_def . then have "concl_of \ \ Sup_llist Ns \ Rf (Sup_llist Ns)" using False \ by auto moreover { assume "concl_of \ \ Sup_llist Ns" then have "\ \ Ri (Sup_llist Ns)" using \ Ri_effective inferences_from_def by blast then have "\ \ Ri (Liminf_llist Ns)" using deriv Ri_Sup_subset_Ri_Liminf by fast } moreover { assume "concl_of \ \ Rf (Sup_llist Ns)" then have "concl_of \ \ Rf (Liminf_llist Ns)" using deriv Rf_Sup_subset_Rf_Liminf by blast then have "\ \ Ri (Liminf_llist Ns)" using \ Ri_effective inferences_from_def by auto } ultimately show "\ \ Ri (Liminf_llist Ns)" by blast qed qed end text \ This corresponds to the trivial redundancy criterion defined on page 36 of Section 4.1. \ locale trivial_redundancy_criterion = inference_system begin definition Rf :: "'a clause set \ 'a clause set" where "Rf _ = {}" definition Ri :: "'a clause set \ 'a inference set" where "Ri N = {\. \ \ \ \ concl_of \ \ N}" sublocale effective_redundancy_criterion \ Rf Ri by unfold_locales (auto simp: Rf_def Ri_def) lemma saturated_upto_iff: "saturated_upto N \ concls_of (inferences_from N) \ N" unfolding saturated_upto_def inferences_from_def Rf_def Ri_def by auto end text \ The following lemmas corresponds to the standard extension of a redundancy criterion defined on page 38 of Section 4.1. \ lemma redundancy_criterion_standard_extension: assumes "\ \ \'" and "redundancy_criterion \ Rf Ri" shows "redundancy_criterion \' Rf (\N. Ri N \ (\' - \))" using assms unfolding redundancy_criterion_def by (intro conjI) ((auto simp: rev_subsetD)[5], sat) lemma redundancy_criterion_standard_extension_saturated_upto_iff: assumes "\ \ \'" and "redundancy_criterion \ Rf Ri" shows "redundancy_criterion.saturated_upto \ Rf Ri M \ redundancy_criterion.saturated_upto \' Rf (\N. Ri N \ (\' - \)) M" using assms redundancy_criterion.saturated_upto_def redundancy_criterion.saturated_upto_def redundancy_criterion_standard_extension unfolding inference_system.inferences_from_def by blast lemma redundancy_criterion_standard_extension_effective: assumes "\ \ \'" and "effective_redundancy_criterion \ Rf Ri" shows "effective_redundancy_criterion \' Rf (\N. Ri N \ (\' - \))" using assms redundancy_criterion_standard_extension[of \] unfolding effective_redundancy_criterion_def effective_redundancy_criterion_axioms_def by auto lemma redundancy_criterion_standard_extension_fair_iff: assumes "\ \ \'" and "effective_redundancy_criterion \ Rf Ri" shows "effective_redundancy_criterion.fair_clss_seq \' Rf (\N. Ri N \ (\' - \)) Ns \ effective_redundancy_criterion.fair_clss_seq \ Rf Ri Ns" using assms redundancy_criterion_standard_extension_effective[of \ \' Rf Ri] effective_redundancy_criterion.fair_clss_seq_def[of \ Rf Ri Ns] effective_redundancy_criterion.fair_clss_seq_def[of \' Rf "(\N. Ri N \ (\' - \))" Ns] unfolding inference_system.inferences_from_def Let_def by auto theorem redundancy_criterion_standard_extension_fair_derive_saturated_upto: assumes subs: "\ \ \'" and red: "redundancy_criterion \ Rf Ri" and red': "sat_preserving_effective_redundancy_criterion \' Rf (\N. Ri N \ (\' - \))" and deriv: "chain (redundancy_criterion.derive \' Rf) Ns" and fair: "effective_redundancy_criterion.fair_clss_seq \' Rf (\N. Ri N \ (\' - \)) Ns" shows "redundancy_criterion.saturated_upto \ Rf Ri (Liminf_llist Ns)" proof - have "redundancy_criterion.saturated_upto \' Rf (\N. Ri N \ (\' - \)) (Liminf_llist Ns)" by (rule sat_preserving_effective_redundancy_criterion.fair_derive_saturated_upto [OF red' deriv fair]) then show ?thesis by (rule redundancy_criterion_standard_extension_saturated_upto_iff[THEN iffD2, OF subs red]) qed end diff --git a/thys/Saturation_Framework/Calculi.thy b/thys/Saturation_Framework/Calculi.thy --- a/thys/Saturation_Framework/Calculi.thy +++ b/thys/Saturation_Framework/Calculi.thy @@ -1,950 +1,951 @@ (* Title: Calculi of the Saturation Framework * Author: Sophie Tourret , 2018-2020 *) section \Calculi\ text \In this section, the section 2.2 to 2.4 of the report are covered. This includes results on calculi equipped with a redundancy criterion or with a family of redundancy criteria, as well as a proof that various notions of redundancy are equivalent\ theory Calculi imports Consequence_Relations_and_Inference_Systems Ordered_Resolution_Prover.Lazy_List_Liminf Ordered_Resolution_Prover.Lazy_List_Chain begin subsection \Calculi with a Redundancy Criterion\ locale calculus_with_red_crit = inference_system Inf + consequence_relation Bot entails for Bot :: "'f set" and Inf :: \'f inference set\ and entails :: "'f set \ 'f set \ bool" (infix "\" 50) + fixes Red_Inf :: "'f set \ 'f inference set" and Red_F :: "'f set \ 'f set" assumes Red_Inf_to_Inf: "Red_Inf N \ Inf" and Red_F_Bot: "B \ Bot \ N \ {B} \ N - Red_F N \ {B}" and Red_F_of_subset: "N \ N' \ Red_F N \ Red_F N'" and Red_Inf_of_subset: "N \ N' \ Red_Inf N \ Red_Inf N'" and Red_F_of_Red_F_subset: "N' \ Red_F N \ Red_F N \ Red_F (N - N')" and Red_Inf_of_Red_F_subset: "N' \ Red_F N \ Red_Inf N \ Red_Inf (N - N')" and Red_Inf_of_Inf_to_N: "\ \ Inf \ concl_of \ \ N \ \ \ Red_Inf N" begin lemma Red_Inf_of_Inf_to_N_subset: "{\ \ Inf. concl_of \ \ N} \ Red_Inf N" using Red_Inf_of_Inf_to_N by blast (* lem:red-concl-implies-red-inf *) lemma red_concl_to_red_inf: assumes i_in: "\ \ Inf" and concl: "concl_of \ \ Red_F N" shows "\ \ Red_Inf N" proof - have "\ \ Red_Inf (Red_F N)" by (simp add: Red_Inf_of_Inf_to_N i_in concl) then have i_in_Red: "\ \ Red_Inf (N \ Red_F N)" by (simp add: Red_Inf_of_Inf_to_N concl i_in) have red_n_subs: "Red_F N \ Red_F (N \ Red_F N)" by (simp add: Red_F_of_subset) then have "\ \ Red_Inf ((N \ Red_F N) - (Red_F N - N))" using Red_Inf_of_Red_F_subset i_in_Red by (meson Diff_subset subsetCE subset_trans) then show ?thesis by (metis Diff_cancel Diff_subset Un_Diff Un_Diff_cancel contra_subsetD calculus_with_red_crit.Red_Inf_of_subset calculus_with_red_crit_axioms sup_bot.right_neutral) qed definition saturated :: "'f set \ bool" where "saturated N \ Inf_from N \ Red_Inf N" definition reduc_saturated :: "'f set \ bool" where "reduc_saturated N \ Inf_from (N - Red_F N) \ Red_Inf N" lemma Red_Inf_without_red_F: "Red_Inf (N - Red_F N) = Red_Inf N" using Red_Inf_of_subset [of "N - Red_F N" N] and Red_Inf_of_Red_F_subset [of "Red_F N" N] by blast lemma saturated_without_red_F: assumes saturated: "saturated N" shows "saturated (N - Red_F N)" proof - have "Inf_from (N - Red_F N) \ Inf_from N" unfolding Inf_from_def by auto also have "Inf_from N \ Red_Inf N" using saturated unfolding saturated_def by auto also have "Red_Inf N \ Red_Inf (N - Red_F N)" using Red_Inf_of_Red_F_subset by auto finally have "Inf_from (N - Red_F N) \ Red_Inf (N - Red_F N)" by auto then show ?thesis unfolding saturated_def by auto qed definition fair :: "'f set llist \ bool" where "fair D \ Inf_from (Liminf_llist D) \ Sup_llist (lmap Red_Inf D)" inductive "derive" :: "'f set \ 'f set \ bool" (infix "\Red" 50) where derive: "M - N \ Red_F N \ M \Red N" lemma gt_Max_notin: \finite A \ A \ {} \ x > Max A \ x \ A\ by auto lemma equiv_Sup_Liminf: assumes in_Sup: "C \ Sup_llist D" and not_in_Liminf: "C \ Liminf_llist D" shows "\i \ {i. enat (Suc i) < llength D}. C \ lnth D i - lnth D (Suc i)" proof - obtain i where C_D_i: "C \ Sup_upto_llist D (enat i)" and "enat i < llength D" using elem_Sup_llist_imp_Sup_upto_llist in_Sup by fast then obtain j where j: "j \ i \ enat j < llength D \ C \ lnth D j" using not_in_Liminf unfolding Sup_upto_llist_def chain_def Liminf_llist_def by auto obtain k where k: "C \ lnth D k" "enat k < llength D" "k \ i" using C_D_i unfolding Sup_upto_llist_def by auto let ?S = "{i. i < j \ i \ k \ C \ lnth D i}" define l where "l = Max ?S" have \k \ {i. i < j \ k \ i \ C \ lnth D i}\ using k j by (auto simp: order.order_iff_strict) then have nempty: "{i. i < j \ k \ i \ C \ lnth D i} \ {}" by auto then have l_prop: "l < j \ l \ k \ C \ lnth D l" using Max_in[of ?S, OF _ nempty] unfolding l_def by auto then have "C \ lnth D l - lnth D (Suc l)" using j gt_Max_notin[OF _ nempty, of "Suc l"] unfolding l_def[symmetric] by (auto intro: Suc_lessI) then show ?thesis proof (rule bexI[of _ l]) show "l \ {i. enat (Suc i) < llength D}" using l_prop j by (clarify, metis Suc_leI dual_order.order_iff_strict enat_ord_simps(2) less_trans) qed qed (* lem:nonpersistent-is-redundant *) lemma Red_in_Sup: assumes deriv: "chain (\Red) D" shows "Sup_llist D - Liminf_llist D \ Red_F (Sup_llist D)" proof fix C assume C_in_subset: "C \ Sup_llist D - Liminf_llist D" { fix C i assume in_ith_elem: "C \ lnth D i - lnth D (Suc i)" and i: "enat (Suc i) < llength D" have "lnth D i \Red lnth D (Suc i)" using i deriv in_ith_elem chain_lnth_rel by auto then have "C \ Red_F (lnth D (Suc i))" using in_ith_elem derive.cases by blast then have "C \ Red_F (Sup_llist D)" using Red_F_of_subset by (meson contra_subsetD i lnth_subset_Sup_llist) } then show "C \ Red_F (Sup_llist D)" using equiv_Sup_Liminf[of C] C_in_subset by fast qed (* lem:redundant-remains-redundant-during-run 1/2 *) lemma Red_Inf_subset_Liminf: assumes deriv: \chain (\Red) D\ and i: \enat i < llength D\ shows \Red_Inf (lnth D i) \ Red_Inf (Liminf_llist D)\ proof - have Sup_in_diff: \Red_Inf (Sup_llist D) \ Red_Inf (Sup_llist D - (Sup_llist D - Liminf_llist D))\ using Red_Inf_of_Red_F_subset[OF Red_in_Sup] deriv by auto also have \Sup_llist D - (Sup_llist D - Liminf_llist D) = Liminf_llist D\ by (simp add: Liminf_llist_subset_Sup_llist double_diff) then have Red_Inf_Sup_in_Liminf: \Red_Inf (Sup_llist D) \ Red_Inf (Liminf_llist D)\ using Sup_in_diff by auto have \lnth D i \ Sup_llist D\ unfolding Sup_llist_def using i by blast then have "Red_Inf (lnth D i) \ Red_Inf (Sup_llist D)" using Red_Inf_of_subset unfolding Sup_llist_def by auto then show ?thesis using Red_Inf_Sup_in_Liminf by auto qed (* lem:redundant-remains-redundant-during-run 2/2 *) lemma Red_F_subset_Liminf: assumes deriv: \chain (\Red) D\ and i: \enat i < llength D\ shows \Red_F (lnth D i) \ Red_F (Liminf_llist D)\ proof - have Sup_in_diff: \Red_F (Sup_llist D) \ Red_F (Sup_llist D - (Sup_llist D - Liminf_llist D))\ using Red_F_of_Red_F_subset[OF Red_in_Sup] deriv by auto also have \Sup_llist D - (Sup_llist D - Liminf_llist D) = Liminf_llist D\ by (simp add: Liminf_llist_subset_Sup_llist double_diff) then have Red_F_Sup_in_Liminf: \Red_F (Sup_llist D) \ Red_F (Liminf_llist D)\ using Sup_in_diff by auto have \lnth D i \ Sup_llist D\ unfolding Sup_llist_def using i by blast then have "Red_F (lnth D i) \ Red_F (Sup_llist D)" using Red_F_of_subset unfolding Sup_llist_def by auto then show ?thesis using Red_F_Sup_in_Liminf by auto qed (* lem:N-i-is-persistent-or-redundant *) lemma i_in_Liminf_or_Red_F: assumes deriv: \chain (\Red) D\ and i: \enat i < llength D\ shows \lnth D i \ Red_F (Liminf_llist D) \ Liminf_llist D\ proof (rule,rule) fix C assume C: \C \ lnth D i\ and C_not_Liminf: \C \ Liminf_llist D\ have \C \ Sup_llist D\ unfolding Sup_llist_def using C i by auto then obtain j where j: \C \ lnth D j - lnth D (Suc j)\ \enat (Suc j) < llength D\ using equiv_Sup_Liminf[of C D] C_not_Liminf by auto then have \C \ Red_F (lnth D (Suc j))\ using deriv by (meson chain_lnth_rel contra_subsetD derive.cases) then show \C \ Red_F (Liminf_llist D)\ using Red_F_subset_Liminf[of D "Suc j"] deriv j(2) by blast qed (* lem:fairness-implies-saturation *) lemma fair_implies_Liminf_saturated: assumes deriv: \chain (\Red) D\ and fair: \fair D\ shows \saturated (Liminf_llist D)\ unfolding saturated_def proof fix \ assume \: \\ \ Inf_from (Liminf_llist D)\ have \\ \ Sup_llist (lmap Red_Inf D)\ using fair \ unfolding fair_def by auto then obtain i where i: \enat i < llength D\ \\ \ Red_Inf (lnth D i)\ unfolding Sup_llist_def by auto then show \\ \ Red_Inf (Liminf_llist D)\ using deriv i_in_Liminf_or_Red_F[of D i] Red_Inf_subset_Liminf by blast qed end locale static_refutational_complete_calculus = calculus_with_red_crit + assumes static_refutational_complete: "B \ Bot \ saturated N \ N \ {B} \ \B'\Bot. B' \ N" begin lemma dynamic_refutational_complete_Liminf: fixes B D assumes bot_elem: \B \ Bot\ and deriv: \chain (\Red) D\ and fair: \fair D\ and - unsat: \lnth D 0 \ {B}\ + unsat: \lhd D \ {B}\ shows \\B'\Bot. B' \ Liminf_llist D\ proof - + note lhd_is = lhd_conv_lnth[OF chain_not_lnull[OF deriv]] have non_empty: \\ lnull D\ using chain_not_lnull[OF deriv] . - have subs: \lnth D 0 \ Sup_llist D\ + have subs: \lhd D \ Sup_llist D\ using lhd_subset_Sup_llist[of D] non_empty by (simp add: lhd_conv_lnth) have \Sup_llist D \ {B}\ - using unsat subset_entailed[OF subs] entails_trans[of "Sup_llist D" "lnth D 0"] by auto + using unsat subset_entailed[OF subs] entails_trans[of "Sup_llist D" "lhd D"] by auto then have Sup_no_Red: \Sup_llist D - Red_F (Sup_llist D) \ {B}\ using bot_elem Red_F_Bot by auto have Sup_no_Red_in_Liminf: \Sup_llist D - Red_F (Sup_llist D) \ Liminf_llist D\ using deriv Red_in_Sup by auto have Liminf_entails_Bot: \Liminf_llist D \ {B}\ using Sup_no_Red subset_entailed[OF Sup_no_Red_in_Liminf] entails_trans by blast have \saturated (Liminf_llist D)\ using deriv fair fair_implies_Liminf_saturated unfolding saturated_def by auto then show ?thesis using bot_elem static_refutational_complete Liminf_entails_Bot by auto qed end locale dynamic_refutational_complete_calculus = calculus_with_red_crit + assumes - dynamic_refutational_complete: "B \ Bot \ chain (\Red) D \ fair D \ lnth D 0 \ {B} \ + dynamic_refutational_complete: "B \ Bot \ chain (\Red) D \ fair D \ lhd D \ {B} \ \i \ {i. enat i < llength D}. \B'\Bot. B' \ lnth D i" begin (* lem:dynamic-ref-compl-implies-static *) sublocale static_refutational_complete_calculus proof fix B N assume bot_elem: \B \ Bot\ and saturated_N: "saturated N" and refut_N: "N \ {B}" define D where "D = LCons N LNil" have[simp]: \\ lnull D\ by (auto simp: D_def) have deriv_D: \chain (\Red) D\ by (simp add: chain.chain_singleton D_def) have liminf_is_N: "Liminf_llist D = N" by (simp add: D_def Liminf_llist_LCons) - have head_D: "N = lnth D 0" by (simp add: D_def) + have head_D: "N = lhd D" by (simp add: D_def) have "Sup_llist (lmap Red_Inf D) = Red_Inf N" by (simp add: D_def) then have fair_D: "fair D" using saturated_N by (simp add: fair_def saturated_def liminf_is_N) obtain i B' where B'_is_bot: \B' \ Bot\ and B'_in: "B' \ lnth D i" and \i < llength D\ using dynamic_refutational_complete[of B D] bot_elem fair_D head_D saturated_N deriv_D refut_N by auto then have "i = 0" by (auto simp: D_def enat_0_iff) show \\B'\Bot. B' \ N\ - using B'_is_bot B'_in unfolding \i = 0\ head_D[symmetric] by auto + using B'_is_bot B'_in unfolding \i = 0\ head_D[symmetric] D_def by auto qed end (* lem:static-ref-compl-implies-dynamic *) sublocale static_refutational_complete_calculus \ dynamic_refutational_complete_calculus proof fix B D assume \B \ Bot\ and \chain (\Red) D\ and \fair D\ and - \lnth D 0 \ {B}\ + \lhd D \ {B}\ then have \\B'\Bot. B' \ Liminf_llist D\ by (rule dynamic_refutational_complete_Liminf) then show \\i\{i. enat i < llength D}. \B'\Bot. B' \ lnth D i\ unfolding Liminf_llist_def by auto qed subsection \Calculi with a Family of Redundancy Criteria\ locale calculus_with_red_crit_family = inference_system Inf + consequence_relation_family Bot Q entails_q for Bot :: "'f set" and Inf :: \'f inference set\ and Q :: "'q set" and entails_q :: "'q \ 'f set \ 'f set \ bool" + fixes Red_Inf_q :: "'q \ 'f set \ 'f inference set" and Red_F_q :: "'q \ 'f set \ 'f set" assumes Q_nonempty: "Q \ {}" and all_red_crit: "\q \ Q. calculus_with_red_crit Bot Inf (entails_q q) (Red_Inf_q q) (Red_F_q q)" begin definition Red_Inf_Q :: "'f set \ 'f inference set" where "Red_Inf_Q N = (\q \ Q. Red_Inf_q q N)" definition Red_F_Q :: "'f set \ 'f set" where "Red_F_Q N = (\q \ Q. Red_F_q q N)" (* lem:intersection-of-red-crit *) sublocale calculus_with_red_crit Bot Inf entails_Q Red_Inf_Q Red_F_Q unfolding calculus_with_red_crit_def calculus_with_red_crit_axioms_def proof (intro conjI) show "consequence_relation Bot entails_Q" using intersect_cons_rel_family . next show "\N. Red_Inf_Q N \ Inf" unfolding Red_Inf_Q_def proof fix N show "(\q \ Q. Red_Inf_q q N) \ Inf" proof (intro Inter_subset) fix Red_Infs assume one_red_inf: "Red_Infs \ (\q. Red_Inf_q q N) ` Q" show "Red_Infs \ Inf" using one_red_inf all_red_crit calculus_with_red_crit.Red_Inf_to_Inf by blast next show "(\q. Red_Inf_q q N) ` Q \ {}" using Q_nonempty by blast qed qed next show "\B N. B \ Bot \ N \Q {B} \ N - Red_F_Q N \Q {B}" proof (intro allI impI) fix B N assume B_in: "B \ Bot" and N_unsat: "N \Q {B}" show "N - Red_F_Q N \Q {B}" unfolding entails_Q_def Red_F_Q_def proof fix qi assume qi_in: "qi \ Q" define entails_qi (infix "\qi" 50) where "entails_qi = entails_q qi" have cons_rel_qi: "consequence_relation Bot entails_qi" unfolding entails_qi_def using qi_in all_red_crit calculus_with_red_crit.axioms(1) by blast define Red_F_qi where "Red_F_qi = Red_F_q qi" have red_qi_in_Q: "Red_F_Q N \ Red_F_qi N" unfolding Red_F_Q_def Red_F_qi_def using qi_in image_iff by blast then have "N - Red_F_qi N \ N - Red_F_Q N" by blast then have entails_1: "N - Red_F_Q N \qi N - Red_F_qi N" using qi_in all_red_crit unfolding calculus_with_red_crit_def consequence_relation_def entails_qi_def by metis have N_unsat_qi: "N \qi {B}" using qi_in N_unsat unfolding entails_qi_def entails_Q_def by simp then have N_unsat_qi: "N - Red_F_qi N \qi {B}" using qi_in all_red_crit Red_F_qi_def calculus_with_red_crit.Red_F_Bot[OF _ B_in] entails_qi_def by fastforce show "N - (\q \ Q. Red_F_q q N) \qi {B}" using consequence_relation.entails_trans[OF cons_rel_qi entails_1 N_unsat_qi] unfolding Red_F_Q_def . qed qed next show "\N1 N2. N1 \ N2 \ Red_F_Q N1 \ Red_F_Q N2" proof (intro allI impI) fix N1 :: "'f set" and N2 :: "'f set" assume N1_in_N2: "N1 \ N2" show "Red_F_Q N1 \ Red_F_Q N2" proof fix C assume "C \ Red_F_Q N1" then have "\qi \ Q. C \ Red_F_q qi N1" unfolding Red_F_Q_def by blast then have "\qi \ Q. C \ Red_F_q qi N2" using N1_in_N2 all_red_crit calculus_with_red_crit.axioms(2) calculus_with_red_crit.Red_F_of_subset by blast then show "C \ Red_F_Q N2" unfolding Red_F_Q_def by blast qed qed next show "\N1 N2. N1 \ N2 \ Red_Inf_Q N1 \ Red_Inf_Q N2" proof (intro allI impI) fix N1 :: "'f set" and N2 :: "'f set" assume N1_in_N2: "N1 \ N2" show "Red_Inf_Q N1 \ Red_Inf_Q N2" proof fix \ assume "\ \ Red_Inf_Q N1" then have "\qi \ Q. \ \ Red_Inf_q qi N1" unfolding Red_Inf_Q_def by blast then have "\qi \ Q. \ \ Red_Inf_q qi N2" using N1_in_N2 all_red_crit calculus_with_red_crit.axioms(2) calculus_with_red_crit.Red_Inf_of_subset by blast then show "\ \ Red_Inf_Q N2" unfolding Red_Inf_Q_def by blast qed qed next show "\N2 N1. N2 \ Red_F_Q N1 \ Red_F_Q N1 \ Red_F_Q (N1 - N2)" proof (intro allI impI) fix N2 N1 assume N2_in_Red_N1: "N2 \ Red_F_Q N1" show "Red_F_Q N1 \ Red_F_Q (N1 - N2)" proof fix C assume "C \ Red_F_Q N1" then have "\qi \ Q. C \ Red_F_q qi N1" unfolding Red_F_Q_def by blast moreover have "\qi \ Q. N2 \ Red_F_q qi N1" using N2_in_Red_N1 unfolding Red_F_Q_def by blast ultimately have "\qi \ Q. C \ Red_F_q qi (N1 - N2)" using all_red_crit calculus_with_red_crit.axioms(2) calculus_with_red_crit.Red_F_of_Red_F_subset by blast then show "C \ Red_F_Q (N1 - N2)" unfolding Red_F_Q_def by blast qed qed next show "\N2 N1. N2 \ Red_F_Q N1 \ Red_Inf_Q N1 \ Red_Inf_Q (N1 - N2)" proof (intro allI impI) fix N2 N1 assume N2_in_Red_N1: "N2 \ Red_F_Q N1" show "Red_Inf_Q N1 \ Red_Inf_Q (N1 - N2)" proof fix \ assume "\ \ Red_Inf_Q N1" then have "\qi \ Q. \ \ Red_Inf_q qi N1" unfolding Red_Inf_Q_def by blast moreover have "\qi \ Q. N2 \ Red_F_q qi N1" using N2_in_Red_N1 unfolding Red_F_Q_def by blast ultimately have "\qi \ Q. \ \ Red_Inf_q qi (N1 - N2)" using all_red_crit calculus_with_red_crit.axioms(2) calculus_with_red_crit.Red_Inf_of_Red_F_subset by blast then show "\ \ Red_Inf_Q (N1 - N2)" unfolding Red_Inf_Q_def by blast qed qed next show "\\ N. \ \ Inf \ concl_of \ \ N \ \ \ Red_Inf_Q N" proof (intro allI impI) fix \ N assume i_in: "\ \ Inf" and concl_in: "concl_of \ \ N" then have "\qi \ Q. \ \ Red_Inf_q qi N" using all_red_crit calculus_with_red_crit.axioms(2) calculus_with_red_crit.Red_Inf_of_Inf_to_N by blast then show "\ \ Red_Inf_Q N" unfolding Red_Inf_Q_def by blast qed qed (* lem:satur-wrt-intersection-of-red *) lemma sat_int_to_sat_q: "calculus_with_red_crit.saturated Inf Red_Inf_Q N \ (\qi \ Q. calculus_with_red_crit.saturated Inf (Red_Inf_q qi) N)" for N proof fix N assume inter_sat: "calculus_with_red_crit.saturated Inf Red_Inf_Q N" show "\qi \ Q. calculus_with_red_crit.saturated Inf (Red_Inf_q qi) N" proof fix qi assume qi_in: "qi \ Q" then interpret one: calculus_with_red_crit Bot Inf "entails_q qi" "Red_Inf_q qi" "Red_F_q qi" by (metis all_red_crit) show "one.saturated N" using qi_in inter_sat unfolding one.saturated_def saturated_def Red_Inf_Q_def by blast qed next fix N assume all_sat: "\qi \ Q. calculus_with_red_crit.saturated Inf (Red_Inf_q qi) N" show "saturated N" unfolding saturated_def Red_Inf_Q_def proof fix \ assume \_in: "\ \ Inf_from N" have "\Red_Inf_qi \ Red_Inf_q ` Q. \ \ Red_Inf_qi N" proof fix Red_Inf_qi assume red_inf_in: "Red_Inf_qi \ Red_Inf_q ` Q" then obtain qi where qi_in: "qi \ Q" and red_inf_qi_def: "Red_Inf_qi = Red_Inf_q qi" by blast then interpret one: calculus_with_red_crit Bot Inf "entails_q qi" "Red_Inf_q qi" "Red_F_q qi" by (metis all_red_crit) have "one.saturated N" using qi_in all_sat red_inf_qi_def by blast then show "\ \ Red_Inf_qi N" unfolding one.saturated_def using \_in red_inf_qi_def by blast qed then show "\ \ (\q \ Q. Red_Inf_q q N)" by blast qed qed (* lem:checking-static-ref-compl-for-intersections *) lemma stat_ref_comp_from_bot_in_sat: "(\N. calculus_with_red_crit.saturated Inf Red_Inf_Q N \ (\B \ Bot. B \ N) \ (\B \ Bot. \qi \ Q. \ entails_q qi N {B})) \ static_refutational_complete_calculus Bot Inf entails_Q Red_Inf_Q Red_F_Q" proof (rule ccontr) assume N_saturated: "\N. (calculus_with_red_crit.saturated Inf Red_Inf_Q N \ (\B \ Bot. B \ N)) \ (\B \ Bot. \qi \ Q. \ entails_q qi N {B})" and no_stat_ref_comp: "\ static_refutational_complete_calculus Bot Inf (\Q) Red_Inf_Q Red_F_Q" obtain N1 B1 where B1_in: "B1 \ Bot" and N1_saturated: "calculus_with_red_crit.saturated Inf Red_Inf_Q N1" and N1_unsat: "N1 \Q {B1}" and no_B_in_N1: "\B \ Bot. B \ N1" using no_stat_ref_comp by (metis calculus_with_red_crit_axioms static_refutational_complete_calculus.intro static_refutational_complete_calculus_axioms.intro) obtain B2 qi where qi_in: "qi \ Q" and no_qi: "\ entails_q qi N1 {B2}" using N_saturated N1_saturated no_B_in_N1 by auto have "N1 \Q {B2}" using N1_unsat B1_in intersect_cons_rel_family unfolding consequence_relation_def by metis then have "entails_q qi N1 {B2}" unfolding entails_Q_def using qi_in by blast then show False using no_qi by simp qed end subsection \Families of Calculi with a Family of Redundancy Criteria\ locale calculus_family_with_red_crit_family = inference_system_family Q Inf_q + consequence_relation_family Bot Q entails_q for Bot :: "'f set" and Q :: "'q set" and Inf_q :: \'q \ 'f inference set\ and entails_q :: "'q \ 'f set \ 'f set \ bool" + fixes Red_Inf_q :: "'q \ 'f set \ 'f inference set" and Red_F_q :: "'q \ 'f set \ 'f set" assumes Q_nonempty: "Q \ {}" and all_red_crit: "\q \ Q. calculus_with_red_crit Bot (Inf_q q) (entails_q q) (Red_Inf_q q) (Red_F_q q)" subsection \Variations on a Theme\ locale calculus_with_reduced_red_crit = calculus_with_red_crit Bot Inf entails Red_Inf Red_F for Bot :: "'f set" and Inf :: \'f inference set\ and entails :: "'f set \ 'f set \ bool" (infix "\" 50) and Red_Inf :: "'f set \ 'f inference set" and Red_F :: "'f set \ 'f set" + assumes inf_in_red_inf: "Inf_from2 UNIV (Red_F N) \ Red_Inf N" begin (* lem:reduced-rc-implies-sat-equiv-reduced-sat *) lemma sat_eq_reduc_sat: "saturated N \ reduc_saturated N" proof fix N assume "saturated N" then show "reduc_saturated N" using Red_Inf_without_red_F saturated_without_red_F unfolding saturated_def reduc_saturated_def by blast next fix N assume red_sat_n: "reduc_saturated N" show "saturated N" unfolding saturated_def using red_sat_n inf_in_red_inf unfolding reduc_saturated_def Inf_from_def Inf_from2_def by blast qed end locale reduc_static_refutational_complete_calculus = calculus_with_red_crit + assumes reduc_static_refutational_complete: "B \ Bot \ reduc_saturated N \ N \ {B} \ \B'\Bot. B' \ N" locale reduc_static_refutational_complete_reduc_calculus = calculus_with_reduced_red_crit + assumes reduc_static_refutational_complete: "B \ Bot \ reduc_saturated N \ N \ {B} \ \B'\Bot. B' \ N" begin sublocale reduc_static_refutational_complete_calculus by (simp add: calculus_with_red_crit_axioms reduc_static_refutational_complete reduc_static_refutational_complete_calculus_axioms.intro reduc_static_refutational_complete_calculus_def) (* cor:reduced-rc-implies-st-ref-comp-equiv-reduced-st-ref-comp 1/2 *) sublocale static_refutational_complete_calculus proof fix B N assume bot_elem: \B \ Bot\ and saturated_N: "saturated N" and refut_N: "N \ {B}" have reduc_saturated_N: "reduc_saturated N" using saturated_N sat_eq_reduc_sat by blast show "\B'\Bot. B' \ N" using reduc_static_refutational_complete[OF bot_elem reduc_saturated_N refut_N] . qed end context calculus_with_reduced_red_crit begin (* cor:reduced-rc-implies-st-ref-comp-equiv-reduced-st-ref-comp 2/2 *) lemma stat_ref_comp_imp_red_stat_ref_comp: "static_refutational_complete_calculus Bot Inf entails Red_Inf Red_F \ reduc_static_refutational_complete_calculus Bot Inf entails Red_Inf Red_F" proof fix B N assume stat_ref_comp: "static_refutational_complete_calculus Bot Inf (\) Red_Inf Red_F" and bot_elem: \B \ Bot\ and saturated_N: "reduc_saturated N" and refut_N: "N \ {B}" have reduc_saturated_N: "saturated N" using saturated_N sat_eq_reduc_sat by blast show "\B'\Bot. B' \ N" using static_refutational_complete_calculus.static_refutational_complete[OF stat_ref_comp bot_elem reduc_saturated_N refut_N] . qed end context calculus_with_red_crit begin definition Red_Red_Inf :: "'f set \ 'f inference set" where "Red_Red_Inf N = Red_Inf N \ Inf_from2 UNIV (Red_F N)" lemma reduced_calc_is_calc: "calculus_with_red_crit Bot Inf entails Red_Red_Inf Red_F" proof fix N show "Red_Red_Inf N \ Inf" unfolding Red_Red_Inf_def Inf_from2_def Inf_from_def using Red_Inf_to_Inf by auto next fix B N assume b_in: "B \ Bot" and n_entails: "N \ {B}" show "N - Red_F N \ {B}" by (simp add: Red_F_Bot b_in n_entails) next fix N N' :: "'f set" assume "N \ N'" then show "Red_F N \ Red_F N'" by (simp add: Red_F_of_subset) next fix N N' :: "'f set" assume n_in: "N \ N'" then have "Inf_from (UNIV - (Red_F N')) \ Inf_from (UNIV - (Red_F N))" using Red_F_of_subset[OF n_in] unfolding Inf_from_def by auto then have "Inf_from2 UNIV (Red_F N) \ Inf_from2 UNIV (Red_F N')" unfolding Inf_from2_def by auto then show "Red_Red_Inf N \ Red_Red_Inf N'" unfolding Red_Red_Inf_def using Red_Inf_of_subset[OF n_in] by blast next fix N N' :: "'f set" assume "N' \ Red_F N" then show "Red_F N \ Red_F (N - N')" by (simp add: Red_F_of_Red_F_subset) next fix N N' :: "'f set" assume np_subs: "N' \ Red_F N" have "Red_F N \ Red_F (N - N')" by (simp add: Red_F_of_Red_F_subset np_subs) then have "Inf_from (UNIV - (Red_F (N - N'))) \ Inf_from (UNIV - (Red_F N))" by (metis Diff_subset Red_F_of_subset eq_iff) then have "Inf_from2 UNIV (Red_F N) \ Inf_from2 UNIV (Red_F (N - N'))" unfolding Inf_from2_def by auto then show "Red_Red_Inf N \ Red_Red_Inf (N - N')" unfolding Red_Red_Inf_def using Red_Inf_of_Red_F_subset[OF np_subs] by blast next fix \ N assume "\ \ Inf" "concl_of \ \ N" then show "\ \ Red_Red_Inf N" by (simp add: Red_Inf_of_Inf_to_N Red_Red_Inf_def) qed lemma inf_subs_reduced_red_inf: "Inf_from2 UNIV (Red_F N) \ Red_Red_Inf N" unfolding Red_Red_Inf_def by simp (* lem:red'-is-reduced-redcrit *) text \The following is a lemma and not a sublocale as was previously used in similar cases. Here, a sublocale cannot be used because it would create an infinitely descending chain of sublocales. \ lemma reduc_calc: "calculus_with_reduced_red_crit Bot Inf entails Red_Red_Inf Red_F" using inf_subs_reduced_red_inf reduced_calc_is_calc by (simp add: calculus_with_reduced_red_crit.intro calculus_with_reduced_red_crit_axioms_def) interpretation reduc_calc: calculus_with_reduced_red_crit Bot Inf entails Red_Red_Inf Red_F by (fact reduc_calc) (* lem:saturation-red-vs-red'-1 *) lemma sat_imp_red_calc_sat: "saturated N \ reduc_calc.saturated N" unfolding saturated_def reduc_calc.saturated_def Red_Red_Inf_def by blast (* lem:saturation-red-vs-red'-2 1/2 (i) \ (ii) *) lemma red_sat_eq_red_calc_sat: "reduc_saturated N \ reduc_calc.saturated N" proof assume red_sat_n: "reduc_saturated N" show "reduc_calc.saturated N" unfolding reduc_calc.saturated_def proof fix \ assume i_in: "\ \ Inf_from N" show "\ \ Red_Red_Inf N" using i_in red_sat_n unfolding reduc_saturated_def Inf_from2_def Inf_from_def Red_Red_Inf_def by blast qed next assume red_sat_n: "reduc_calc.saturated N" show "reduc_saturated N" unfolding reduc_saturated_def proof fix \ assume i_in: "\ \ Inf_from (N - Red_F N)" show "\ \ Red_Inf N" using i_in red_sat_n unfolding Inf_from_def reduc_calc.saturated_def Red_Red_Inf_def Inf_from2_def by blast qed qed (* lem:saturation-red-vs-red'-2 2/2 (i) \ (iii) *) lemma red_sat_eq_sat: "reduc_saturated N \ saturated (N - Red_F N)" unfolding reduc_saturated_def saturated_def by (simp add: Red_Inf_without_red_F) (* thm:reduced-stat-ref-compl 1/3 (i) \ (iii) *) theorem stat_is_stat_red: "static_refutational_complete_calculus Bot Inf entails Red_Inf Red_F \ static_refutational_complete_calculus Bot Inf entails Red_Red_Inf Red_F" proof assume stat_ref1: "static_refutational_complete_calculus Bot Inf entails Red_Inf Red_F" show "static_refutational_complete_calculus Bot Inf entails Red_Red_Inf Red_F" using reduc_calc.calculus_with_red_crit_axioms unfolding static_refutational_complete_calculus_def static_refutational_complete_calculus_axioms_def proof show "\B N. B \ Bot \ reduc_calc.saturated N \ N \ {B} \ (\B'\Bot. B' \ N)" proof (clarify) fix B N assume b_in: "B \ Bot" and n_sat: "reduc_calc.saturated N" and n_imp_b: "N \ {B}" have "saturated (N - Red_F N)" using red_sat_eq_red_calc_sat[of N] red_sat_eq_sat[of N] n_sat by blast moreover have "(N - Red_F N) \ {B}" using n_imp_b b_in by (simp add: reduc_calc.Red_F_Bot) ultimately show "\B'\Bot. B'\ N" using stat_ref1 by (meson DiffD1 b_in static_refutational_complete_calculus.static_refutational_complete) qed qed next assume stat_ref3: "static_refutational_complete_calculus Bot Inf entails Red_Red_Inf Red_F" show "static_refutational_complete_calculus Bot Inf entails Red_Inf Red_F" unfolding static_refutational_complete_calculus_def static_refutational_complete_calculus_axioms_def using calculus_with_red_crit_axioms proof show "\B N. B \ Bot \ saturated N \ N \ {B} \ (\B'\Bot. B' \ N)" proof clarify fix B N assume b_in: "B \ Bot" and n_sat: "saturated N" and n_imp_b: "N \ {B}" then show "\B'\ Bot. B' \ N" using stat_ref3 sat_imp_red_calc_sat[OF n_sat] by (meson static_refutational_complete_calculus.static_refutational_complete) qed qed qed (* thm:reduced-stat-ref-compl 2/3 (iv) \ (iii) *) theorem red_stat_red_is_stat_red: "reduc_static_refutational_complete_calculus Bot Inf entails Red_Red_Inf Red_F \ static_refutational_complete_calculus Bot Inf entails Red_Red_Inf Red_F" using reduc_calc.stat_ref_comp_imp_red_stat_ref_comp by (metis reduc_calc.sat_eq_reduc_sat reduc_static_refutational_complete_calculus.axioms(2) reduc_static_refutational_complete_calculus_axioms_def reduced_calc_is_calc static_refutational_complete_calculus.intro static_refutational_complete_calculus_axioms.intro) (* thm:reduced-stat-ref-compl 3/3 (ii) \ (iii) *) theorem red_stat_is_stat_red: "reduc_static_refutational_complete_calculus Bot Inf entails Red_Inf Red_F \ static_refutational_complete_calculus Bot Inf entails Red_Red_Inf Red_F" using reduc_calc.calculus_with_red_crit_axioms calculus_with_red_crit_axioms red_sat_eq_red_calc_sat unfolding static_refutational_complete_calculus_def static_refutational_complete_calculus_axioms_def reduc_static_refutational_complete_calculus_def reduc_static_refutational_complete_calculus_axioms_def by blast lemma sup_red_f_in_red_liminf: "chain derive D \ Sup_llist (lmap Red_F D) \ Red_F (Liminf_llist D)" proof fix N assume deriv: "chain derive D" and n_in_sup: "N \ Sup_llist (lmap Red_F D)" obtain i0 where i_smaller: "enat i0 < llength D" and n_in: "N \ Red_F (lnth D i0)" using n_in_sup by (metis Sup_llist_imp_exists_index llength_lmap lnth_lmap) have "Red_F (lnth D i0) \ Red_F (Liminf_llist D)" using i_smaller by (simp add: deriv Red_F_subset_Liminf) then show "N \ Red_F (Liminf_llist D)" using n_in by fast qed lemma sup_red_inf_in_red_liminf: "chain derive D \ Sup_llist (lmap Red_Inf D) \ Red_Inf (Liminf_llist D)" proof fix \ assume deriv: "chain derive D" and i_in_sup: "\ \ Sup_llist (lmap Red_Inf D)" obtain i0 where i_smaller: "enat i0 < llength D" and n_in: "\ \ Red_Inf (lnth D i0)" using i_in_sup unfolding Sup_llist_def by auto have "Red_Inf (lnth D i0) \ Red_Inf (Liminf_llist D)" using i_smaller by (simp add: deriv Red_Inf_subset_Liminf) then show "\ \ Red_Inf (Liminf_llist D)" using n_in by fast qed definition reduc_fair :: "'f set llist \ bool" where "reduc_fair D \ Inf_from (Liminf_llist D - Sup_llist (lmap Red_F D)) \ Sup_llist (lmap Red_Inf D)" (* lem:red-fairness-implies-red-saturation *) lemma reduc_fair_imp_Liminf_reduc_sat: "chain derive D \ reduc_fair D \ reduc_saturated (Liminf_llist D)" unfolding reduc_saturated_def proof - fix D assume deriv: "chain derive D" and red_fair: "reduc_fair D" have "Inf_from (Liminf_llist D - Red_F (Liminf_llist D)) \ Inf_from (Liminf_llist D - Sup_llist (lmap Red_F D))" using sup_red_f_in_red_liminf[OF deriv] unfolding Inf_from_def by blast then have "Inf_from (Liminf_llist D - Red_F (Liminf_llist D)) \ Sup_llist (lmap Red_Inf D)" using red_fair unfolding reduc_fair_def by simp then show "Inf_from (Liminf_llist D - Red_F (Liminf_llist D)) \ Red_Inf (Liminf_llist D)" using sup_red_inf_in_red_liminf[OF deriv] by fast qed end locale reduc_dynamic_refutational_complete_calculus = calculus_with_red_crit + assumes - reduc_dynamic_refutational_complete: "B \ Bot \ chain derive D \ reduc_fair D - \ lnth D 0 \ {B} \ \i \ {i. enat i < llength D}. \ B'\Bot. B' \ lnth D i" + reduc_dynamic_refutational_complete: "B \ Bot \ chain derive D \ reduc_fair D \ + lhd D \ {B} \ \i \ {i. enat i < llength D}. \ B'\Bot. B' \ lnth D i" begin sublocale reduc_static_refutational_complete_calculus proof fix B N assume bot_elem: \B \ Bot\ and saturated_N: "reduc_saturated N" and refut_N: "N \ {B}" define D where "D = LCons N LNil" have[simp]: \\ lnull D\ by (auto simp: D_def) have deriv_D: \chain (\Red) D\ by (simp add: chain.chain_singleton D_def) have liminf_is_N: "Liminf_llist D = N" by (simp add: D_def Liminf_llist_LCons) - have head_D: "N = lnth D 0" by (simp add: D_def) + have head_D: "N = lhd D" by (simp add: D_def) have "Sup_llist (lmap Red_F D) = Red_F N" by (simp add: D_def) moreover have "Sup_llist (lmap Red_Inf D) = Red_Inf N" by (simp add: D_def) ultimately have fair_D: "reduc_fair D" using saturated_N liminf_is_N unfolding reduc_fair_def reduc_saturated_def by (simp add: reduc_fair_def reduc_saturated_def liminf_is_N) obtain i B' where B'_is_bot: \B' \ Bot\ and B'_in: "B' \ lnth D i" and \i < llength D\ using reduc_dynamic_refutational_complete[of B D] bot_elem fair_D head_D saturated_N deriv_D refut_N by auto then have "i = 0" by (auto simp: D_def enat_0_iff) show \\B'\Bot. B' \ N\ - using B'_is_bot B'_in unfolding \i = 0\ head_D[symmetric] by auto + using B'_is_bot B'_in unfolding \i = 0\ head_D[symmetric] D_def by auto qed end sublocale reduc_static_refutational_complete_calculus \ reduc_dynamic_refutational_complete_calculus proof fix B D assume bot_elem: \B \ Bot\ and deriv: \chain (\Red) D\ and fair: \reduc_fair D\ and - unsat: \lnth D 0 \ {B}\ + unsat: \lhd D \ {B}\ have non_empty: \\ lnull D\ using chain_not_lnull[OF deriv] . - have subs: \lnth D 0 \ Sup_llist D\ + have subs: \lhd D \ Sup_llist D\ using lhd_subset_Sup_llist[of D] non_empty by (simp add: lhd_conv_lnth) have \Sup_llist D \ {B}\ - using unsat subset_entailed[OF subs] entails_trans[of "Sup_llist D" "lnth D 0"] by auto + using unsat subset_entailed[OF subs] entails_trans[of "Sup_llist D" "lhd D"] by auto then have Sup_no_Red: \Sup_llist D - Red_F (Sup_llist D) \ {B}\ using bot_elem Red_F_Bot by auto have Sup_no_Red_in_Liminf: \Sup_llist D - Red_F (Sup_llist D) \ Liminf_llist D\ using deriv Red_in_Sup by auto have Liminf_entails_Bot: \Liminf_llist D \ {B}\ using Sup_no_Red subset_entailed[OF Sup_no_Red_in_Liminf] entails_trans by blast have \reduc_saturated (Liminf_llist D)\ using deriv fair reduc_fair_imp_Liminf_reduc_sat unfolding reduc_saturated_def by auto then have \\B'\Bot. B' \ Liminf_llist D\ using bot_elem reduc_static_refutational_complete Liminf_entails_Bot by auto then show \\i\{i. enat i < llength D}. \B'\Bot. B' \ lnth D i\ unfolding Liminf_llist_def by auto qed context calculus_with_red_crit begin lemma dyn_equiv_stat: "dynamic_refutational_complete_calculus Bot Inf entails Red_Inf Red_F = static_refutational_complete_calculus Bot Inf entails Red_Inf Red_F" proof assume "dynamic_refutational_complete_calculus Bot Inf entails Red_Inf Red_F" then interpret dynamic_refutational_complete_calculus Bot Inf entails Red_Inf Red_F by simp show "static_refutational_complete_calculus Bot Inf entails Red_Inf Red_F" by (simp add: static_refutational_complete_calculus_axioms) next assume "static_refutational_complete_calculus Bot Inf entails Red_Inf Red_F" then interpret static_refutational_complete_calculus Bot Inf entails Red_Inf Red_F by simp show "dynamic_refutational_complete_calculus Bot Inf entails Red_Inf Red_F" by (simp add: dynamic_refutational_complete_calculus_axioms) qed lemma red_dyn_equiv_red_stat: "reduc_dynamic_refutational_complete_calculus Bot Inf entails Red_Inf Red_F = reduc_static_refutational_complete_calculus Bot Inf entails Red_Inf Red_F" proof assume "reduc_dynamic_refutational_complete_calculus Bot Inf entails Red_Inf Red_F" then interpret reduc_dynamic_refutational_complete_calculus Bot Inf entails Red_Inf Red_F by simp show "reduc_static_refutational_complete_calculus Bot Inf entails Red_Inf Red_F" by (simp add: reduc_static_refutational_complete_calculus_axioms) next assume "reduc_static_refutational_complete_calculus Bot Inf entails Red_Inf Red_F" then interpret reduc_static_refutational_complete_calculus Bot Inf entails Red_Inf Red_F by simp show "reduc_dynamic_refutational_complete_calculus Bot Inf entails Red_Inf Red_F" by (simp add: reduc_dynamic_refutational_complete_calculus_axioms) qed interpretation reduc_calc: calculus_with_reduced_red_crit Bot Inf entails Red_Red_Inf Red_F by (fact reduc_calc) (* thm:reduced-dyn-ref-compl 1/3 (v) \ (vii) *) theorem dyn_ref_eq_dyn_ref_red: "dynamic_refutational_complete_calculus Bot Inf entails Red_Inf Red_F \ dynamic_refutational_complete_calculus Bot Inf entails Red_Red_Inf Red_F" using dyn_equiv_stat stat_is_stat_red reduc_calc.dyn_equiv_stat by meson (* thm:reduced-dyn-ref-compl 2/3 (viii) \ (vii) *) theorem red_dyn_ref_red_eq_dyn_ref_red: "reduc_dynamic_refutational_complete_calculus Bot Inf entails Red_Red_Inf Red_F \ dynamic_refutational_complete_calculus Bot Inf entails Red_Red_Inf Red_F" using red_dyn_equiv_red_stat dyn_equiv_stat red_stat_red_is_stat_red by (simp add: reduc_calc.dyn_equiv_stat reduc_calc.red_dyn_equiv_red_stat) (* thm:reduced-dyn-ref-compl 3/3 (vi) \ (vii) *) theorem red_dyn_ref_eq_dyn_ref_red: "reduc_dynamic_refutational_complete_calculus Bot Inf entails Red_Inf Red_F \ dynamic_refutational_complete_calculus Bot Inf entails Red_Red_Inf Red_F" using red_dyn_equiv_red_stat dyn_equiv_stat red_stat_is_stat_red reduc_calc.dyn_equiv_stat reduc_calc.red_dyn_equiv_red_stat by blast end end diff --git a/thys/Saturation_Framework/Prover_Architectures.thy b/thys/Saturation_Framework/Prover_Architectures.thy --- a/thys/Saturation_Framework/Prover_Architectures.thy +++ b/thys/Saturation_Framework/Prover_Architectures.thy @@ -1,1167 +1,1172 @@ (* Title: Prover Architectures of the Saturation Framework * Author: Sophie Tourret , 2019-2020 *) section \Prover Architectures\ text \This section covers all the results presented in the section 4 of the report. This is where abstract architectures of provers are defined and proven dynamically refutationally complete.\ theory Prover_Architectures imports Lambda_Free_RPOs.Lambda_Free_Util Labeled_Lifting_to_Non_Ground_Calculi begin subsection \Basis of the Prover Architectures\ locale prover_architecture_basis = std?: labeled_lifting_with_red_crit_family Bot_F Inf_F Bot_G Q entails_q Inf_G_q Red_Inf_q Red_F_q \_F_q \_Inf_q Inf_FL for Bot_F :: "'f set" and Inf_F :: "'f inference set" and Bot_G :: "'g set" and Q :: "'q set" and entails_q :: "'q \ 'g set \ 'g set \ bool" and Inf_G_q :: \'q \ 'g inference set\ and Red_Inf_q :: "'q \ 'g set \ 'g inference set" and Red_F_q :: "'q \ 'g set \ 'g set" and \_F_q :: "'q \ 'f \ 'g set" and \_Inf_q :: "'q \ 'f inference \ 'g inference set option" and Inf_FL :: \('f \ 'l) inference set\ + fixes Equiv_F :: "'f \ 'f \ bool" (infix "\" 50) and Prec_F :: "'f \ 'f \ bool" (infix "\\" 50) and Prec_l :: "'l \ 'l \ bool" (infix "\l" 50) and active :: "'l" assumes equiv_equiv_F: "equivp (\)" and wf_prec_F: "minimal_element (\\) UNIV" and wf_prec_l: "minimal_element (\l) UNIV" and compat_equiv_prec: "C1 \ D1 \ C2 \ D2 \ C1 \\ C2 \ D1 \\ D2" and equiv_F_grounding: "q \ Q \ C1 \ C2 \ \_F_q q C1 \ \_F_q q C2" and prec_F_grounding: "q \ Q \ C2 \\ C1 \ \_F_q q C1 \ \_F_q q C2" and active_minimal: "l2 \ active \ active \l l2" and at_least_two_labels: "\l2. active \l l2" and inf_never_active: "\ \ Inf_FL \ snd (concl_of \) \ active" and static_ref_comp: "static_refutational_complete_calculus Bot_F Inf_F (\\\) no_labels.Red_Inf_\_Q no_labels.Red_F_\_empty" begin abbreviation Prec_eq_F :: "'f \ 'f \ bool" (infix "\\" 50) where "C \\ D \ C \ D \ C \\ D" definition Prec_FL :: "('f \ 'l) \ ('f \ 'l) \ bool" (infix "\" 50) where "Cl1 \ Cl2 \ fst Cl1 \\ fst Cl2 \ (fst Cl1 \ fst Cl2 \ snd Cl1 \l snd Cl2)" lemma irrefl_prec_F: "\ C \\ C" by (simp add: minimal_element.po[OF wf_prec_F, unfolded po_on_def irreflp_on_def]) lemma trans_prec_F: "C1 \\ C2 \ C2 \\ C3 \ C1 \\ C3" by (auto intro: minimal_element.po[OF wf_prec_F, unfolded po_on_def transp_on_def, THEN conjunct2, simplified, rule_format]) lemma wf_prec_FL: "minimal_element (\) UNIV" proof show "po_on (\) UNIV" unfolding po_on_def proof show "irreflp_on (\) UNIV" unfolding irreflp_on_def Prec_FL_def proof fix Cl assume a_in: "Cl \ (UNIV::('f \ 'l) set)" have "\ (fst Cl \\ fst Cl)" using wf_prec_F minimal_element.min_elt_ex by force moreover have "\ (snd Cl \l snd Cl)" using wf_prec_l minimal_element.min_elt_ex by force ultimately show "\ (fst Cl \\ fst Cl \ fst Cl \ fst Cl \ snd Cl \l snd Cl)" by blast qed next show "transp_on (\) UNIV" unfolding transp_on_def Prec_FL_def proof (simp, intro allI impI) fix C1 l1 C2 l2 C3 l3 assume trans_hyp: "(C1 \\ C2 \ C1 \ C2 \ l1 \l l2) \ (C2 \\ C3 \ C2 \ C3 \ l2 \l l3)" have "C1 \\ C2 \ C2 \ C3 \ C1 \\ C3" using compat_equiv_prec by (metis equiv_equiv_F equivp_def) moreover have "C1 \ C2 \ C2 \\ C3 \ C1 \\ C3" using compat_equiv_prec by (metis equiv_equiv_F equivp_def) moreover have "l1 \l l2 \ l2 \l l3 \ l1 \l l3" using wf_prec_l unfolding minimal_element_def po_on_def transp_on_def by (meson UNIV_I) moreover have "C1 \ C2 \ C2 \ C3 \ C1 \ C3" using equiv_equiv_F by (meson equivp_transp) ultimately show "C1 \\ C3 \ C1 \ C3 \ l1 \l l3" using trans_hyp using trans_prec_F by blast qed qed next show "wfp_on (\) UNIV" unfolding wfp_on_def proof assume contra: "\f. \i. f i \ UNIV \ f (Suc i) \ f i" then obtain f where f_suc: "\i. f (Suc i) \ f i" by blast define R :: "(('f \ 'l) \ ('f \ 'l)) set" where "R = {(Cl1, Cl2). fst Cl1 \\ fst Cl2}" define S :: "(('f \ 'l) \ ('f \ 'l)) set" where "S = {(Cl1, Cl2). fst Cl1 \ fst Cl2 \ snd Cl1 \l snd Cl2}" obtain k where f_chain: "\i. (f (Suc (i + k)), f (i + k)) \ S" proof (atomize_elim, rule wf_infinite_down_chain_compatible[of R f S]) show "wf R" unfolding R_def using wf_app[OF wf_prec_F[unfolded minimal_element_def, THEN conjunct2, unfolded wfp_on_UNIV wfP_def]] by force next show "\i. (f (Suc i), f i) \ R \ S" using f_suc unfolding R_def S_def Prec_FL_def by blast next show "R O S \ R" unfolding R_def S_def using compat_equiv_prec equiv_equiv_F equivp_reflp by fastforce qed define g where "\i. g i = f (i + k)" have g_chain: "\i. (g (Suc i), g i) \ S" unfolding g_def using f_chain by simp have wf_s: "wf S" unfolding S_def by (rule wf_subset[OF wf_app[OF wf_prec_l[unfolded minimal_element_def, THEN conjunct2, unfolded wfp_on_UNIV wfP_def], of snd]]) fast show False using g_chain[unfolded S_def] wf_s[unfolded S_def, folded wfP_def wfp_on_UNIV, unfolded wfp_on_def] by auto qed qed definition active_subset :: "('f \ 'l) set \ ('f \ 'l) set" where "active_subset M = {CL \ M. snd CL = active}" definition passive_subset :: "('f \ 'l) set \ ('f \ 'l) set" where "passive_subset M = {CL \ M. snd CL \ active}" lemma active_subset_insert[simp]: "active_subset (insert Cl N) = (if snd Cl = active then {Cl} else {}) \ active_subset N" unfolding active_subset_def by auto lemma active_subset_union[simp]: "active_subset (M \ N) = active_subset M \ active_subset N" unfolding active_subset_def by auto lemma passive_subset_insert[simp]: "passive_subset (insert Cl N) = (if snd Cl \ active then {Cl} else {}) \ passive_subset N" unfolding passive_subset_def by auto lemma passive_subset_union[simp]: "passive_subset (M \ N) = passive_subset M \ passive_subset N" unfolding passive_subset_def by auto sublocale std?: static_refutational_complete_calculus Bot_FL Inf_FL "(\\\L)" Red_Inf_Q Red_F_Q using labeled_static_ref[OF static_ref_comp] . lemma standard_labeled_lifting_family: assumes q_in: "q \ Q" shows "lifting_with_wf_ordering_family Bot_FL Inf_FL Bot_G (entails_q q) (Inf_G_q q) (Red_Inf_q q) (Red_F_q q) (\_F_L_q q) (\_Inf_L_q q) (\g. Prec_FL)" proof - have "lifting_with_wf_ordering_family Bot_FL Inf_FL Bot_G (entails_q q) (Inf_G_q q) (Red_Inf_q q) (Red_F_q q) (\_F_L_q q) (\_Inf_L_q q) (\g Cl Cl'. False)" using ord_fam_lifted_q[OF q_in] . then have "standard_lifting Bot_FL Inf_FL Bot_G (Inf_G_q q) (entails_q q) (Red_Inf_q q) (Red_F_q q) (\_F_L_q q) (\_Inf_L_q q)" using lifted_q[OF q_in] by blast then show "lifting_with_wf_ordering_family Bot_FL Inf_FL Bot_G (entails_q q) (Inf_G_q q) (Red_Inf_q q) (Red_F_q q) (\_F_L_q q) (\_Inf_L_q q) (\g. Prec_FL)" using wf_prec_FL by (simp add: lifting_with_wf_ordering_family.intro lifting_with_wf_ordering_family_axioms.intro) qed sublocale standard_lifting_with_red_crit_family Inf_FL Bot_G Q Inf_G_q entails_q Red_Inf_q Red_F_q Bot_FL \_F_L_q \_Inf_L_q "\g. Prec_FL" using standard_labeled_lifting_family no_labels.ground.calculus_family_with_red_crit_family_axioms by (simp add: standard_lifting_with_red_crit_family.intro standard_lifting_with_red_crit_family_axioms.intro) notation derive (infix "\RedL" 50) lemma std_Red_Inf_Q_eq: "std.Red_Inf_Q = Red_Inf_\_Q" unfolding Red_Inf_\_q_def Red_Inf_\_L_q_def by simp lemma std_Red_F_Q_eq: "std.Red_F_Q = Red_F_\_empty" unfolding Red_F_\_empty_q_def Red_F_\_empty_L_q_def by simp sublocale static_refutational_complete_calculus Bot_FL Inf_FL "(\\\L)" Red_Inf_Q Red_F_Q by unfold_locales (use static_refutational_complete std_Red_Inf_Q_eq in auto) (* lem:redundant-labeled-inferences *) lemma labeled_red_inf_eq_red_inf: assumes i_in: "\ \ Inf_FL" shows "\ \ Red_Inf_Q N \ to_F \ \ no_labels.Red_Inf_\_Q (fst ` N)" proof assume i_in2: "\ \ Red_Inf_Q N" then have "X \ Red_Inf_\_q ` Q \ \ \ X N" for X unfolding Red_Inf_Q_def by blast obtain X0 where "X0 \ Red_Inf_\_q ` Q" using Q_nonempty by blast then obtain q0 where x0_is: "X0 N = Red_Inf_\_q q0 N" by blast then obtain Y0 where y0_is: "Y0 (fst ` N) = to_F ` (X0 N)" by auto have "Y0 (fst ` N) = no_labels.Red_Inf_\_q q0 (fst ` N)" unfolding y0_is proof show "to_F ` X0 N \ no_labels.Red_Inf_\_q q0 (fst ` N)" proof fix \0 assume i0_in: "\0 \ to_F ` X0 N" then have i0_in2: "\0 \ to_F ` Red_Inf_\_q q0 N" using x0_is by argo then obtain \0_FL where i0_FL_in: "\0_FL \ Inf_FL" and i0_to_i0_FL: "\0 = to_F \0_FL" and subs1: "((\_Inf_L_q q0 \0_FL) \ None \ the (\_Inf_L_q q0 \0_FL) \ Red_Inf_q q0 (\_set_q q0 N)) \ ((\_Inf_L_q q0 \0_FL = None) \ \_F_L_q q0 (concl_of \0_FL) \ \_set_q q0 N \ Red_F_q q0 (\_set_q q0 N))" unfolding Red_Inf_\_q_def by blast have concl_swap: "fst (concl_of \0_FL) = concl_of \0" unfolding concl_of_def i0_to_i0_FL to_F_def by simp have i0_in3: "\0 \ Inf_F" using i0_to_i0_FL Inf_FL_to_Inf_F[OF i0_FL_in] unfolding to_F_def by blast { assume not_none: "\_Inf_q q0 \0 \ None" and "the (\_Inf_q q0 \0) \ {}" then obtain \1 where i1_in: "\1 \ the (\_Inf_q q0 \0)" by blast have "the (\_Inf_q q0 \0) \ Red_Inf_q q0 (no_labels.\_set_q q0 (fst ` N))" using subs1 i0_to_i0_FL not_none by auto } moreover { assume is_none: "\_Inf_q q0 \0 = None" then have "\_F_q q0 (concl_of \0) \ no_labels.\_set_q q0 (fst ` N) \ Red_F_q q0 (no_labels.\_set_q q0 (fst ` N))" using subs1 i0_to_i0_FL concl_swap by simp } ultimately show "\0 \ no_labels.Red_Inf_\_q q0 (fst ` N)" unfolding no_labels.Red_Inf_\_q_def using i0_in3 by auto qed next show "no_labels.Red_Inf_\_q q0 (fst ` N) \ to_F ` X0 N" proof fix \0 assume i0_in: "\0 \ no_labels.Red_Inf_\_q q0 (fst ` N)" then have i0_in2: "\0 \ Inf_F" unfolding no_labels.Red_Inf_\_q_def by blast obtain \0_FL where i0_FL_in: "\0_FL \ Inf_FL" and i0_to_i0_FL: "\0 = to_F \0_FL" using Inf_F_to_Inf_FL[OF i0_in2] unfolding to_F_def by (metis Ex_list_of_length fst_conv inference.exhaust_sel inference.inject map_fst_zip) have concl_swap: "fst (concl_of \0_FL) = concl_of \0" unfolding concl_of_def i0_to_i0_FL to_F_def by simp have subs1: "((\_Inf_L_q q0 \0_FL) \ None \ the (\_Inf_L_q q0 \0_FL) \ Red_Inf_q q0 (\_set_q q0 N)) \ ((\_Inf_L_q q0 \0_FL = None) \ \_F_L_q q0 (concl_of \0_FL) \ (\_set_q q0 N \ Red_F_q q0 (\_set_q q0 N)))" using i0_in i0_to_i0_FL concl_swap unfolding no_labels.Red_Inf_\_q_def by simp then have "\0_FL \ Red_Inf_\_q q0 N" using i0_FL_in unfolding Red_Inf_\_q_def by simp then show "\0 \ to_F ` X0 N" using x0_is i0_to_i0_FL i0_in2 by blast qed qed then have "Y \ no_labels.Red_Inf_\_q ` Q \ to_F \ \ Y (fst ` N)" for Y using i_in2 no_labels.Red_Inf_Q_def std_Red_Inf_Q_eq red_inf_impl by force then show "to_F \ \ no_labels.Red_Inf_\_Q (fst ` N)" unfolding Red_Inf_Q_def no_labels.Red_Inf_\_Q_def by blast next assume to_F_in: "to_F \ \ no_labels.Red_Inf_\_Q (fst ` N)" have imp_to_F: "X \ no_labels.Red_Inf_\_q ` Q \ to_F \ \ X (fst ` N)" for X using to_F_in unfolding no_labels.Red_Inf_\_Q_def by blast then have to_F_in2: "to_F \ \ no_labels.Red_Inf_\_q q (fst ` N)" if "q \ Q" for q using that by auto have "Red_Inf_\_q q N = {\0_FL \ Inf_FL. to_F \0_FL \ no_labels.Red_Inf_\_q q (fst ` N)}" for q proof show "Red_Inf_\_q q N \ {\0_FL \ Inf_FL. to_F \0_FL \ no_labels.Red_Inf_\_q q (fst ` N)}" proof fix q0 \1 assume i1_in: "\1 \ Red_Inf_\_q q0 N" have i1_in2: "\1 \ Inf_FL" using i1_in unfolding Red_Inf_\_q_def by blast then have to_F_i1_in: "to_F \1 \ Inf_F" using Inf_FL_to_Inf_F unfolding to_F_def by simp have concl_swap: "fst (concl_of \1) = concl_of (to_F \1)" unfolding concl_of_def to_F_def by simp then have i1_to_F_in: "to_F \1 \ no_labels.Red_Inf_\_q q0 (fst ` N)" using i1_in to_F_i1_in unfolding Red_Inf_\_q_def no_labels.Red_Inf_\_q_def by force show "\1 \ {\0_FL \ Inf_FL. to_F \0_FL \ no_labels.Red_Inf_\_q q0 (fst ` N)}" using i1_in2 i1_to_F_in by blast qed next show "{\0_FL \ Inf_FL. to_F \0_FL \ no_labels.Red_Inf_\_q q (fst ` N)} \ Red_Inf_\_q q N" proof fix q0 \1 assume i1_in: "\1 \ {\0_FL \ Inf_FL. to_F \0_FL \ no_labels.Red_Inf_\_q q0 (fst ` N)}" then have i1_in2: "\1 \ Inf_FL" by blast then have to_F_i1_in: "to_F \1 \ Inf_F" using Inf_FL_to_Inf_F unfolding to_F_def by simp have concl_swap: "fst (concl_of \1) = concl_of (to_F \1)" unfolding concl_of_def to_F_def by simp then have "((\_Inf_L_q q0 \1) \ None \ the (\_Inf_L_q q0 \1) \ Red_Inf_q q0 (\_set_q q0 N)) \ (\_Inf_L_q q0 \1 = None \ \_F_L_q q0 (concl_of \1) \ \_set_q q0 N \ Red_F_q q0 (\_set_q q0 N))" using i1_in unfolding no_labels.Red_Inf_\_q_def by auto then show "\1 \ Red_Inf_\_q q0 N" using i1_in2 unfolding Red_Inf_\_q_def by blast qed qed then have "\ \ Red_Inf_\_q q N" if "q \ Q" for q using that to_F_in2 i_in unfolding Red_Inf_\_q_def no_labels.Red_Inf_\_q_def by auto then show "\ \ Red_Inf_\_Q N" unfolding Red_Inf_\_Q_def by blast qed (* lem:redundant-labeled-formulas *) lemma red_labeled_clauses: assumes \C \ no_labels.Red_F_\_empty (fst ` N) \ (\C' \ fst ` N. C' \\ C) \ (\(C', L') \ N. L' \l L \ C' \\ C)\ shows \(C, L) \ Red_F_Q N\ proof - note assms moreover have i: \C \ no_labels.Red_F_\_empty (fst ` N) \ (C, L) \ Red_F_Q N\ proof - assume "C \ no_labels.Red_F_\_empty (fst ` N)" then have "C \ no_labels.Red_F_\_empty_q q (fst ` N)" if "q \ Q" for q unfolding no_labels.Red_F_\_empty_def using that by fast then have g_in_red: "\_F_q q C \ Red_F_q q (no_labels.\_set_q q (fst ` N))" if "q \ Q" for q unfolding no_labels.Red_F_\_empty_q_def using that by blast have "\_F_L_q q (C, L) \ Red_F_q q (\_set_q q N)" if "q \ Q" for q using that g_in_red by simp then show ?thesis unfolding Red_F_Q_def Red_F_\_q_g_def by blast qed moreover have ii: \\C' \ fst ` N. C' \\ C \ (C, L) \ Red_F_Q N\ proof - assume "\C' \ fst ` N. C' \\ C" then obtain C' where c'_in: "C' \ fst ` N" and c_prec_c': "C' \\ C" by blast obtain L' where c'_l'_in: "(C', L') \ N" using c'_in by auto have c'_l'_prec: "(C', L') \ (C, L)" using c_prec_c' unfolding Prec_FL_def by simp have c_in_c'_g: "\_F_q q C \ \_F_q q C'" if "q \ Q" for q using prec_F_grounding[OF that c_prec_c'] by presburger then have "\_F_L_q q (C, L) \ \_F_L_q q (C', L')" if "q \ Q" for q using that by auto then have "(C, L) \ Red_F_\_q_g q N" if "q \ Q" for q unfolding Red_F_\_q_g_def using that c'_l'_in c'_l'_prec by blast then show ?thesis unfolding Red_F_Q_def by blast qed moreover have iii: \\(C', L') \ N. L' \l L \ C' \\ C \ (C, L) \ Red_F_Q N\ proof - assume "\(C', L') \ N. L' \l L \ C' \\ C" then obtain C' L' where c'_l'_in: "(C', L') \ N" and l'_sub_l: "L' \l L" and c'_sub_c: "C' \\ C" by fast have "(C, L) \ Red_F_Q N" if "C' \\ C" using that c'_l'_in ii by fastforce moreover { assume equiv_c_c': "C \ C'" then have equiv_c'_c: "C' \ C" using equiv_equiv_F by (simp add: equivp_symp) then have c'_l'_prec: "(C', L') \ (C, L)" using l'_sub_l unfolding Prec_FL_def by simp have "\_F_q q C = \_F_q q C'" if "q \ Q" for q using that equiv_F_grounding equiv_c_c' equiv_c'_c by (simp add: set_eq_subset) then have "\_F_L_q q (C, L) = \_F_L_q q (C', L')" if "q \ Q" for q using that by auto then have "(C, L) \ Red_F_\_q_g q N" if "q \ Q" for q unfolding Red_F_\_q_g_def using that c'_l'_in c'_l'_prec by blast then have ?thesis unfolding Red_F_Q_def by blast } ultimately show ?thesis using c'_sub_c equiv_equiv_F equivp_symp by fastforce qed ultimately show ?thesis by blast qed end subsection \Given Clause Architecture\ locale given_clause = prover_architecture_basis Bot_F Inf_F Bot_G Q entails_q Inf_G_q Red_Inf_q Red_F_q \_F_q \_Inf_q Inf_FL Equiv_F Prec_F Prec_l active for Bot_F :: "'f set" and Inf_F :: "'f inference set" and Bot_G :: "'g set" and Q :: "'q set" and entails_q :: "'q \ 'g set \ 'g set \ bool" and Inf_G_q :: \'q \ 'g inference set\ and Red_Inf_q :: "'q \ 'g set \ 'g inference set" and Red_F_q :: "'q \ 'g set \ 'g set" and \_F_q :: "'q \ 'f \ 'g set" and \_Inf_q :: "'q \ 'f inference \ 'g inference set option" and Inf_FL :: \('f \ 'l) inference set\ and Equiv_F :: "'f \ 'f \ bool" (infix "\" 50) and Prec_F :: "'f \ 'f \ bool" (infix "\\" 50) and Prec_l :: "'l \ 'l \ bool" (infix "\l" 50) and active :: 'l + assumes inf_have_prems: "\F \ Inf_F \ prems_of \F \ []" begin lemma labeled_inf_have_prems: "\ \ Inf_FL \ prems_of \ \ []" using inf_have_prems Inf_FL_to_Inf_F by fastforce inductive step :: "('f \ 'l) set \ ('f \ 'l) set \ bool" (infix "\GC" 50) where process: "N1 = N \ M \ N2 = N \ M' \ M \ Red_F_Q (N \ M') \ active_subset M' = {} \ N1 \GC N2" | infer: "N1 = N \ {(C, L)} \ N2 = N \ {(C, active)} \ M \ L \ active \ active_subset M = {} \ no_labels.Inf_from2 (fst ` (active_subset N)) {C} \ no_labels.Red_Inf_Q (fst ` (N \ {(C, active)} \ M)) \ N1 \GC N2" lemma one_step_equiv: "N1 \GC N2 \ N1 \RedL N2" proof (cases N1 N2 rule: step.cases) show "N1 \GC N2 \ N1 \GC N2" by blast next fix N M M' assume gc_step: "N1 \GC N2" and n1_is: "N1 = N \ M" and n2_is: "N2 = N \ M'" and m_red: "M \ Red_F_Q (N \ M')" and active_empty: "active_subset M' = {}" have "N1 - N2 \ Red_F_Q N2" using n1_is n2_is m_red by auto then show "N1 \RedL N2" unfolding derive.simps by blast next fix N C L M assume gc_step: "N1 \GC N2" and n1_is: "N1 = N \ {(C, L)}" and not_active: "L \ active" and n2_is: "N2 = N \ {(C, active)} \ M" and active_empty: "active_subset M = {}" have "(C, active) \ N2" using n2_is by auto moreover have "C \\ C" using equiv_equiv_F by (metis equivp_def) moreover have "active \l L" using active_minimal[OF not_active] . ultimately have "{(C, L)} \ Red_F_Q N2" using red_labeled_clauses by blast moreover have "N1 - N2 = {} \ N1 - N2 = {(C, L)}" using n1_is n2_is by blast ultimately have "N1 - N2 \ Red_F_Q N2" using std_Red_F_Q_eq by blast then show "N1 \RedL N2" unfolding derive.simps by blast qed (* lem:gc-derivations-are-red-derivations *) lemma gc_to_red: "chain (\GC) D \ chain (\RedL) D" using one_step_equiv Lazy_List_Chain.chain_mono by blast lemma (in-) all_ex_finite_set: "(\(j::nat)\{0..(n::nat). P j n) \ (\n1 n2. \j\{0.. P j n2 \ n1 = n2) \ finite {n. \j \ {0.. nat \ bool" assume allj_exn: "\j\{0..n. P j n" and uniq_n: "\n1 n2. \j\{0.. P j n2 \ n1 = n2" have "{n. \j \ {0..((\j. {n. P j n}) ` {0..j\{0.. finite {n. \j \ {0..j. {n. P j n}"] by simp have "\j\{0..!n. P j n" using allj_exn uniq_n by blast then have "\j\{0..j \ {0..GC) D" and - init_state: "active_subset (lnth D 0) = {}" and + init_state: "active_subset (lhd D) = {}" and final_state: "passive_subset (Liminf_llist D) = {}" shows "fair D" unfolding fair_def proof fix \ assume i_in: "\ \ Inf_from (Liminf_llist D)" + note lhd_is = lhd_conv_lnth[OF chain_not_lnull[OF deriv]] have i_in_inf_fl: "\ \ Inf_FL" using i_in unfolding Inf_from_def by blast have "Liminf_llist D = active_subset (Liminf_llist D)" using final_state unfolding passive_subset_def active_subset_def by blast then have i_in2: "\ \ Inf_from (active_subset (Liminf_llist D))" using i_in by simp define m where "m = length (prems_of \)" then have m_def_F: "m = length (prems_of (to_F \))" unfolding to_F_def by simp have i_in_F: "to_F \ \ Inf_F" using i_in Inf_FL_to_Inf_F unfolding Inf_from_def to_F_def by blast then have m_pos: "m > 0" using m_def_F using inf_have_prems by blast have exist_nj: "\j \ {0..nj. enat (Suc nj) < llength D \ prems_of \ ! j \ active_subset (lnth D nj) \ (\k. k > nj \ enat k < llength D \ prems_of \ ! j \ active_subset (lnth D k)))" proof clarify fix j assume j_in: "j \ {0.. ! j" using i_in2 unfolding m_def Inf_from_def active_subset_def by (smt Collect_mem_eq Collect_mono_iff atLeastLessThan_iff nth_mem old.prod.exhaust snd_conv) then have "(C, active) \ Liminf_llist D" using j_in i_in unfolding m_def Inf_from_def by force then obtain nj where nj_is: "enat nj < llength D" and c_in2: "(C, active) \ \ (lnth D ` {k. nj \ k \ enat k < llength D})" unfolding Liminf_llist_def using init_state by blast - then have c_in3: "\k. k \ nj \ enat k < llength D \ (C, active) \ (lnth D k)" by blast - have nj_pos: "nj > 0" using init_state c_in2 nj_is unfolding active_subset_def by fastforce + then have c_in3: "\k. k \ nj \ enat k < llength D \ (C, active) \ lnth D k" by blast + have nj_pos: "nj > 0" using init_state c_in2 nj_is + unfolding active_subset_def lhd_is by force obtain nj_min where nj_min_is: "nj_min = (LEAST nj. enat nj < llength D \ (C, active) \ \ (lnth D ` {k. nj \ k \ enat k < llength D}))" by blast then have in_allk: "\k. k \ nj_min \ enat k < llength D \ (C, active) \ (lnth D k)" using c_in3 nj_is c_in2 by (metis (mono_tags, lifting) INT_E LeastI_ex mem_Collect_eq) have njm_smaller_D: "enat nj_min < llength D" using nj_min_is by (smt LeastI_ex \\thesis. (\nj. \enat nj < llength D; (C, active) \ \ (lnth D ` {k. nj \ k \ enat k < llength D})\ \ thesis) \ thesis\) have "nj_min > 0" - using nj_is c_in2 nj_pos nj_min_is + using nj_is c_in2 nj_pos nj_min_is lhd_is by (metis (mono_tags, lifting) Collect_empty_eq \(C, active) \ Liminf_llist D\ \Liminf_llist D = active_subset (Liminf_llist D)\ \\k\nj_min. enat k < llength D \ (C, active) \ lnth D k\ active_subset_def init_state linorder_not_less mem_Collect_eq zero_enat_def chain_length_pos[OF deriv]) then obtain njm_prec where nj_prec_is: "Suc njm_prec = nj_min" using gr0_conv_Suc by auto then have njm_prec_njm: "njm_prec < nj_min" by blast then have njm_prec_njm_enat: "enat njm_prec < enat nj_min" by simp have njm_prec_smaller_d: "njm_prec < llength D" - using HOL.no_atp(15)[OF njm_smaller_D njm_prec_njm_enat] . + using HOL.no_atp(15)[OF njm_smaller_D njm_prec_njm_enat] . have njm_prec_all_suc: "\k>njm_prec. enat k < llength D \ (C, active) \ lnth D k" using nj_prec_is in_allk by simp have notin_njm_prec: "(C, active) \ lnth D njm_prec" proof (rule ccontr) assume "\ (C, active) \ lnth D njm_prec" then have absurd_hyp: "(C, active) \ lnth D njm_prec" by simp have prec_smaller: "enat njm_prec < llength D" using nj_min_is nj_prec_is by (smt LeastI_ex Suc_leD \\thesis. (\nj. \enat nj < llength D; (C, active) \ \ (lnth D ` {k. nj \ k \ enat k < llength D})\ \ thesis) \ thesis\ enat_ord_simps(1) le_eq_less_or_eq le_less_trans) have "(C, active) \ \ (lnth D ` {k. njm_prec \ k \ enat k < llength D})" proof - { fix k assume k_in: "njm_prec \ k \ enat k < llength D" have "k = njm_prec \ (C, active) \ lnth D k" using absurd_hyp by simp moreover have "njm_prec < k \ (C, active) \ lnth D k" using nj_prec_is in_allk k_in by simp ultimately have "(C, active) \ lnth D k" using k_in by fastforce } then show "(C, active) \ \ (lnth D ` {k. njm_prec \ k \ enat k < llength D})" by blast qed then have "enat njm_prec < llength D \ (C, active) \ \ (lnth D ` {k. njm_prec \ k \ enat k < llength D})" using prec_smaller by blast then show False using nj_min_is nj_prec_is Orderings.wellorder_class.not_less_Least njm_prec_njm by blast qed then have notin_active_subs_njm_prec: "(C, active) \ active_subset (lnth D njm_prec)" unfolding active_subset_def by blast then show "\nj. enat (Suc nj) < llength D \ prems_of \ ! j \ active_subset (lnth D nj) \ (\k. k > nj \ enat k < llength D \ prems_of \ ! j \ active_subset (lnth D k))" using c_is njm_prec_all_suc njm_prec_smaller_d by (metis (mono_tags, lifting) active_subset_def mem_Collect_eq nj_prec_is njm_smaller_D snd_conv) qed define nj_set where "nj_set = {nj. (\j\{0.. prems_of \ ! j \ active_subset (lnth D nj) \ (\k. k > nj \ enat k < llength D \ prems_of \ ! j \ active_subset (lnth D k)))}" then have nj_not_empty: "nj_set \ {}" proof - have zero_in: "0 \ {0.. ! 0 \ active_subset (lnth D n0)" and "\k>n0. enat k < llength D \ prems_of \ ! 0 \ active_subset (lnth D k)" using exist_nj by fast then have "n0 \ nj_set" unfolding nj_set_def using zero_in by blast then show "nj_set \ {}" by auto qed have nj_finite: "finite nj_set" using all_ex_finite_set[OF exist_nj] by (metis (no_types, lifting) Suc_ile_eq dual_order.strict_implies_order linorder_neqE_nat nj_set_def) (* the n below in the n-1 from the pen-and-paper proof *) have "\n \ nj_set. \nj \ nj_set. nj \ n" using nj_not_empty nj_finite using Max_ge Max_in by blast then obtain n where n_in: "n \ nj_set" and n_bigger: "\nj \ nj_set. nj \ n" by blast then obtain j0 where j0_in: "j0 \ {0.. ! j0 \ active_subset (lnth D n)" and j0_allin: "(\k. k > n \ enat k < llength D \ prems_of \ ! j0 \ active_subset (lnth D k))" unfolding nj_set_def by blast obtain C0 where C0_is: "prems_of \ ! j0 = (C0, active)" using j0_in using i_in2 unfolding m_def Inf_from_def active_subset_def by (smt Collect_mem_eq Collect_mono_iff atLeastLessThan_iff nth_mem old.prod.exhaust snd_conv) then have C0_prems_i: "(C0, active) \ set (prems_of \)" using in_set_conv_nth j0_in m_def by force have C0_in: "(C0, active) \ (lnth D (Suc n))" using C0_is j0_allin suc_n_length by (simp add: active_subset_def) have C0_notin: "(C0, active) \ (lnth D n)" using C0_is j0_notin unfolding active_subset_def by simp have step_n: "lnth D n \GC lnth D (Suc n)" using deriv chain_lnth_rel n_in unfolding nj_set_def by blast have "\N C L M. (lnth D n = N \ {(C, L)} \ lnth D (Suc n) = N \ {(C, active)} \ M \ L \ active \ active_subset M = {} \ no_labels.Inf_from2 (fst ` (active_subset N)) {C} \ no_labels.Red_Inf_Q (fst ` (N \ {(C, active)} \ M)))" proof - have proc_or_infer: "(\N1 N M N2 M'. lnth D n = N1 \ lnth D (Suc n) = N2 \ N1 = N \ M \ N2 = N \ M' \ M \ Red_F_Q (N \ M') \ active_subset M' = {}) \ (\N1 N C L N2 M. lnth D n = N1 \ lnth D (Suc n) = N2 \ N1 = N \ {(C, L)} \ N2 = N \ {(C, active)} \ M \ L \ active \ active_subset M = {} \ no_labels.Inf_from2 (fst ` (active_subset N)) {C} \ no_labels.Red_Inf_Q (fst ` (N \ {(C, active)} \ M)))" using step.simps[of "lnth D n" "lnth D (Suc n)"] step_n by blast show ?thesis using C0_in C0_notin proc_or_infer j0_in C0_is by (smt Un_iff active_subset_def mem_Collect_eq snd_conv sup_bot.right_neutral) qed then obtain N M L where inf_from_subs: "no_labels.Inf_from2 (fst ` (active_subset N)) {C0} \ no_labels.Red_Inf_Q (fst ` (N \ {(C0, active)} \ M))" and nth_d_is: "lnth D n = N \ {(C0, L)}" and suc_nth_d_is: "lnth D (Suc n) = N \ {(C0, active)} \ M" and l_not_active: "L \ active" using C0_in C0_notin j0_in C0_is using active_subset_def by fastforce have "j \ {0.. prems_of \ ! j \ prems_of \ ! j0 \ prems_of \ ! j \ (active_subset N)" for j proof - fix j assume j_in: "j \ {0.. ! j \ prems_of \ ! j0" obtain nj where nj_len: "enat (Suc nj) < llength D" and nj_prems: "prems_of \ ! j \ active_subset (lnth D nj)" and nj_greater: "(\k. k > nj \ enat k < llength D \ prems_of \ ! j \ active_subset (lnth D k))" using exist_nj j_in by blast then have "nj \ nj_set" unfolding nj_set_def using j_in by blast moreover have "nj \ n" proof (rule ccontr) assume "\ nj \ n" then have "prems_of \ ! j = (C0, active)" using C0_in C0_notin step.simps[of "lnth D n" "lnth D (Suc n)"] step_n by (smt Un_iff nth_d_is suc_nth_d_is l_not_active active_subset_def insertCI insertE lessI mem_Collect_eq nj_greater nj_prems snd_conv suc_n_length) then show False using j_not_j0 C0_is by simp qed ultimately have "nj < n" using n_bigger by force then have "prems_of \ ! j \ (active_subset (lnth D n))" using nj_greater n_in Suc_ile_eq dual_order.strict_implies_order unfolding nj_set_def by blast then show "prems_of \ ! j \ (active_subset N)" using nth_d_is l_not_active unfolding active_subset_def by force qed then have "set (prems_of \) \ active_subset N \ {(C0, active)}" using C0_prems_i C0_is m_def by (metis Un_iff atLeast0LessThan in_set_conv_nth insertCI lessThan_iff subrelI) moreover have "\ (set (prems_of \) \ active_subset N - {(C0, active)})" using C0_prems_i by blast ultimately have "\ \ Inf_from2 (active_subset N) {(C0, active)}" using i_in_inf_fl unfolding Inf_from2_def Inf_from_def by blast then have "to_F \ \ no_labels.Inf_from2 (fst ` (active_subset N)) {C0}" unfolding to_F_def Inf_from2_def Inf_from_def no_labels.Inf_from2_def no_labels.Inf_from_def using Inf_FL_to_Inf_F by force then have "to_F \ \ no_labels.Red_Inf_Q (fst ` (lnth D (Suc n)))" using suc_nth_d_is inf_from_subs by fastforce then have "\q \ Q. (\_Inf_q q (to_F \) \ None \ the (\_Inf_q q (to_F \)) \ Red_Inf_q q (\ (\_F_q q ` fst ` lnth D (Suc n)))) \ (\_Inf_q q (to_F \) = None \ \_F_q q (concl_of (to_F \)) \ \ (\_F_q q ` fst ` lnth D (Suc n)) \ Red_F_q q (\ (\_F_q q ` fst ` lnth D (Suc n))))" unfolding to_F_def no_labels.Red_Inf_Q_def no_labels.Red_Inf_\_q_def by blast then have "\ \ Red_Inf_\_Q (lnth D (Suc n))" using i_in_inf_fl unfolding Red_Inf_\_Q_def Red_Inf_\_q_def by (simp add: to_F_def) then show "\ \ Sup_llist (lmap Red_Inf_\_Q D)" unfolding Sup_llist_def using suc_n_length by auto qed theorem gc_complete_Liminf: assumes deriv: "chain (\GC) D" and - init_state: "active_subset (lnth D 0) = {}" and + init_state: "active_subset (lhd D) = {}" and final_state: "passive_subset (Liminf_llist D) = {}" and b_in: "B \ Bot_F" and - bot_entailed: "no_labels.entails_\_Q (fst ` lnth D 0) {B}" + bot_entailed: "no_labels.entails_\_Q (fst ` lhd D) {B}" shows "\BL \ Bot_FL. BL \ Liminf_llist D" proof - + note lhd_is = lhd_conv_lnth[OF chain_not_lnull[OF deriv]] have labeled_b_in: "(B, active) \ Bot_FL" using b_in by simp - have labeled_bot_entailed: "entails_\_L_Q (lnth D 0) {(B, active)}" - using labeled_entailment_lifting bot_entailed by fastforce + have labeled_bot_entailed: "entails_\_L_Q (lhd D) {(B, active)}" + using labeled_entailment_lifting bot_entailed lhd_is by fastforce have fair: "fair D" using gc_fair[OF deriv init_state final_state] . then show ?thesis using dynamic_refutational_complete_Liminf[OF labeled_b_in gc_to_red[OF deriv] fair labeled_bot_entailed] by blast qed (* thm:gc-completeness *) theorem gc_complete: assumes deriv: "chain (\GC) D" and - init_state: "active_subset (lnth D 0) = {}" and + init_state: "active_subset (lhd D) = {}" and final_state: "passive_subset (Liminf_llist D) = {}" and b_in: "B \ Bot_F" and - bot_entailed: "no_labels.entails_\_Q (fst ` (lnth D 0)) {B}" + bot_entailed: "no_labels.entails_\_Q (fst ` lhd D) {B}" shows "\i. enat i < llength D \ (\BL \ Bot_FL. BL \ lnth D i)" proof - + note lhd_is = lhd_conv_lnth[OF chain_not_lnull[OF deriv]] have "\BL\Bot_FL. BL \ Liminf_llist D" using assms by (rule gc_complete_Liminf) then show ?thesis unfolding Liminf_llist_def by auto qed end subsection \Lazy Given Clause Architecture\ locale lazy_given_clause = prover_architecture_basis Bot_F Inf_F Bot_G Q entails_q Inf_G_q Red_Inf_q Red_F_q \_F_q \_Inf_q Inf_FL Equiv_F Prec_F Prec_l active for Bot_F :: "'f set" and Inf_F :: "'f inference set" and Bot_G :: "'g set" and Q :: "'q set" and entails_q :: "'q \ 'g set \ 'g set \ bool" and Inf_G_q :: \'q \ 'g inference set\ and Red_Inf_q :: "'q \ 'g set \ 'g inference set" and Red_F_q :: "'q \ 'g set \ 'g set" and \_F_q :: "'q \ 'f \ 'g set" and \_Inf_q :: "'q \ 'f inference \ 'g inference set option" and Inf_FL :: \('f \ 'l) inference set\ and Equiv_F :: "'f \ 'f \ bool" (infix "\" 50) and Prec_F :: "'f \ 'f \ bool" (infix "\\" 50) and Prec_l :: "'l \ 'l \ bool" (infix "\l" 50) and active :: 'l begin inductive step :: "'f inference set \ ('f \ 'l) set \ 'f inference set \ ('f \ 'l) set \ bool" (infix "\LGC" 50) where process: "N1 = N \ M \ N2 = N \ M' \ M \ Red_F_Q (N \ M') \ active_subset M' = {} \ (T, N1) \LGC (T, N2)" | schedule_infer: "T2 = T1 \ T' \ N1 = N \ {(C, L)} \ N2 = N \ {(C, active)} \ L \ active \ T' = no_labels.Inf_from2 (fst ` (active_subset N)) {C} \ (T1, N1) \LGC (T2, N2)" | compute_infer: "T1 = T2 \ {\} \ N2 = N1 \ M \ active_subset M = {} \ \ \ no_labels.Red_Inf_Q (fst ` (N1 \ M)) \ (T1, N1) \LGC (T2, N2)" | delete_orphans: "T1 = T2 \ T' \ T' \ no_labels.Inf_from (fst ` (active_subset N)) = {} \ (T1, N) \LGC (T2, N)" lemma premise_free_inf_always_from: "\ \ Inf_F \ length (prems_of \) = 0 \ \ \ no_labels.Inf_from N" unfolding no_labels.Inf_from_def by simp lemma one_step_equiv: "(T1, N1) \LGC (T2, N2) \ N1 \RedL N2" proof (cases "(T1, N1)" "(T2, N2)" rule: step.cases) show "(T1, N1) \LGC (T2, N2) \ (T1, N1) \LGC (T2, N2)" by blast next fix N M M' assume n1_is: "N1 = N \ M" and n2_is: "N2 = N \ M'" and m_red: "M \ Red_F_Q (N \ M')" have "N1 - N2 \ Red_F_Q N2" using n1_is n2_is m_red by auto then show "N1 \RedL N2" unfolding derive.simps by blast next fix N C L M assume n1_is: "N1 = N \ {(C, L)}" and not_active: "L \ active" and n2_is: "N2 = N \ {(C, active)}" have "(C, active) \ N2" using n2_is by auto moreover have "C \\ C" by (metis equivp_def equiv_equiv_F) moreover have "active \l L" using active_minimal[OF not_active] . ultimately have "{(C, L)} \ Red_F_Q N2" using red_labeled_clauses by blast then have "N1 - N2 \ Red_F_Q N2" using std_Red_F_Q_eq using n1_is n2_is by blast then show "N1 \RedL N2" unfolding derive.simps by blast next fix M assume n2_is: "N2 = N1 \ M" have "N1 - N2 \ Red_F_Q N2" using n2_is by blast then show "N1 \RedL N2" unfolding derive.simps by blast next assume n2_is: "N2 = N1" have "N1 - N2 \ Red_F_Q N2" using n2_is by blast then show "N1 \RedL N2" unfolding derive.simps by blast qed (* lem:lgc-derivations-are-red-derivations *) lemma lgc_to_red: "chain (\LGC) D \ chain (\RedL) (lmap snd D)" using one_step_equiv Lazy_List_Chain.chain_mono by (smt chain_lmap prod.collapse) (* lem:fair-lgc-derivations *) lemma lgc_fair: assumes deriv: "chain (\LGC) D" and - init_state: "active_subset (snd (lnth D 0)) = {}" and + init_state: "active_subset (snd (lhd D)) = {}" and final_state: "passive_subset (Liminf_llist (lmap snd D)) = {}" and - no_prems_init_active: "\\ \ Inf_F. length (prems_of \) = 0 \ \ \ fst (lnth D 0)" and + no_prems_init_active: "\\ \ Inf_F. length (prems_of \) = 0 \ \ \ fst (lhd D)" and final_schedule: "Liminf_llist (lmap fst D) = {}" shows "fair (lmap snd D)" unfolding fair_def proof fix \ assume i_in: "\ \ Inf_from (Liminf_llist (lmap snd D))" + note lhd_is = lhd_conv_lnth[OF chain_not_lnull[OF deriv]] have i_in_inf_fl: "\ \ Inf_FL" using i_in unfolding Inf_from_def by blast have "Liminf_llist (lmap snd D) = active_subset (Liminf_llist (lmap snd D))" using final_state unfolding passive_subset_def active_subset_def by blast then have i_in2: "\ \ Inf_from (active_subset (Liminf_llist (lmap snd D)))" using i_in by simp define m where "m = length (prems_of \)" then have m_def_F: "m = length (prems_of (to_F \))" unfolding to_F_def by simp have i_in_F: "to_F \ \ Inf_F" using i_in Inf_FL_to_Inf_F unfolding Inf_from_def to_F_def by blast have exist_nj: "\j \ {0..nj. enat (Suc nj) < llength D \ prems_of \ ! j \ active_subset (snd (lnth D nj)) \ (\k. k > nj \ enat k < llength D \ prems_of \ ! j \ active_subset (snd (lnth D k))))" proof clarify fix j assume j_in: "j \ {0.. ! j" using i_in2 unfolding m_def Inf_from_def active_subset_def by (smt Collect_mem_eq Collect_mono_iff atLeastLessThan_iff nth_mem old.prod.exhaust snd_conv) then have "(C, active) \ Liminf_llist (lmap snd D)" using j_in i_in unfolding m_def Inf_from_def by force then obtain nj where nj_is: "enat nj < llength D" and c_in2: "(C, active) \ \ (snd ` (lnth D ` {k. nj \ k \ enat k < llength D}))" unfolding Liminf_llist_def using init_state by fastforce then have c_in3: "\k. k \ nj \ enat k < llength D \ (C, active) \ snd (lnth D k)" by blast - have nj_pos: "nj > 0" using init_state c_in2 nj_is unfolding active_subset_def by fastforce + have nj_pos: "nj > 0" using init_state c_in2 nj_is unfolding active_subset_def lhd_is by fastforce obtain nj_min where nj_min_is: "nj_min = (LEAST nj. enat nj < llength D \ (C, active) \ \ (snd ` (lnth D ` {k. nj \ k \ enat k < llength D})))" by blast then have in_allk: "\k. k \ nj_min \ enat k < llength D \ (C, active) \ snd (lnth D k)" using c_in3 nj_is c_in2 INT_E LeastI_ex by (smt INT_iff INT_simps(10) c_is image_eqI mem_Collect_eq) have njm_smaller_D: "enat nj_min < llength D" using nj_min_is by (smt LeastI_ex \\thesis. (\nj. \enat nj < llength D; (C, active) \ \ (snd ` (lnth D ` {k. nj \ k \ enat k < llength D}))\ \ thesis) \ thesis\) have "nj_min > 0" - using nj_is c_in2 nj_pos nj_min_is + using nj_is c_in2 nj_pos nj_min_is lhd_is by (metis (mono_tags, lifting) active_subset_def emptyE in_allk init_state mem_Collect_eq not_less snd_conv zero_enat_def chain_length_pos[OF deriv]) then obtain njm_prec where nj_prec_is: "Suc njm_prec = nj_min" using gr0_conv_Suc by auto then have njm_prec_njm: "njm_prec < nj_min" by blast then have njm_prec_njm_enat: "enat njm_prec < enat nj_min" by simp have njm_prec_smaller_d: "njm_prec < llength D" using HOL.no_atp(15)[OF njm_smaller_D njm_prec_njm_enat] . have njm_prec_all_suc: "\k>njm_prec. enat k < llength D \ (C, active) \ snd (lnth D k)" using nj_prec_is in_allk by simp have notin_njm_prec: "(C, active) \ snd (lnth D njm_prec)" proof (rule ccontr) assume "\ (C, active) \ snd (lnth D njm_prec)" then have absurd_hyp: "(C, active) \ snd (lnth D njm_prec)" by simp have prec_smaller: "enat njm_prec < llength D" using nj_min_is nj_prec_is by (smt LeastI_ex Suc_leD \\thesis. (\nj. \enat nj < llength D; (C, active) \ \ (snd ` (lnth D ` {k. nj \ k \ enat k < llength D}))\ \ thesis) \ thesis\ enat_ord_simps(1) le_eq_less_or_eq le_less_trans) have "(C, active) \ \ (snd ` (lnth D ` {k. njm_prec \ k \ enat k < llength D}))" proof - { fix k assume k_in: "njm_prec \ k \ enat k < llength D" have "k = njm_prec \ (C, active) \ snd (lnth D k)" using absurd_hyp by simp moreover have "njm_prec < k \ (C, active) \ snd (lnth D k)" using nj_prec_is in_allk k_in by simp ultimately have "(C, active) \ snd (lnth D k)" using k_in by fastforce } then show "(C, active) \ \ (snd ` (lnth D ` {k. njm_prec \ k \ enat k < llength D}))" by blast qed then have "enat njm_prec < llength D \ (C, active) \ \ (snd ` (lnth D ` {k. njm_prec \ k \ enat k < llength D}))" using prec_smaller by blast then show False using nj_min_is nj_prec_is Orderings.wellorder_class.not_less_Least njm_prec_njm by blast qed then have notin_active_subs_njm_prec: "(C, active) \ active_subset (snd (lnth D njm_prec))" unfolding active_subset_def by blast then show "\nj. enat (Suc nj) < llength D \ prems_of \ ! j \ active_subset (snd (lnth D nj)) \ (\k. k > nj \ enat k < llength D \ prems_of \ ! j \ active_subset (snd (lnth D k)))" using c_is njm_prec_all_suc njm_prec_smaller_d by (metis (mono_tags, lifting) active_subset_def mem_Collect_eq nj_prec_is njm_smaller_D snd_conv) qed define nj_set where "nj_set = {nj. (\j\{0.. prems_of \ ! j \ active_subset (snd (lnth D nj)) \ (\k. k > nj \ enat k < llength D \ prems_of \ ! j \ active_subset (snd (lnth D k))))}" { assume m_null: "m = 0" - then have "enat 0 < llength D \ to_F \ \ fst (lnth D 0)" + then have "enat 0 < llength D \ to_F \ \ fst (lhd D)" using no_prems_init_active i_in_F m_def_F zero_enat_def chain_length_pos[OF deriv] by auto then have "\n. enat n < llength D \ to_F \ \ fst (lnth D n)" - by blast + unfolding lhd_is by blast } moreover { assume m_pos: "m > 0" have nj_not_empty: "nj_set \ {}" proof - have zero_in: "0 \ {0.. ! 0 \ active_subset (snd (lnth D n0))" and "\k>n0. enat k < llength D \ prems_of \ ! 0 \ active_subset (snd (lnth D k))" using exist_nj by fast then have "n0 \ nj_set" unfolding nj_set_def using zero_in by blast then show "nj_set \ {}" by auto qed have nj_finite: "finite nj_set" using all_ex_finite_set[OF exist_nj] by (metis (no_types, lifting) Suc_ile_eq dual_order.strict_implies_order linorder_neqE_nat nj_set_def) have "\n \ nj_set. \nj \ nj_set. nj \ n" using nj_not_empty nj_finite using Max_ge Max_in by blast then obtain n where n_in: "n \ nj_set" and n_bigger: "\nj \ nj_set. nj \ n" by blast then obtain j0 where j0_in: "j0 \ {0.. ! j0 \ active_subset (snd (lnth D n))" and j0_allin: "(\k. k > n \ enat k < llength D \ prems_of \ ! j0 \ active_subset (snd (lnth D k)))" unfolding nj_set_def by blast obtain C0 where C0_is: "prems_of \ ! j0 = (C0, active)" using j0_in i_in2 unfolding m_def Inf_from_def active_subset_def by (smt Collect_mem_eq Collect_mono_iff atLeastLessThan_iff nth_mem old.prod.exhaust snd_conv) then have C0_prems_i: "(C0, active) \ set (prems_of \)" using in_set_conv_nth j0_in m_def by force have C0_in: "(C0, active) \ (snd (lnth D (Suc n)))" using C0_is j0_allin suc_n_length by (simp add: active_subset_def) have C0_notin: "(C0, active) \ (snd (lnth D n))" using C0_is j0_notin unfolding active_subset_def by simp have step_n: "lnth D n \LGC lnth D (Suc n)" using deriv chain_lnth_rel n_in unfolding nj_set_def by blast have is_scheduled: "\T2 T1 T' N1 N C L N2. lnth D n = (T1, N1) \ lnth D (Suc n) = (T2, N2) \ T2 = T1 \ T' \ N1 = N \ {(C, L)} \ N2 = N \ {(C, active)} \ L \ active \ T' = no_labels.Inf_from2 (fst ` active_subset N) {C}" using step.simps[of "lnth D n" "lnth D (Suc n)"] step_n C0_in C0_notin unfolding active_subset_def by fastforce then obtain T2 T1 T' N1 N L N2 where nth_d_is: "lnth D n = (T1, N1)" and suc_nth_d_is: "lnth D (Suc n) = (T2, N2)" and t2_is: "T2 = T1 \ T'" and n1_is: "N1 = N \ {(C0, L)}" "N2 = N \ {(C0, active)}" and l_not_active: "L \ active" and tp_is: "T' = no_labels.Inf_from2 (fst ` active_subset N) {C0}" using C0_in C0_notin j0_in C0_is using active_subset_def by fastforce have "j \ {0.. prems_of \ ! j \ prems_of \ ! j0 \ prems_of \ ! j \ (active_subset N)" for j proof - fix j assume j_in: "j \ {0.. ! j \ prems_of \ ! j0" obtain nj where nj_len: "enat (Suc nj) < llength D" and nj_prems: "prems_of \ ! j \ active_subset (snd (lnth D nj))" and nj_greater: "(\k. k > nj \ enat k < llength D \ prems_of \ ! j \ active_subset (snd (lnth D k)))" using exist_nj j_in by blast then have "nj \ nj_set" unfolding nj_set_def using j_in by blast moreover have "nj \ n" proof (rule ccontr) assume "\ nj \ n" then have "prems_of \ ! j = (C0, active)" using C0_in C0_notin step.simps[of "lnth D n" "lnth D (Suc n)"] step_n active_subset_def is_scheduled nj_greater nj_prems suc_n_length by auto then show False using j_not_j0 C0_is by simp qed ultimately have "nj < n" using n_bigger by force then have "prems_of \ ! j \ (active_subset (snd (lnth D n)))" using nj_greater n_in Suc_ile_eq dual_order.strict_implies_order unfolding nj_set_def by blast then show "prems_of \ ! j \ (active_subset N)" using nth_d_is l_not_active n1_is unfolding active_subset_def by force qed then have prems_i_active: "set (prems_of \) \ active_subset N \ {(C0, active)}" using C0_prems_i C0_is m_def by (metis Un_iff atLeast0LessThan in_set_conv_nth insertCI lessThan_iff subrelI) moreover have "\ (set (prems_of \) \ active_subset N - {(C0, active)})" using C0_prems_i by blast ultimately have "\ \ Inf_from2 (active_subset N) {(C0, active)}" using i_in_inf_fl prems_i_active unfolding Inf_from2_def Inf_from_def by blast then have "to_F \ \ no_labels.Inf_from2 (fst ` (active_subset N)) {C0}" unfolding to_F_def Inf_from2_def Inf_from_def no_labels.Inf_from2_def no_labels.Inf_from_def using Inf_FL_to_Inf_F by force then have i_in_t2: "to_F \ \ T2" using tp_is t2_is by simp have "j \ {0.. (\k. k > n \ enat k < llength D \ prems_of \ ! j \ active_subset (snd (lnth D k)))" for j proof (cases "j = j0") case True assume "j = j0" then show "(\k. k > n \ enat k < llength D \ prems_of \ ! j \ active_subset (snd (lnth D k)))" using j0_allin by simp next case False assume j_in: "j \ {0.. j0" obtain nj where nj_len: "enat (Suc nj) < llength D" and nj_prems: "prems_of \ ! j \ active_subset (snd (lnth D nj))" and nj_greater: "(\k. k > nj \ enat k < llength D \ prems_of \ ! j \ active_subset (snd (lnth D k)))" using exist_nj j_in by blast then have "nj \ nj_set" unfolding nj_set_def using j_in by blast then show "(\k. k > n \ enat k < llength D \ prems_of \ ! j \ active_subset (snd (lnth D k)))" using nj_greater n_bigger by auto qed then have allj_allk: "(\c\ set (prems_of \). (\k. k > n \ enat k < llength D \ c \ active_subset (snd (lnth D k))))" using m_def by (metis atLeast0LessThan in_set_conv_nth lessThan_iff) have "\c\ set (prems_of \). snd c = active" using prems_i_active unfolding active_subset_def by auto then have ex_n_i_in: "\n. enat (Suc n) < llength D \ to_F \ \ fst (lnth D (Suc n)) \ (\c\ set (prems_of \). snd c = active) \ (\c\ set (prems_of \). (\k. k > n \ enat k < llength D \ c \ active_subset (snd (lnth D k))))" using allj_allk i_in_t2 suc_nth_d_is fstI n_in nj_set_def by auto then have "\n. enat n < llength D \ to_F \ \ fst (lnth D n) \ (\c\ set (prems_of \). snd c = active) \ (\c\ set (prems_of \). (\k. k \ n \ enat k < llength D \ c \ active_subset (snd (lnth D k))))" by auto } ultimately obtain n T2 N2 where i_in_suc_n: "to_F \ \ fst (lnth D n)" and all_prems_active_after: "m > 0 \ (\c\ set (prems_of \). (\k. k \ n \ enat k < llength D \ c \ active_subset (snd (lnth D k))))" and suc_n_length: "enat n < llength D" and suc_nth_d_is: "lnth D n = (T2, N2)" by (metis less_antisym old.prod.exhaust zero_less_Suc) then have i_in_t2: "to_F \ \ T2" by simp have "\p\n. enat (Suc p) < llength D \ to_F \ \ (fst (lnth D p)) \ to_F \ \ (fst (lnth D (Suc p)))" proof (rule ccontr) assume contra: "\ (\p\n. enat (Suc p) < llength D \ to_F \ \ (fst (lnth D p)) \ to_F \ \ (fst (lnth D (Suc p))))" then have i_in_suc: "p0 \ n \ enat (Suc p0) < llength D \ to_F \ \ (fst (lnth D p0)) \ to_F \ \ (fst (lnth D (Suc p0)))" for p0 by blast have "p0 \ n \ enat p0 < llength D \ to_F \ \ (fst (lnth D p0))" for p0 proof (induction rule: nat_induct_at_least) case base then show ?case using i_in_t2 suc_nth_d_is by simp next case (Suc p0) assume p_bigger_n: "n \ p0" and induct_hyp: "enat p0 < llength D \ to_F \ \ fst (lnth D p0)" and sucsuc_smaller_d: "enat (Suc p0) < llength D" have suc_p_bigger_n: "n \ p0" using p_bigger_n by simp have suc_smaller_d: "enat p0 < llength D" using sucsuc_smaller_d Suc_ile_eq dual_order.strict_implies_order by blast then have "to_F \ \ fst (lnth D p0)" using induct_hyp by blast then show ?case using i_in_suc[OF suc_p_bigger_n sucsuc_smaller_d] by blast qed then have i_in_all_bigger_n: "\j. j \ n \ enat j < llength D \ to_F \ \ (fst (lnth D j))" by presburger have "llength (lmap fst D) = llength D" by force then have "to_F \ \ \ (lnth (lmap fst D) ` {j. n \ j \ enat j < llength (lmap fst D)})" using i_in_all_bigger_n using Suc_le_D by auto then have "to_F \ \ Liminf_llist (lmap fst D)" unfolding Liminf_llist_def using suc_n_length by auto then show False using final_schedule by fast qed then obtain p where p_greater_n: "p \ n" and p_smaller_d: "enat (Suc p) < llength D" and i_in_p: "to_F \ \ (fst (lnth D p))" and i_notin_suc_p: "to_F \ \ (fst (lnth D (Suc p)))" by blast have p_neq_n: "Suc p \ n" using i_notin_suc_p i_in_suc_n by blast have step_p: "lnth D p \LGC lnth D (Suc p)" using deriv p_smaller_d chain_lnth_rel by blast then have "\T1 T2 \ N2 N1 M. lnth D p = (T1, N1) \ lnth D (Suc p) = (T2, N2) \ T1 = T2 \ {\} \ N2 = N1 \ M \ active_subset M = {} \ \ \ no_labels.Red_Inf_\_Q (fst ` (N1 \ M))" proof - have ci_or_do: "(\T1 T2 \ N2 N1 M. lnth D p = (T1, N1) \ lnth D (Suc p) = (T2, N2) \ T1 = T2 \ {\} \ N2 = N1 \ M \ active_subset M = {} \ \ \ no_labels.Red_Inf_\_Q (fst ` (N1 \ M))) \ (\T1 T2 T' N. lnth D p = (T1, N) \ lnth D (Suc p) = (T2, N) \ T1 = T2 \ T' \ T' \ no_labels.Inf_from (fst ` active_subset N) = {})" using step.simps[of "lnth D p" "lnth D (Suc p)"] step_p i_in_p i_notin_suc_p by fastforce then have p_greater_n_strict: "n < Suc p" using suc_nth_d_is p_greater_n i_in_t2 i_notin_suc_p le_eq_less_or_eq by force have "m > 0 \ j \ {0.. prems_of (to_F \) ! j \ fst ` active_subset (snd (lnth D p))" for j proof - fix j assume m_pos: "m > 0" and j_in: "j \ {0.. ! j \ (active_subset (snd (lnth D p)))" using all_prems_active_after[OF m_pos] p_smaller_d m_def p_greater_n p_neq_n by (meson Suc_ile_eq atLeastLessThan_iff dual_order.strict_implies_order nth_mem p_greater_n_strict) then have "fst (prems_of \ ! j) \ fst ` active_subset (snd (lnth D p))" by blast then show "prems_of (to_F \) ! j \ fst ` active_subset (snd (lnth D p))" unfolding to_F_def using j_in m_def by simp qed then have prems_i_active_p: "m > 0 \ to_F \ \ no_labels.Inf_from (fst ` active_subset (snd (lnth D p)))" using i_in_F unfolding no_labels.Inf_from_def by (smt atLeast0LessThan in_set_conv_nth lessThan_iff m_def_F mem_Collect_eq subsetI) have "m = 0 \ (\T1 T2 \ N2 N1 M. lnth D p = (T1, N1) \ lnth D (Suc p) = (T2, N2) \ T1 = T2 \ {\} \ N2 = N1 \ M \ active_subset M = {} \ \ \ no_labels.Red_Inf_\_Q (fst ` (N1 \ M)))" using ci_or_do premise_free_inf_always_from[of "to_F \" "fst ` active_subset _", OF i_in_F] m_def i_in_p i_notin_suc_p m_def_F by auto then show "(\T1 T2 \ N2 N1 M. lnth D p = (T1, N1) \ lnth D (Suc p) = (T2, N2) \ T1 = T2 \ {\} \ N2 = N1 \ M \ active_subset M = {} \ \ \ no_labels.Red_Inf_\_Q (fst ` (N1 \ M)))" using ci_or_do i_in_p i_notin_suc_p prems_i_active_p unfolding active_subset_def by force qed then obtain T1p T2p N1p N2p Mp where "lnth D p = (T1p, N1p)" and suc_p_is: "lnth D (Suc p) = (T2p, N2p)" and "T1p = T2p \ {to_F \}" and "T2p \ {to_F \} = {}" and n2p_is: "N2p = N1p \ Mp"and "active_subset Mp = {}" and i_in_red_inf: "to_F \ \ no_labels.Red_Inf_\_Q (fst ` (N1p \ Mp))" using i_in_p i_notin_suc_p by fastforce have "to_F \ \ no_labels.Red_Inf_Q (fst ` (snd (lnth D (Suc p))))" using i_in_red_inf suc_p_is n2p_is by fastforce then have "\q \ Q. (\_Inf_q q (to_F \) \ None \ the (\_Inf_q q (to_F \)) \ Red_Inf_q q (\ (\_F_q q ` fst ` snd (lnth D (Suc p))))) \ (\_Inf_q q (to_F \) = None \ \_F_q q (concl_of (to_F \)) \ \ (\_F_q q ` fst ` snd (lnth D (Suc p))) \ Red_F_q q (\ (\_F_q q ` fst ` snd (lnth D (Suc p)))))" unfolding to_F_def no_labels.Red_Inf_Q_def no_labels.Red_Inf_\_q_def by blast then have "\ \ Red_Inf_\_Q (snd (lnth D (Suc p)))" using i_in_inf_fl unfolding Red_Inf_\_Q_def Red_Inf_\_q_def by (simp add: to_F_def) then show "\ \ Sup_llist (lmap Red_Inf_\_Q (lmap snd D))" unfolding Sup_llist_def using suc_n_length p_smaller_d by auto qed theorem lgc_complete_Liminf: assumes deriv: "chain (\LGC) D" and - init_state: "active_subset (snd (lnth D 0)) = {}" and + init_state: "active_subset (snd (lhd D)) = {}" and final_state: "passive_subset (Liminf_llist (lmap snd D)) = {}" and - no_prems_init_active: "\\ \ Inf_F. length (prems_of \) = 0 \ \ \ fst (lnth D 0)" and + no_prems_init_active: "\\ \ Inf_F. length (prems_of \) = 0 \ \ \ fst (lhd D)" and final_schedule: "Liminf_llist (lmap fst D) = {}" and b_in: "B \ Bot_F" and - bot_entailed: "no_labels.entails_\_Q (fst ` (snd (lnth D 0))) {B}" + bot_entailed: "no_labels.entails_\_Q (fst ` snd (lhd D)) {B}" shows "\BL \ Bot_FL. BL \ Liminf_llist (lmap snd D)" proof - have labeled_b_in: "(B, active) \ Bot_FL" using b_in by simp - have simp_snd_lmap: "lnth (lmap snd D) 0 = snd (lnth D 0)" - using lnth_lmap[of 0 D snd] chain_length_pos[OF deriv] by (simp add: zero_enat_def) - have labeled_bot_entailed: "entails_\_L_Q (snd (lnth D 0)) {(B, active)}" + have simp_snd_lmap: "lhd (lmap snd D) = snd (lhd D)" + by (rule llist.map_sel(1)[OF chain_not_lnull[OF deriv]]) + have labeled_bot_entailed: "entails_\_L_Q (snd (lhd D)) {(B, active)}" using labeled_entailment_lifting bot_entailed by fastforce have "fair (lmap snd D)" using lgc_fair[OF deriv init_state final_state no_prems_init_active final_schedule] . then show ?thesis using dynamic_refutational_complete_Liminf labeled_b_in lgc_to_red[OF deriv] labeled_bot_entailed simp_snd_lmap std_Red_Inf_Q_eq by presburger qed (* thm:lgc-completeness *) theorem lgc_complete: assumes deriv: "chain (\LGC) D" and - init_state: "active_subset (snd (lnth D 0)) = {}" and + init_state: "active_subset (snd (lhd D)) = {}" and final_state: "passive_subset (Liminf_llist (lmap snd D)) = {}" and - no_prems_init_active: "\\ \ Inf_F. length (prems_of \) = 0 \ \ \ fst (lnth D 0)" and + no_prems_init_active: "\\ \ Inf_F. length (prems_of \) = 0 \ \ \ fst (lhd D)" and final_schedule: "Liminf_llist (lmap fst D) = {}" and b_in: "B \ Bot_F" and - bot_entailed: "no_labels.entails_\_Q (fst ` snd (lnth D 0)) {B}" + bot_entailed: "no_labels.entails_\_Q (fst ` snd (lhd D)) {B}" shows "\i. enat i < llength D \ (\BL \ Bot_FL. BL \ snd (lnth D i))" proof - have "\BL\Bot_FL. BL \ Liminf_llist (lmap snd D)" using assms by (rule lgc_complete_Liminf) then show ?thesis unfolding Liminf_llist_def by auto qed end end