diff --git a/thys/Saturation_Framework/Labeled_Lifting_to_Non_Ground_Calculi.thy b/thys/Saturation_Framework/Labeled_Lifting_to_Non_Ground_Calculi.thy --- a/thys/Saturation_Framework/Labeled_Lifting_to_Non_Ground_Calculi.thy +++ b/thys/Saturation_Framework/Labeled_Lifting_to_Non_Ground_Calculi.thy @@ -1,394 +1,381 @@ (* Title: Labeled Lifting to Non-Ground Calculi of the Saturation Framework * Author: Sophie Tourret , 2019-2020 *) section \Labeled Liftings\ text \This section formalizes the extension of the lifting results to labeled calculi. This corresponds to section 3.4 of the report.\ theory Labeled_Lifting_to_Non_Ground_Calculi imports Lifting_to_Non_Ground_Calculi begin subsection \Labeled Lifting with a Family of Well-founded Orderings\ locale labeled_lifting_w_wf_ord_family = - lifting_with_wf_ordering_family Bot_F Inf_F Bot_G entails_G Inf_G Red_Inf_G Red_F_G \_F \_Inf Prec_F + no_labels: lifting_with_wf_ordering_family Bot_F Inf_F Bot_G entails_G Inf_G Red_Inf_G Red_F_G \_F + \_Inf Prec_F for Bot_F :: "'f set" and Inf_F :: "'f inference set" and Bot_G :: "'g set" and entails_G :: "'g set \ 'g set \ bool" (infix "\G" 50) and Inf_G :: "'g inference set" and Red_Inf_G :: "'g set \ 'g inference set" and Red_F_G :: "'g set \ 'g set" and \_F :: "'f \ 'g set" and \_Inf :: "'f inference \ 'g inference set option" and Prec_F :: "'g \ 'f \ 'f \ bool" (infix "\" 50) + fixes Inf_FL :: \('f \ 'l) inference set\ assumes Inf_F_to_Inf_FL: \\\<^sub>F \ Inf_F \ length (Ll :: 'l list) = length (prems_of \\<^sub>F) \ \L0. Infer (zip (prems_of \\<^sub>F) Ll) (concl_of \\<^sub>F, L0) \ Inf_FL\ and Inf_FL_to_Inf_F: \\\<^sub>F\<^sub>L \ Inf_FL \ Infer (map fst (prems_of \\<^sub>F\<^sub>L)) (fst (concl_of \\<^sub>F\<^sub>L)) \ Inf_F\ begin definition to_F :: \('f \ 'l) inference \ 'f inference\ where \to_F \\<^sub>F\<^sub>L = Infer (map fst (prems_of \\<^sub>F\<^sub>L)) (fst (concl_of \\<^sub>F\<^sub>L))\ abbreviation Bot_FL :: \('f \ 'l) set\ where \Bot_FL \ Bot_F \ UNIV\ abbreviation \_F_L :: \('f \ 'l) \ 'g set\ where \\_F_L CL \ \_F (fst CL)\ abbreviation \_Inf_L :: \('f \ 'l) inference \ 'g inference set option\ where \\_Inf_L \\<^sub>F\<^sub>L \ \_Inf (to_F \\<^sub>F\<^sub>L)\ (* lem:labeled-grounding-function *) -sublocale labeled_standard_lifting: standard_lifting +sublocale standard_lifting where Bot_F = Bot_FL and Inf_F = Inf_FL and \_F = \_F_L and \_Inf = \_Inf_L proof show "Bot_FL \ {}" - using Bot_F_not_empty by simp + using no_labels.Bot_F_not_empty by simp next - show "B\Bot_FL \ \_F_L B \ {}" for B - using Bot_map_not_empty by auto + show "B \ Bot_FL \ \_F_L B \ {}" for B + using no_labels.Bot_map_not_empty by auto next - show "B\Bot_FL \ \_F_L B \ Bot_G" for B - using Bot_map by force + show "B \ Bot_FL \ \_F_L B \ Bot_G" for B + using no_labels.Bot_map by force next fix CL show "\_F_L CL \ Bot_G \ {} \ CL \ Bot_FL" - using Bot_cond by (metis SigmaE UNIV_I UNIV_Times_UNIV mem_Sigma_iff prod.sel(1)) + using no_labels.Bot_cond by (metis SigmaE UNIV_I UNIV_Times_UNIV mem_Sigma_iff prod.sel(1)) next fix \ assume i_in: \\ \ Inf_FL\ and ground_not_none: \\_Inf_L \ \ None\ then show "the (\_Inf_L \) \ Red_Inf_G (\_F_L (concl_of \))" - unfolding to_F_def using inf_map Inf_FL_to_Inf_F by fastforce + unfolding to_F_def using no_labels.inf_map Inf_FL_to_Inf_F by fastforce qed abbreviation Labeled_Empty_Order :: \ ('f \ 'l) \ ('f \ 'l) \ bool\ where "Labeled_Empty_Order C1 C2 \ False" sublocale labeled_lifting_w_empty_ord_family: lifting_with_wf_ordering_family Bot_FL Inf_FL Bot_G entails_G Inf_G Red_Inf_G Red_F_G \_F_L \_Inf_L "\g. Labeled_Empty_Order" proof show "po_on Labeled_Empty_Order UNIV" unfolding po_on_def by (simp add: transp_onI wfp_on_imp_irreflp_on) show "wfp_on Labeled_Empty_Order UNIV" unfolding wfp_on_def by simp qed -notation labeled_standard_lifting.entails_\ (infix "\\L" 50) +notation entails_\ (infix "\\L" 50) (* lem:labeled-consequence *) lemma labeled_entailment_lifting: "NL1 \\L NL2 \ fst ` NL1 \\ fst ` NL2" by simp -lemma red_inf_impl: "\ \ labeled_lifting_w_empty_ord_family.Red_Inf_\ NL \ to_F \ \ Red_Inf_\ (fst ` NL)" - unfolding labeled_lifting_w_empty_ord_family.Red_Inf_\_def Red_Inf_\_def using Inf_FL_to_Inf_F - by (auto simp: to_F_def) +lemma red_inf_impl: + "\ \ labeled_lifting_w_empty_ord_family.Red_Inf_\ NL \ to_F \ \ no_labels.Red_Inf_\ (fst ` NL)" + unfolding labeled_lifting_w_empty_ord_family.Red_Inf_\_def no_labels.Red_Inf_\_def + using Inf_FL_to_Inf_F by (auto simp: to_F_def) (* lem:labeled-saturation *) lemma labeled_saturation_lifting: - "labeled_lifting_w_empty_ord_family.saturated NL \ empty_order_lifting.saturated (fst ` NL)" - unfolding labeled_lifting_w_empty_ord_family.saturated_def empty_order_lifting.saturated_def - labeled_standard_lifting.Inf_from_def Inf_from_def + "labeled_lifting_w_empty_ord_family.saturated NL \ no_labels.saturated (fst ` NL)" + unfolding labeled_lifting_w_empty_ord_family.saturated_def no_labels.saturated_def Inf_from_def + no_labels.Inf_from_def proof clarify fix \ assume subs_Red_Inf: "{\ \ Inf_FL. set (prems_of \) \ NL} \ labeled_lifting_w_empty_ord_family.Red_Inf_\ NL" and i_in: "\ \ Inf_F" and i_prems: "set (prems_of \) \ fst ` NL" define Lli where "Lli i = (SOME x. ((prems_of \)!i,x) \ NL)" for i have [simp]:"((prems_of \)!i,Lli i) \ NL" if "i < length (prems_of \)" for i using that i_prems unfolding Lli_def by (metis nth_mem someI_ex DomainE Domain_fst subset_eq) define Ll where "Ll = map Lli [0..)]" have Ll_length: "length Ll = length (prems_of \)" unfolding Ll_def by auto have subs_NL: "set (zip (prems_of \) Ll) \ NL" unfolding Ll_def by (auto simp:in_set_zip) obtain L0 where L0: "Infer (zip (prems_of \) Ll) (concl_of \, L0) \ Inf_FL" using Inf_F_to_Inf_FL[OF i_in Ll_length] .. define \_FL where "\_FL = Infer (zip (prems_of \) Ll) (concl_of \, L0)" then have "set (prems_of \_FL) \ NL" using subs_NL by simp then have "\_FL \ {\ \ Inf_FL. set (prems_of \) \ NL}" unfolding \_FL_def using L0 by blast then have "\_FL \ labeled_lifting_w_empty_ord_family.Red_Inf_\ NL" using subs_Red_Inf by fast moreover have "\ = to_F \_FL" unfolding to_F_def \_FL_def using Ll_length by (cases \) auto - ultimately show "\ \ Red_Inf_\ (fst ` NL)" by (auto intro:red_inf_impl) + ultimately show "\ \ no_labels.Red_Inf_\ (fst ` NL)" by (auto intro: red_inf_impl) qed (* lem:labeled-static-ref-compl *) -lemma stat_ref_comp_to_labeled_sta_ref_comp: "static_refutational_complete_calculus Bot_F Inf_F (\\) Red_Inf_\ Red_F_\ \ - static_refutational_complete_calculus Bot_FL Inf_FL (\\L) +lemma stat_ref_comp_to_labeled_sta_ref_comp: + assumes static: + "static_refutational_complete_calculus Bot_F Inf_F (\\) no_labels.Red_Inf_\ no_labels.Red_F_\" + shows "static_refutational_complete_calculus Bot_FL Inf_FL (\\L) labeled_lifting_w_empty_ord_family.Red_Inf_\ labeled_lifting_w_empty_ord_family.Red_F_\" - unfolding static_refutational_complete_calculus_def -proof (rule conjI impI; clarify) - interpret calculus_with_red_crit Bot_FL Inf_FL labeled_standard_lifting.entails_\ - labeled_lifting_w_empty_ord_family.Red_Inf_\ labeled_lifting_w_empty_ord_family.Red_F_\ - by (simp add: labeled_lifting_w_empty_ord_family.calculus_with_red_crit_axioms) - show "calculus_with_red_crit Bot_FL Inf_FL (\\L) labeled_lifting_w_empty_ord_family.Red_Inf_\ - labeled_lifting_w_empty_ord_family.Red_F_\" - by standard -next +proof + fix Bl :: \'f \ 'l\ and Nl :: \('f \ 'l) set\ assume - calc: "calculus_with_red_crit Bot_F Inf_F (\\) Red_Inf_\ Red_F_\" and - static: "static_refutational_complete_calculus_axioms Bot_F Inf_F (\\) Red_Inf_\" - show "static_refutational_complete_calculus_axioms Bot_FL Inf_FL (\\L) - labeled_lifting_w_empty_ord_family.Red_Inf_\" - unfolding static_refutational_complete_calculus_axioms_def - proof (intro conjI impI allI) - fix Bl :: \'f \ 'l\ and Nl :: \('f \ 'l) set\ - assume - Bl_in: \Bl \ Bot_FL\ and - Nl_sat: \labeled_lifting_w_empty_ord_family.saturated Nl\ and - Nl_entails_Bl: \Nl \\L {Bl}\ - have static_axioms: "B \ Bot_F \ empty_order_lifting.saturated N \ N \\ {B} \ - (\B'\Bot_F. B' \ N)" for B N - using static[unfolded static_refutational_complete_calculus_axioms_def] by fast - define B where "B = fst Bl" - have B_in: "B \ Bot_F" using Bl_in B_def SigmaE by force - define N where "N = fst ` Nl" - have N_sat: "empty_order_lifting.saturated N" - using N_def Nl_sat labeled_saturation_lifting by blast - have N_entails_B: "N \\ {B}" - using Nl_entails_Bl unfolding labeled_entailment_lifting N_def B_def by force - have "\B' \ Bot_F. B' \ N" using B_in N_sat N_entails_B static_axioms[of B N] by blast - then obtain B' where in_Bot: "B' \ Bot_F" and in_N: "B' \ N" by force - then have "B' \ fst ` Bot_FL" by fastforce - obtain Bl' where in_Nl: "Bl' \ Nl" and fst_Bl': "fst Bl' = B'" - using in_N unfolding N_def by blast - have "Bl' \ Bot_FL" using fst_Bl' in_Bot vimage_fst by fastforce - then show \\Bl'\Bot_FL. Bl' \ Nl\ using in_Nl by blast - qed + Bl_in: \Bl \ Bot_FL\ and + Nl_sat: \labeled_lifting_w_empty_ord_family.saturated Nl\ and + Nl_entails_Bl: \Nl \\L {Bl}\ + define B where "B = fst Bl" + have B_in: "B \ Bot_F" using Bl_in B_def SigmaE by force + define N where "N = fst ` Nl" + have N_sat: "no_labels.saturated N" + using N_def Nl_sat labeled_saturation_lifting by blast + have N_entails_B: "N \\ {B}" + using Nl_entails_Bl unfolding labeled_entailment_lifting N_def B_def by force + have "\B' \ Bot_F. B' \ N" using B_in N_sat N_entails_B + using static[unfolded static_refutational_complete_calculus_def + static_refutational_complete_calculus_axioms_def] by blast + then obtain B' where in_Bot: "B' \ Bot_F" and in_N: "B' \ N" by force + then have "B' \ fst ` Bot_FL" by fastforce + obtain Bl' where in_Nl: "Bl' \ Nl" and fst_Bl': "fst Bl' = B'" + using in_N unfolding N_def by blast + have "Bl' \ Bot_FL" using fst_Bl' in_Bot vimage_fst by fastforce + then show \\Bl'\Bot_FL. Bl' \ Nl\ using in_Nl by blast qed end subsection \Labeled Lifting with a Family of Redundancy Criteria\ locale labeled_lifting_with_red_crit_family = no_labels: standard_lifting_with_red_crit_family Inf_F Bot_G Q Inf_G_q entails_q Red_Inf_q Red_F_q Bot_F \_F_q \_Inf_q "\g. Empty_Order" 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" + fixes Inf_FL :: \('f \ 'l) inference set\ assumes Inf_F_to_Inf_FL: \\\<^sub>F \ Inf_F \ length (Ll :: 'l list) = length (prems_of \\<^sub>F) \ \L0. Infer (zip (prems_of \\<^sub>F) Ll) (concl_of \\<^sub>F, L0) \ Inf_FL\ and Inf_FL_to_Inf_F: \\\<^sub>F\<^sub>L \ Inf_FL \ Infer (map fst (prems_of \\<^sub>F\<^sub>L)) (fst (concl_of \\<^sub>F\<^sub>L)) \ Inf_F\ begin definition to_F :: \('f \ 'l) inference \ 'f inference\ where \to_F \\<^sub>F\<^sub>L = Infer (map fst (prems_of \\<^sub>F\<^sub>L)) (fst (concl_of \\<^sub>F\<^sub>L))\ abbreviation Bot_FL :: \('f \ 'l) set\ where \Bot_FL \ Bot_F \ UNIV\ abbreviation \_F_L_q :: \'q \ ('f \ 'l) \ 'g set\ where \\_F_L_q q CL \ \_F_q q (fst CL)\ abbreviation \_Inf_L_q :: \'q \ ('f \ 'l) inference \ 'g inference set option\ where \\_Inf_L_q q \\<^sub>F\<^sub>L \ \_Inf_q q (to_F \\<^sub>F\<^sub>L)\ abbreviation \_set_L_q :: "'q \ ('f \ 'l) set \ 'g set" where "\_set_L_q q N \ \ (\_F_L_q q ` N)" definition Red_Inf_\_L_q :: "'q \ ('f \ 'l) set \ ('f \ 'l) inference set" where "Red_Inf_\_L_q q N = {\ \ Inf_FL. ((\_Inf_L_q q \) \ None \ the (\_Inf_L_q q \) \ Red_Inf_q q (\_set_L_q q N)) \ (\_Inf_L_q q \ = None \ \_F_L_q q (concl_of \) \ \_set_L_q q N \ Red_F_q q (\_set_L_q q N))}" abbreviation Red_Inf_\_L_Q :: "('f \ 'l) set \ ('f \ 'l) inference set" where "Red_Inf_\_L_Q N \ (\q \ Q. Red_Inf_\_L_q q N)" abbreviation Labeled_Empty_Order :: \ ('f \ 'l) \ ('f \ 'l) \ bool\ where "Labeled_Empty_Order C1 C2 \ False" definition Red_F_\_empty_L_q :: "'q \ ('f \ 'l) set \ ('f \ 'l) set" where "Red_F_\_empty_L_q q N = {C. \D \ \_F_L_q q C. D \ Red_F_q q (\_set_L_q q N) \ (\E \ N. Labeled_Empty_Order E C \ D \ \_F_L_q q E)}" abbreviation Red_F_\_empty_L :: "('f \ 'l) set \ ('f \ 'l) set" where "Red_F_\_empty_L N \ (\q \ Q. Red_F_\_empty_L_q q N)" abbreviation entails_\_L_q :: "'q \ ('f \ 'l) set \ ('f \ 'l) set \ bool" where "entails_\_L_q q N1 N2 \ entails_q q (\_set_L_q q N1) (\_set_L_q q N2)" lemma lifting_q: assumes "q \ Q" shows "labeled_lifting_w_wf_ord_family Bot_F Inf_F Bot_G (entails_q q) (Inf_G_q q) (Red_Inf_q q) (Red_F_q q) (\_F_q q) (\_Inf_q q) (\g. Empty_Order) Inf_FL" using assms no_labels.standard_lifting_family Inf_F_to_Inf_FL Inf_FL_to_Inf_F by (simp add: labeled_lifting_w_wf_ord_family_axioms_def labeled_lifting_w_wf_ord_family_def) lemma lifted_q: assumes q_in: "q \ Q" shows "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)" proof - interpret q_lifting: labeled_lifting_w_wf_ord_family Bot_F Inf_F Bot_G "entails_q q" "Inf_G_q q" "Red_Inf_q q" "Red_F_q q" "\_F_q q" "\_Inf_q q" "\g. Empty_Order" Inf_FL using lifting_q[OF q_in] . have "\_Inf_L_q q = q_lifting.\_Inf_L" unfolding to_F_def q_lifting.to_F_def by simp then show ?thesis - using q_lifting.labeled_standard_lifting.standard_lifting_axioms by simp + using q_lifting.standard_lifting_axioms by simp qed lemma ord_fam_lifted_q: 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. Empty_Order)" proof - interpret standard_q_lifting: 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] . have "minimal_element Labeled_Empty_Order UNIV" by (simp add: minimal_element.intro po_on_def transp_onI wfp_on_imp_irreflp_on) then show ?thesis using standard_q_lifting.standard_lifting_axioms by (simp add: lifting_with_wf_ordering_family_axioms.intro lifting_with_wf_ordering_family_def) qed lemma all_lifted_red_crit: assumes q_in: "q \ Q" shows "calculus_with_red_crit Bot_FL Inf_FL (entails_\_L_q q) (Red_Inf_\_L_q q) (Red_F_\_empty_L_q q)" proof - interpret ord_q_lifting: 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. Empty_Order" using ord_fam_lifted_q[OF q_in] . have "Red_Inf_\_L_q q = ord_q_lifting.Red_Inf_\" unfolding Red_Inf_\_L_q_def ord_q_lifting.Red_Inf_\_def by simp moreover have "Red_F_\_empty_L_q q = ord_q_lifting.Red_F_\" unfolding Red_F_\_empty_L_q_def ord_q_lifting.Red_F_\_def by simp ultimately show ?thesis using ord_q_lifting.calculus_with_red_crit_axioms by argo qed lemma all_lifted_cons_rel: assumes q_in: "q \ Q" shows "consequence_relation Bot_FL (entails_\_L_q q)" using all_lifted_red_crit calculus_with_red_crit_def q_in by blast sublocale labeled_cons_rel_family: consequence_relation_family Bot_FL Q entails_\_L_q using all_lifted_cons_rel by (simp add: consequence_relation_family.intro no_labels.Q_nonempty) sublocale calculus_with_red_crit_family Bot_FL Inf_FL Q entails_\_L_q Red_Inf_\_L_q Red_F_\_empty_L_q using calculus_with_red_crit_family.intro[OF labeled_cons_rel_family.consequence_relation_family_axioms] by (simp add: all_lifted_red_crit calculus_with_red_crit_family_axioms_def no_labels.Q_nonempty) notation no_labels.entails_\_Q (infix "\\\" 50) abbreviation entails_\_L_Q :: "('f \ 'l) set \ ('f \ 'l) set \ bool" (infix "\\\L" 50) where "(\\\L) \ labeled_cons_rel_family.entails_Q" lemmas entails_\_L_Q_def = labeled_cons_rel_family.entails_Q_def (* lem:labeled-consequence-intersection *) lemma labeled_entailment_lifting: "NL1 \\\L NL2 \ fst ` NL1 \\\ fst ` NL2" unfolding no_labels.entails_\_Q_def entails_\_L_Q_def by force lemma red_inf_impl: "\ \ Red_Inf_Q NL \ to_F \ \ no_labels.Red_Inf_\_Q (fst ` NL)" unfolding no_labels.Red_Inf_\_Q_def Red_Inf_Q_def proof clarify fix X Xa q assume q_in: "q \ Q" and i_in_inter: "\ \ (\q \ Q. Red_Inf_\_L_q q NL)" have i_in_q: "\ \ Red_Inf_\_L_q q NL" using q_in i_in_inter image_eqI by blast then have i_in: "\ \ Inf_FL" unfolding Red_Inf_\_L_q_def by blast have to_F_in: "to_F \ \ Inf_F" unfolding to_F_def using Inf_FL_to_Inf_F[OF i_in] . have rephrase1: "(\CL\NL. \_F_q q (fst CL)) = (\ (\_F_q q ` fst ` NL))" by blast have rephrase2: "fst (concl_of \) = concl_of (to_F \)" unfolding concl_of_def to_F_def by simp have subs_red: "(\_Inf_L_q q \ \ None \ the (\_Inf_L_q q \) \ Red_Inf_q q (\_set_L_q q NL)) \ (\_Inf_L_q q \ = None \ \_F_L_q q (concl_of \) \ \_set_L_q q NL \ Red_F_q q (\_set_L_q q NL))" using i_in_q unfolding Red_Inf_\_L_q_def by blast then have to_F_subs_red: "(\_Inf_q q (to_F \) \ None \ the (\_Inf_q q (to_F \)) \ Red_Inf_q q (no_labels.\_set_q q (fst ` NL))) \ (\_Inf_q q (to_F \) = None \ \_F_q q (concl_of (to_F \)) \ no_labels.\_set_q q (fst ` NL) \ Red_F_q q (no_labels.\_set_q q (fst ` NL)))" using rephrase1 rephrase2 by metis then show "to_F \ \ no_labels.Red_Inf_\_q q (fst ` NL)" using to_F_in unfolding no_labels.Red_Inf_\_q_def by simp qed (* lem:labeled-saturation-intersection *) lemma labeled_family_saturation_lifting: "saturated NL \ no_labels.saturated (fst ` NL)" unfolding saturated_def no_labels.saturated_def Inf_from_def no_labels.Inf_from_def proof clarify fix \F assume labeled_sat: "{\ \ Inf_FL. set (prems_of \) \ NL} \ Red_Inf_Q NL" and iF_in: "\F \ Inf_F" and iF_prems: "set (prems_of \F) \ fst ` NL" define Lli where "Lli i = (SOME x. ((prems_of \F)!i,x) \ NL)" for i have [simp]:"((prems_of \F)!i,Lli i) \ NL" if "i < length (prems_of \F)" for i using that iF_prems nth_mem someI_ex unfolding Lli_def by (metis DomainE Domain_fst subset_eq) define Ll where "Ll = map Lli [0..F)]" have Ll_length: "length Ll = length (prems_of \F)" unfolding Ll_def by auto have subs_NL: "set (zip (prems_of \F) Ll) \ NL" unfolding Ll_def by (auto simp:in_set_zip) obtain L0 where L0: "Infer (zip (prems_of \F) Ll) (concl_of \F, L0) \ Inf_FL" using Inf_F_to_Inf_FL[OF iF_in Ll_length] .. define \FL where "\FL = Infer (zip (prems_of \F) Ll) (concl_of \F, L0)" then have "set (prems_of \FL) \ NL" using subs_NL by simp then have "\FL \ {\ \ Inf_FL. set (prems_of \) \ NL}" unfolding \FL_def using L0 by blast then have "\FL \ Red_Inf_Q NL" using labeled_sat by fast moreover have "\F = to_F \FL" unfolding to_F_def \FL_def using Ll_length by (cases \F) auto ultimately show "\F \ no_labels.Red_Inf_\_Q (fst ` NL)" by (auto intro: red_inf_impl) qed (* thm:labeled-static-ref-compl-intersection *) theorem labeled_static_ref: assumes calc: "static_refutational_complete_calculus Bot_F Inf_F (\\\) no_labels.Red_Inf_\_Q no_labels.Red_F_\_empty" shows "static_refutational_complete_calculus Bot_FL Inf_FL (\\\L) Red_Inf_Q Red_F_Q" proof fix Bl :: \'f \ 'l\ and Nl :: \('f \ 'l) set\ assume Bl_in: \Bl \ Bot_FL\ and Nl_sat: \saturated Nl\ and Nl_entails_Bl: \Nl \\\L {Bl}\ define B where "B = fst Bl" have B_in: "B \ Bot_F" using Bl_in B_def SigmaE by force define N where "N = fst ` Nl" have N_sat: "no_labels.saturated N" using N_def Nl_sat labeled_family_saturation_lifting by blast have N_entails_B: "N \\\ {B}" using Nl_entails_Bl unfolding labeled_entailment_lifting N_def B_def by force have "\B' \ Bot_F. B' \ N" using B_in N_sat N_entails_B calc[unfolded static_refutational_complete_calculus_def static_refutational_complete_calculus_axioms_def] by blast then obtain B' where in_Bot: "B' \ Bot_F" and in_N: "B' \ N" by force then have "B' \ fst ` Bot_FL" by fastforce obtain Bl' where in_Nl: "Bl' \ Nl" and fst_Bl': "fst Bl' = B'" using in_N unfolding N_def by blast have "Bl' \ Bot_FL" using fst_Bl' in_Bot vimage_fst by fastforce then show \\Bl'\Bot_FL. Bl' \ Nl\ using in_Nl by blast qed end end