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,278 +1,283 @@ (* Title: Liminf of Lazy Lists Author: Jasmin Blanchette , 2014, 2017 Author: Dmitriy Traytel , 2014 Maintainer: Jasmin Blanchette *) section \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. \ 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 definition Sup_upto_llist :: "'a set llist \ nat \ 'a set" where "Sup_upto_llist Xs j = (\i \ {i. enat i < llength Xs \ i \ j}. lnth Xs i)" 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) lemma Sup_upto_llist_Suc[simp]: "Sup_upto_llist Xs (Suc j) = Sup_upto_llist Xs 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_mono: "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: "j \ k \ 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 j" unfolding Sup_llist_def Sup_upto_llist_def by blast lemma lnth_subset_Sup_upto_llist: "enat 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 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 k" using elem_Sup_llist_imp_Sup_upto_llist by fast from X obtain k' where k': "X \ Sup_upto_llist Xs k'" using insert.hyps(3) by fast have "insert x X \ Sup_upto_llist Xs (max k k')" using k k' by (metis insert_absorb insert_subset Sup_upto_llist_mono max.cobounded2 max.commute order.trans) then show ?case by fast qed simp 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 \ \ (lnth Xs ` {j. i \ j \ enat j < llength Xs})" 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 \ \ (lnth Xs ` {j. i \ j \ enat j < llength Xs})" 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 \ \ (lnth Xs ` {j. i_of x \ j \ enat j < llength Xs})" 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 \ \ (lnth Xs ` {j. i_max \ j \ enat j < llength Xs})" proof fix x assume x_in: "x \ X" then have x_in_inter: "x \ \ (lnth Xs ` {j. i_of x \ j \ enat j < llength Xs})" 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 \ \ (lnth Xs ` {j. i_max \ j \ enat j < llength Xs})" 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 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,950 @@ (* 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 i" and "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_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}\ shows \\B'\Bot. B' \ Liminf_llist D\ proof - have non_empty: \\ lnull D\ using chain_not_lnull[OF deriv] . have subs: \lnth D 0 \ 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 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} \ \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 "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 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}\ 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" 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 "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 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: \lnth D 0 \ {B}\ have non_empty: \\ lnull D\ using chain_not_lnull[OF deriv] . - have subs: \(lnth D 0) \ Sup_llist D\ + have subs: \lnth D 0 \ 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 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