diff --git a/thys/Auto2_HOL/Auto2_Setup.thy b/thys/Auto2_HOL/Auto2_Setup.thy new file mode 100644 --- /dev/null +++ b/thys/Auto2_HOL/Auto2_Setup.thy @@ -0,0 +1,35 @@ +(* + File: Auto2_Base.thy + Author: Bohua Zhan + + Base file for auto2 setup. +*) + +theory Auto2_Setup + imports Pure +begin + +ML_file \util.ML\ +ML_file \util_base.ML\ +ML_file \util_logic.ML\ +ML_file \box_id.ML\ +ML_file \consts.ML\ +ML_file \property.ML\ +ML_file \wellform.ML\ +ML_file \wfterm.ML\ +ML_file \rewrite.ML\ +ML_file \propertydata.ML\ +ML_file \matcher.ML\ +ML_file \items.ML\ +ML_file \wfdata.ML\ +ML_file \auto2_data.ML\ +ML_file \status.ML\ +ML_file \normalize.ML\ +ML_file \proofsteps.ML\ +ML_file \auto2_state.ML\ +ML_file \logic_steps.ML\ +ML_file \auto2.ML\ +ML_file \auto2_outer.ML\ +ML_file \auto2_setup.ML\ + +end diff --git a/thys/Auto2_HOL/HOL/Auto2_HOL.thy b/thys/Auto2_HOL/HOL/Auto2_HOL.thy deleted file mode 100644 --- a/thys/Auto2_HOL/HOL/Auto2_HOL.thy +++ /dev/null @@ -1,63 +0,0 @@ -(* - File: Auto2_HOL.thy - Author: Bohua Zhan - - Main file for auto2 setup in HOL. -*) - -theory Auto2_HOL - imports HOL_Base - keywords "@proof" :: prf_block % "proof" - and "@have" "@case" "@obtain" "@let" "@contradiction" "@strong_induct" :: prf_decl % "proof" - and "@unfold" :: prf_decl % "proof" - and "@induct" "@fun_induct" "@case_induct" "@prop_induct" "@cases" :: prf_decl % "proof" - and "@apply_induct_hyp" :: prf_decl % "proof" - and "@subgoal" "@endgoal" "@end" :: prf_decl % "proof" - and "@qed" :: qed_block % "proof" - and "@with" "where" "arbitrary" "@rule" :: quasi_command -begin - -ML_file \../util.ML\ -ML_file \../util_base.ML\ -ML_file \auto2_hol.ML\ -ML_file \../util_logic.ML\ -ML_file \../box_id.ML\ -ML_file \../consts.ML\ -ML_file \../property.ML\ -ML_file \../wellform.ML\ -ML_file \../wfterm.ML\ -ML_file \../rewrite.ML\ -ML_file \../propertydata.ML\ -ML_file \../matcher.ML\ -ML_file \../items.ML\ -ML_file \../wfdata.ML\ -ML_file \../auto2_data.ML\ -ML_file \../status.ML\ -ML_file \../normalize.ML\ -ML_file \../proofsteps.ML\ -ML_file \../auto2_state.ML\ -ML_file \../logic_steps.ML\ -ML_file \../auto2.ML\ -ML_file \../auto2_outer.ML\ - -ML_file \acdata.ML\ -ML_file \ac_steps.ML\ -ML_file \unfolding.ML\ -ML_file \induct_outer.ML\ -ML_file \extra_hol.ML\ - -method_setup auto2 = \Scan.succeed (SIMPLE_METHOD o Auto2.auto2_tac)\ "auto2 prover" - -attribute_setup forward = \setup_attrib add_forward_prfstep\ -attribute_setup backward = \setup_attrib add_backward_prfstep\ -attribute_setup backward1 = \setup_attrib add_backward1_prfstep\ -attribute_setup backward2 = \setup_attrib add_backward2_prfstep\ -attribute_setup resolve = \setup_attrib add_resolve_prfstep\ -attribute_setup rewrite = \setup_attrib add_rewrite_rule\ -attribute_setup rewrite_back = \setup_attrib add_rewrite_rule_back\ -attribute_setup rewrite_bidir = \setup_attrib add_rewrite_rule_bidir\ -attribute_setup forward_arg1 = \setup_attrib add_forward_arg1_prfstep\ -attribute_setup forward_arg = \setup_attrib add_forward_arg_prfstep\ -attribute_setup rewrite_arg = \setup_attrib add_rewrite_arg_rule\ - -end diff --git a/thys/Auto2_HOL/HOL/Auto2_HOL_Extra_Setup.thy b/thys/Auto2_HOL/HOL/Auto2_HOL_Extra_Setup.thy new file mode 100644 --- /dev/null +++ b/thys/Auto2_HOL/HOL/Auto2_HOL_Extra_Setup.thy @@ -0,0 +1,21 @@ +(* + File: Auto2_HOL_Extra_Setup.thy + Author: Kevin Kappelmann + + Extra setup files for auto2 in HOL. +*) + +theory Auto2_HOL_Extra_Setup + imports + HOL.HOL + Auto2_Setup +begin + +ML_file \acdata.ML\ +ML_file \ac_steps.ML\ +ML_file \unfolding.ML\ +ML_file \induct_outer.ML\ +ML_file \extra_hol.ML\ +ML_file \auto2_hol_extra_setup.ML\ + +end diff --git a/thys/Auto2_HOL/HOL/Auto2_HOL_Setup.thy b/thys/Auto2_HOL/HOL/Auto2_HOL_Setup.thy new file mode 100644 --- /dev/null +++ b/thys/Auto2_HOL/HOL/Auto2_HOL_Setup.thy @@ -0,0 +1,104 @@ +(* + File: Auto2_HOL_Setup.thy + Author: Bohua Zhan + + Setup of auto2 in HOL. +*) + +theory Auto2_HOL_Setup + imports + Auto2_HOL_Extra_Setup + HOL_Base + keywords "@proof" :: prf_block % "proof" + and "@have" "@case" "@obtain" "@let" "@contradiction" "@strong_induct" :: prf_decl % "proof" + and "@unfold" :: prf_decl % "proof" + and "@induct" "@fun_induct" "@case_induct" "@prop_induct" "@cases" :: prf_decl % "proof" + and "@apply_induct_hyp" :: prf_decl % "proof" + and "@subgoal" "@endgoal" "@end" :: prf_decl % "proof" + and "@qed" :: qed_block % "proof" + and "@with" "where" "arbitrary" "@rule" :: quasi_command +begin + +ML_file \auto2_hol_util_base.ML\ +ML\ + structure Auto2_Keywords : AUTO2_KEYWORDS = + struct + val case' = @{command_keyword "@case"} + val contradiction = @{command_keyword "@contradiction"} + val end' = @{command_keyword "@end"} + val endgoal = @{command_keyword "@endgoal"} + val have = @{command_keyword "@have"} + val let' = @{command_keyword "@let"} + val obtain = @{command_keyword "@obtain"} + val proof = @{command_keyword "@proof"} + val qed = @{command_keyword "@qed"} + val subgoal = @{command_keyword "@subgoal"} + val rule = @{keyword "@rule"} + val with' = @{keyword "@with"} + end; + structure Auto2_Setup = Auto2_Setup( + structure Auto2_Keywords = Auto2_Keywords; + structure UtilBase = Auto2_UtilBase; + ); +\ + +ML\ + structure Unfolding_Keyword : UNFOLDING_KEYWORD = + struct val unfold = @{command_keyword "@unfold"} end; + structure Induct_ProofSteps_Keywords : INDUCT_PROOFSTEPS_KEYWORDS = + struct + val apply_induct_hyp = @{command_keyword "@apply_induct_hyp"} + val case_induct = @{command_keyword "@case_induct"} + val cases = @{command_keyword "@cases"} + val fun_induct = @{command_keyword "@fun_induct"} + val induct = @{command_keyword "@induct"} + val prop_induct = @{command_keyword "@prop_induct"} + val strong_induct = @{command_keyword "@strong_induct"} + val arbitrary = @{keyword "arbitrary"} + val with' = @{keyword "@with"} + end; + structure Auto2_HOL_Extra_Setup = Auto2_HOL_Extra_Setup( + structure Auto2_Setup = Auto2_Setup; + structure Unfolding_Keyword = Unfolding_Keyword; + structure Induct_ProofSteps_Keywords = Induct_ProofSteps_Keywords; + ); +\ +method_setup auto2 = \Scan.succeed (SIMPLE_METHOD o Auto2_Setup.Auto2.auto2_tac)\ "auto2 prover" + +ML\ + open Auto2_Basic_UtilBase + open Auto2_Setup.Basic_UtilLogic + + val WithTerm = Auto2_Setup.ProofStep.WithTerm + val WithGoal = Auto2_Setup.ProofStep.WithGoal + val WithProp = Auto2_Setup.ProofStep.WithProp + val neq_filter = Auto2_Setup.ProofStep.neq_filter + val order_filter = Auto2_Setup.ProofStep.order_filter + val size1_filter = Auto2_Setup.ProofStep.size1_filter + val not_type_filter = Auto2_Setup.ProofStep.not_type_filter + + open Auto2_Setup.ProofStepData + open Auto2_HOL_Extra_Setup.Extra_HOL + + val add_strong_induct_rule = Auto2_HOL_Extra_Setup.Induct_ProofSteps.add_strong_induct_rule + val add_case_induct_rule = Auto2_HOL_Extra_Setup.Induct_ProofSteps.add_case_induct_rule + val add_prop_induct_rule = Auto2_HOL_Extra_Setup.Induct_ProofSteps.add_prop_induct_rule + val add_var_induct_rule = Auto2_HOL_Extra_Setup.Induct_ProofSteps.add_var_induct_rule + val add_fun_induct_rule = Auto2_HOL_Extra_Setup.Induct_ProofSteps.add_fun_induct_rule + val add_cases_rule = Auto2_HOL_Extra_Setup.Induct_ProofSteps.add_cases_rule +\ + +attribute_setup forward = \setup_attrib add_forward_prfstep\ +attribute_setup backward = \setup_attrib add_backward_prfstep\ +attribute_setup backward1 = \setup_attrib add_backward1_prfstep\ +attribute_setup backward2 = \setup_attrib add_backward2_prfstep\ +attribute_setup resolve = \setup_attrib add_resolve_prfstep\ +attribute_setup rewrite = \setup_attrib add_rewrite_rule\ +attribute_setup rewrite_back = \setup_attrib add_rewrite_rule_back\ +attribute_setup rewrite_bidir = \setup_attrib add_rewrite_rule_bidir\ +attribute_setup forward_arg1 = \setup_attrib add_forward_arg1_prfstep\ +attribute_setup forward_arg = \setup_attrib add_forward_arg_prfstep\ +attribute_setup rewrite_arg = \setup_attrib add_rewrite_arg_rule\ + + +end diff --git a/thys/Auto2_HOL/HOL/Lists_Thms.thy b/thys/Auto2_HOL/HOL/Lists_Thms.thy --- a/thys/Auto2_HOL/HOL/Lists_Thms.thy +++ b/thys/Auto2_HOL/HOL/Lists_Thms.thy @@ -1,312 +1,312 @@ (* File: Lists_Thms.thy Author: Bohua Zhan Setup for proof steps related to lists. *) section \Setup for lists\ theory Lists_Thms imports Set_Thms begin subsection \Definition of lists\ setup \add_resolve_prfstep @{thm list.distinct(2)}\ setup \add_forward_prfstep (equiv_forward_th @{thm list.simps(1)})\ setup \fold add_rewrite_rule @{thms List.list.sel}\ setup \add_rewrite_rule @{thm list.collapse}\ setup \add_var_induct_rule @{thm list.induct}\ subsection \Length\ setup \add_rewrite_rule @{thm List.list.size(3)}\ lemma length_one [rewrite]: "length [x] = 1" by simp lemma length_Cons [rewrite]: "length (a # b) = length b + 1" by simp lemma length_snoc [rewrite]: "length (xs @ [x]) = length xs + 1" by auto lemma length_zero_is_nil [forward]: "length xs = 0 \ xs = []" by simp lemma length_gt_zero [forward]: "length xs > 0 \ xs \ []" by simp subsection \Append\ setup \add_rewrite_rule @{thm List.length_append}\ setup \add_rewrite_rule_cond @{thm List.append.simps(2)} [with_cond "?xs \ []"]\ setup \add_rewrite_rule @{thm List.hd_append2}\ lemma append_is_empty [forward]: "xs @ ys = [] \ xs = [] \ ys = []" by simp lemma cons_to_append [rewrite_back]: "a # b = [a] @ b" by simp ML_file \list_ac.ML\ ML_file \list_ac_test.ML\ subsection \Showing two lists are equal\ setup \add_backward2_prfstep_cond @{thm nth_equalityI} [with_filt (order_filter "xs" "ys")]\ subsection \Set of elements of a list\ setup \add_rewrite_rule @{thm List.set_simps(1)}\ lemma set_one [rewrite]: "set [u] = {u}" by simp lemma set_two [rewrite]: "set [u, v] = {u, v}" by simp lemma set_simps2: "set (x # xs) = {x} \ set xs" by simp setup \add_rewrite_rule_cond @{thm set_simps2} [with_cond "?xs \ []", with_cond "?xs \ [?y]"]\ setup \add_rewrite_rule @{thm List.set_append}\ setup \add_rewrite_rule @{thm List.set_rev}\ setup \add_resolve_prfstep @{thm List.finite_set}\ setup \add_backward_prfstep (equiv_forward_th @{thm in_set_conv_nth})\ subsection \hd\ setup \register_wellform_data ("hd xs", ["xs \ []"])\ setup \add_forward_prfstep_cond @{thm List.hd_in_set} [with_term "hd ?xs"]\ subsection \tl\ setup \add_rewrite_rule @{thm length_tl}\ lemma nth_tl' [rewrite]: "i < length (tl xs) \ tl xs ! i = xs ! (i + 1)" by (simp add: nth_tl) lemma set_tl_subset [forward_arg1]: "set (tl xs) \ set xs" by (metis list.set_sel(2) subsetI tl_Nil) subsection \nth\ setup \register_wellform_data ("xs ! i", ["i < length xs"])\ setup \add_prfstep_check_req ("xs ! i", "i < length xs")\ setup \add_rewrite_rule_back @{thm hd_conv_nth}\ setup \add_rewrite_rule @{thm List.nth_Cons'}\ setup \add_rewrite_rule @{thm List.nth_append}\ setup \add_forward_prfstep_cond @{thm nth_mem} [with_term "?xs ! ?n"]\ subsection \sorted\ (*TEMPORARY VERSION BY BOHUA pending a treatment of sorted_wrt*) definition sorted :: "'a::linorder list \ bool" where "sorted = List.sorted" declare sorted_def[simp] lemma sorted_Nil [resolve]: "sorted []" by simp lemma sorted_single [resolve]: "sorted [x]" by simp lemma sorted_simps2: "sorted (x # ys) = (Ball (set ys) ((\) x) \ sorted ys)" by simp setup \add_backward_prfstep (equiv_backward_th @{thm sorted_simps2})\ lemma sorted_ConsD1 [forward]: "sorted (x # xs) \ sorted xs" using sorted_simps(2) by simp lemma sorted_ConsD2 [forward, backward2]: "sorted (x # xs) \ y \ set xs \ x \ y" - using sorted_simps(2) by simp + using sorted_simps(2) by auto lemma sorted_appendI [backward]: "sorted xs \ sorted ys \ (\x\set xs. \y\set ys. x \ y) \ sorted (xs @ ys)" by (simp add: sorted_append) lemma sorted_appendE [forward]: "sorted (xs @ ys) \ sorted xs \ sorted ys" by (simp add: sorted_append) lemma sorted_appendE2 [forward]: "sorted (xs @ ys) \ x \ set xs \ \y\set ys. x \ y" - using sorted_append by (simp add: sorted_append) + using sorted_append by (auto simp add: sorted_append) lemma sorted_nth_mono' [backward]: "sorted xs \ j < length xs \ i \ j \ xs ! i \ xs ! j" using sorted_nth_mono by auto lemma sorted_nth_mono_less [forward]: "sorted xs \ i < length xs \ xs ! i < xs ! j \ i < j" by (meson leD not_le_imp_less sorted_nth_mono') subsection \sort\ lemma sorted_sort: "sorted (sort xs)" by simp setup \add_forward_prfstep_cond @{thm sorted_sort} [with_term "sort ?xs"]\ setup \add_rewrite_rule @{thm length_sort}\ setup \add_rewrite_rule @{thm mset_sort}\ setup \add_rewrite_rule @{thm set_sort}\ lemma properties_for_sort: "mset ys = mset xs \ sorted ys \ sort xs = ys" using properties_for_sort by simp setup \add_backward_prfstep @{thm properties_for_sort}\ lemma sort_Nil [rewrite]: "sort [] = []" by auto lemma sort_singleton [rewrite]: "sort [a] = [a]" by auto subsection \distinct\ - + lemma distinct_Nil [resolve]: "distinct []" by simp setup \add_resolve_prfstep @{thm List.distinct_singleton}\ setup \add_rewrite_rule_cond @{thm distinct.simps(2)} [with_cond "?xs \ []"]\ setup \add_rewrite_rule @{thm distinct_append}\ setup \add_rewrite_rule @{thm distinct_rev}\ setup \add_rewrite_rule @{thm distinct_sort}\ setup \add_resolve_prfstep (equiv_backward_th @{thm distinct_conv_nth})\ lemma distinct_nthE [forward]: "distinct xs \ i < length xs \ j < length xs \ xs ! i = xs ! j \ i = j" using nth_eq_iff_index_eq by blast subsection \map function\ setup \fold add_rewrite_rule @{thms List.list.map}\ setup \add_rewrite_rule @{thm List.length_map}\ setup \add_rewrite_rule @{thm List.nth_map}\ setup \add_rewrite_rule @{thm List.map_append}\ subsection \Replicate\ setup \add_rewrite_arg_rule @{thm length_replicate}\ setup \add_rewrite_rule @{thm List.nth_replicate}\ subsection \last\ setup \register_wellform_data ("last xs", ["xs \ []"])\ lemma last_eval1 [rewrite]: "last [x] = x" by simp lemma last_eval2 [rewrite]: "last [u, v] = v" by simp setup \add_rewrite_rule @{thm List.last_ConsR}\ setup \add_rewrite_rule @{thm List.last_appendR}\ setup \add_rewrite_rule @{thm List.last_snoc}\ setup \add_rewrite_rule_back @{thm last_conv_nth}\ setup \add_forward_prfstep_cond @{thm List.last_in_set} [with_term "last ?as"]\ subsection \butlast\ setup \add_rewrite_arg_rule @{thm List.length_butlast}\ setup \add_rewrite_rule @{thm List.nth_butlast}\ setup \add_rewrite_rule_back @{thm List.butlast_conv_take}\ setup \add_rewrite_rule @{thm List.butlast_snoc}\ lemma butlast_eval1 [rewrite]: "butlast [x] = []" by simp lemma butlast_eval2 [rewrite]: "butlast [x, y] = [x]" by simp lemma butlast_cons [rewrite]: "as \ [] \ butlast (a # as) = a # butlast as" by simp lemma butlast_append' [rewrite]: "bs \ [] \ butlast (as @ bs) = as @ butlast bs" by (simp add: butlast_append) setup \add_rewrite_rule @{thm List.append_butlast_last_id}\ lemma set_butlast_is_subset: "set (butlast xs) \ set xs" by (simp add: in_set_butlastD subsetI) setup \add_forward_arg1_prfstep @{thm set_butlast_is_subset}\ subsection \List update\ setup \register_wellform_data ("xs[i := x]", ["i < length xs"])\ setup \add_rewrite_arg_rule @{thm List.length_list_update}\ setup \add_rewrite_rule @{thm List.nth_list_update_eq}\ setup \add_rewrite_rule @{thm List.nth_list_update_neq}\ setup \add_rewrite_rule @{thm List.nth_list_update}\ subsection \take\ setup \register_wellform_data ("take n xs", ["n \ length xs"])\ setup \add_prfstep_check_req ("take n xs", "n \ length xs")\ lemma length_take [rewrite_arg]: "n \ length xs \ length (take n xs) = n" by simp lemma nth_take [rewrite]: "i < length (take n xs) \ take n xs ! i = xs ! i" by simp setup \add_rewrite_rule @{thm List.take_0}\ setup \add_rewrite_rule @{thm List.take_Suc_conv_app_nth}\ lemma take_length [rewrite]: "take (length xs) xs = xs" by simp setup \add_forward_arg1_prfstep @{thm List.set_take_subset}\ lemma take_Suc [rewrite]: "Suc n \ length xs \ take (Suc n) xs = take n xs @ [nth xs n]" using Suc_le_lessD take_Suc_conv_app_nth by blast setup \add_rewrite_rule @{thm List.take_update_cancel}\ setup \add_rewrite_rule @{thm List.append_take_drop_id}\ setup \add_rewrite_rule @{thm List.take_all}\ subsection \drop\ setup \add_rewrite_arg_rule @{thm List.length_drop}\ lemma nth_drop [rewrite]: "i < length (drop n xs) \ drop n xs ! i = xs ! (n + i)" by simp setup \add_rewrite_rule @{thm List.drop_0}\ setup \add_rewrite_rule @{thm List.drop_all}\ setup \add_rewrite_rule_back @{thm List.take_drop}\ setup \add_rewrite_rule @{thm List.drop_drop}\ subsection \rev\ setup \add_rewrite_arg_rule @{thm List.length_rev}\ setup \fold add_rewrite_rule @{thms List.rev.simps}\ setup \add_rewrite_rule @{thm List.rev_append}\ setup \add_rewrite_rule @{thm List.rev_rev_ident}\ subsection \filter\ setup \fold add_rewrite_rule @{thms filter.simps}\ setup \add_rewrite_rule @{thm filter_append}\ setup \add_rewrite_rule_bidir @{thm rev_filter}\ subsection \concat\ setup \fold add_rewrite_rule @{thms concat.simps}\ subsection \mset\ setup \add_rewrite_rule @{thm mset.simps(1)}\ lemma mset_simps_2 [rewrite]: "mset (a # x) = mset x + {#a#}" by simp setup \add_rewrite_rule @{thm mset_append}\ setup \add_rewrite_rule @{thm mset_eq_setD}\ setup \add_rewrite_rule_cond @{thm in_multiset_in_set} [with_term "set ?xs"]\ setup \add_rewrite_rule_back_cond @{thm in_multiset_in_set} [with_term "mset ?xs"]\ setup \add_backward_prfstep @{thm Multiset.nth_mem_mset}\ lemma in_mset_conv_nth [resolve]: "x \# mset xs \ \i [] \ hd xs \# mset xs" by simp setup \add_forward_prfstep_cond @{thm hd_in_mset} [with_term "hd ?xs", with_term "mset ?xs"]\ lemma last_in_mset: "xs \ [] \ last xs \# mset xs" by simp setup \add_forward_prfstep_cond @{thm last_in_mset} [with_term "last ?xs", with_term "mset ?xs"]\ subsection \Relationship between mset and set of lists\ lemma mset_butlast [rewrite]: "xs \ [] \ mset (butlast xs) = mset xs - {# last xs #}" by (metis add_diff_cancel_right' append_butlast_last_id mset.simps(1) mset.simps(2) union_code) lemma insert_mset_to_set [rewrite]: "mset xs' = mset xs + {# x #} \ set xs' = set xs \ {x}" by (metis set_mset_mset set_mset_single set_mset_union) lemma delete_mset_to_set [rewrite]: "distinct xs \ mset xs' = mset xs - {# x #} \ set xs' = set xs - {x}" by (metis mset_eq_setD mset_remove1 set_remove1_eq) lemma update_mset_to_set [rewrite]: "distinct xs \ mset xs' = {# y #} + (mset xs - {# x #}) \ set xs' = (set xs - {x}) \ {y}" by (metis insert_mset_to_set mset_remove1 set_remove1_eq union_commute) lemma mset_update' [rewrite]: "i < length ls \ mset (ls[i := v]) = {#v#} + (mset ls - {# ls ! i #})" using mset_update by fastforce subsection \swap\ setup \add_rewrite_rule @{thm mset_swap}\ setup \add_rewrite_rule @{thm set_swap}\ subsection \upto lists\ lemma upt_zero_length [rewrite]: "length [0.. [0..Lambda lists\ definition list :: "(nat \ 'a) \ nat \ 'a list" where "list s n = map s [0 ..< n]" lemma list_length [rewrite_arg]: "length (list s n) = n" by (simp add: list_def) lemma list_nth [rewrite]: "i < length (list s n) \ (list s n) ! i = s i" by (simp add: list_def) subsection \Splitting of lists\ setup \add_resolve_prfstep @{thm split_list}\ setup \add_backward_prfstep @{thm not_distinct_decomp}\ subsection \Finiteness\ setup \add_resolve_prfstep @{thm finite_lists_length_le}\ subsection \Cardinality\ setup \add_rewrite_rule @{thm distinct_card}\ end diff --git a/thys/Auto2_HOL/HOL/Logic_Thms.thy b/thys/Auto2_HOL/HOL/Logic_Thms.thy --- a/thys/Auto2_HOL/HOL/Logic_Thms.thy +++ b/thys/Auto2_HOL/HOL/Logic_Thms.thy @@ -1,112 +1,113 @@ (* File: Logic_Thms.thy Author: Bohua Zhan Setup for proof steps related to logic. *) theory Logic_Thms -imports Auto2_HOL +imports Auto2_HOL_Setup begin (* Trivial contradictions. *) setup \add_resolve_prfstep @{thm HOL.refl}\ setup \add_forward_prfstep @{thm contra_triv}\ setup \add_resolve_prfstep @{thm TrueI}\ theorem FalseD [resolve]: "\False" by simp lemma exists_triv_eq [resolve]: "\x. x = x" by auto (* Not. *) -setup \add_forward_prfstep_cond @{thm HOL.not_sym} [with_filt (not_type_filter "s" boolT)]\ +setup \add_forward_prfstep_cond @{thm HOL.not_sym} + [with_filt (Auto2_Setup.ProofStep.not_type_filter "s" boolT)]\ (* Iff. *) setup \add_gen_prfstep ("iff_intro1", [WithGoal @{term_pat "(?A::bool) = ?B"}, CreateCase @{term_pat "?A::bool"}, WithScore 25])\ theorem iff_goal: "A \ B \ A \ \B" "A \ B \ B \ \A" "A \ B \ \A \ B" "A \ B \ \B \ A" "(\A) \ B \ A \ B" "A \ (\B) \ B \ A" by auto setup \fold (fn th => add_forward_prfstep_cond th [with_score 1]) @{thms iff_goal}\ (* Quantifiers: normalization *) theorem exists_split: "(\x y. P x \ Q y) = ((\x. P x) \ (\y. Q y))" by simp setup \add_backward_prfstep (equiv_backward_th @{thm exists_split})\ (* Case analysis. *) setup \add_gen_prfstep ("case_intro", [WithTerm @{term_pat "if ?cond then (?yes::?'a) else ?no"}, CreateCase @{term_pat "?cond::bool"}])\ setup \add_gen_prfstep ("case_intro_fact", [WithFact @{term_pat "if ?cond then (?yes::bool) else ?no"}, CreateCase @{term_pat "?cond::bool"}])\ setup \add_gen_prfstep ("case_intro_goal", [WithGoal @{term_pat "if ?cond then (?yes::bool) else ?no"}, CreateCase @{term_pat "?cond::bool"}])\ lemma if_eval': "P \ (if \P then x else y) = y" by auto lemma ifb_eval: "P \ (if P then (x::bool) else y) = x" "\P \ (if P then (x::bool) else y) = y" "P \ (if \P then (x::bool) else y) = y" by auto setup \fold (fn th => add_rewrite_rule_cond th [with_score 1]) ([@{thm HOL.if_P}, @{thm HOL.if_not_P}, @{thm if_eval'}] @ @{thms ifb_eval})\ (* THE and \! *) setup \add_forward_prfstep_cond @{thm theI'} [with_term "THE x. ?P x"]\ setup \add_gen_prfstep ("ex1_case", [WithGoal @{term_pat "\!x. ?P x"}, CreateConcl @{term_pat "\x. ?P x"}])\ theorem ex_ex1I' [backward1]: "\y. P y \ x = y \ P x \ \!x. P x" by auto theorem the1_equality': "P a \ \!x. P x \ (THE x. P x) = a" by (simp add: the1_equality) setup \add_forward_prfstep_cond @{thm the1_equality'} [with_term "THE x. ?P x"]\ (* Hilbert choice. *) setup \add_gen_prfstep ("SOME_case_intro", [WithTerm @{term_pat "SOME k. ?P k"}, CreateConcl @{term_pat "\k. ?P k"}])\ setup \add_forward_prfstep_cond @{thm someI} [with_term "SOME x. ?P x"]\ setup \add_forward_prfstep_cond @{thm someI_ex} [with_term "SOME x. ?P x"]\ (* Axiom of choice *) setup \add_prfstep_custom ("ex_choice", [WithGoal @{term_pat "EX f. !x. ?Q f x"}], (fn ((id, _), ths) => fn _ => fn _ => let - val choice = @{thm choice} |> apply_to_thm (Conv.rewr_conv UtilBase.backward_conv_th) + val choice = @{thm choice} |> apply_to_thm (Conv.rewr_conv Auto2_UtilBase.backward_conv_th) in - [Update.thm_update (id, (ths MRS choice))] + [Auto2_Setup.Update.thm_update (id, (ths MRS choice))] end handle THM _ => [])) \ (* Least operator. *) theorem Least_equality' [backward1]: "P (x::('a::order)) \ \y. P y \ x \ y \ Least P = x" by (simp add: Least_equality) (* Pairs. *) lemma pair_inj: "(a,b) = c \ a = fst c \ b = snd c" by auto -setup \Normalizer.add_inj_struct_data @{thm pair_inj}\ +setup \Auto2_Setup.Normalizer.add_inj_struct_data @{thm pair_inj}\ setup \add_rewrite_rule @{thm fst_conv}\ setup \add_rewrite_rule @{thm snd_conv}\ setup \add_forward_prfstep (equiv_forward_th @{thm prod.simps(1)})\ setup \add_rewrite_rule_cond @{thm surjective_pairing} [with_cond "?t \ (?a, ?b)"]\ -setup \Normalizer.add_rewr_normalizer ("rewr_case", (to_meta_eq @{thm case_prod_beta'}))\ +setup \Auto2_Setup.Normalizer.add_rewr_normalizer ("rewr_case", (to_meta_eq @{thm case_prod_beta'}))\ (* Let. *) -setup \Normalizer.add_rewr_normalizer ("rewr_let", @{thm Let_def})\ +setup \Auto2_Setup.Normalizer.add_rewr_normalizer ("rewr_let", @{thm Let_def})\ -(* Equivalence relations *) +(* Equivalence relations *) setup \add_forward_prfstep @{thm Relation.symD}\ setup \add_backward_prfstep @{thm Relation.symI}\ setup \add_forward_prfstep @{thm Relation.transD}\ setup \add_backward_prfstep @{thm Relation.transI}\ (* Options *) setup \add_resolve_prfstep @{thm option.distinct(1)}\ setup \add_rewrite_rule @{thm Option.option.sel}\ setup \add_forward_prfstep @{thm option.collapse}\ setup \add_forward_prfstep (equiv_forward_th @{thm option.simps(1)})\ setup \fold (fn th => add_rewrite_rule_cond th [with_score 1]) @{thms Option.option.case}\ end diff --git a/thys/Auto2_HOL/HOL/Primes_Ex.thy b/thys/Auto2_HOL/HOL/Primes_Ex.thy --- a/thys/Auto2_HOL/HOL/Primes_Ex.thy +++ b/thys/Auto2_HOL/HOL/Primes_Ex.thy @@ -1,146 +1,146 @@ (* File: Primes_Ex.thy Author: Bohua Zhan Elementary number theory of primes, up to the proof of infinitude of primes and the unique factorization theorem. Follows the development in HOL/Computational_Algebra/Primes.thy. *) section \Primes\ theory Primes_Ex imports Auto2_Main begin subsection \Basic definition\ definition prime :: "nat \ bool" where [rewrite]: "prime p = (1 < p \ (\m. m dvd p \ m = 1 \ m = p))" lemma primeD1 [forward]: "prime p \ 1 < p" by auto2 lemma primeD2: "prime p \ m dvd p \ m = 1 \ m = p" by auto2 setup \add_forward_prfstep_cond @{thm primeD2} [with_cond "?m \ 1", with_cond "?m \ ?p"]\ setup \del_prfstep_thm_eqforward @{thm prime_def}\ (* Exists a prime p. *) theorem exists_prime [resolve]: "\p. prime p" @proof @have "prime 2" @qed lemma prime_odd_nat: "prime p \ p > 2 \ odd p" by auto2 lemma prime_imp_coprime_nat [backward2]: "prime p \ \ p dvd n \ coprime p n" by auto2 lemma prime_dvd_mult_nat: "prime p \ p dvd m * n \ p dvd m \ p dvd n" by auto2 setup \add_forward_prfstep_cond @{thm prime_dvd_mult_nat} (with_conds ["?m \ ?p", "?n \ ?p", "?m \ ?p * ?m'", "?n \ ?p * ?n'"])\ theorem prime_dvd_intro: "prime p \ p * q = m * n \ p dvd m \ p dvd n" @proof @have "p dvd m * n" @qed setup \add_forward_prfstep_cond @{thm prime_dvd_intro} (with_conds ["?m \ ?p", "?n \ ?p", "?m \ ?p * ?m'", "?n \ ?p * ?n'"])\ lemma prime_dvd_mult_eq_nat: "prime p \ p dvd m * n = (p dvd m \ p dvd n)" by auto2 lemma not_prime_eq_prod_nat [backward1]: "n > 1 \ \ prime n \ \m k. n = m * k \ 1 < m \ m < n \ 1 < k \ k < n" @proof @obtain m where "m dvd n \ m \ 1 \ m \ n" @obtain k where "n = m * k" @have "m \ m * k" @have "k \ m * k" @qed lemma prime_dvd_power_nat: "prime p \ p dvd x^n \ p dvd x" by auto2 setup \add_forward_prfstep_cond @{thm prime_dvd_power_nat} [with_cond "?p \ ?x"]\ lemma prime_dvd_power_nat_iff: "prime p \ n > 0 \ p dvd x^n \ p dvd x" by auto2 lemma prime_nat_code: "prime p = (1 < p \ (\x. 1 < x \ x < p \ \ x dvd p))" by auto2 lemma prime_factor_nat [backward]: "n \ 1 \ \p. p dvd n \ prime p" @proof @strong_induct n @case "prime n" @case "n = 0" @obtain k where "k \ 1" "k \ n" "k dvd n" @apply_induct_hyp k @qed lemma prime_divprod_pow_nat: "prime p \ coprime a b \ p^n dvd a * b \ p^n dvd a \ p^n dvd b" by auto2 lemma prime_product [forward]: "prime (p * q) \ p = 1 \ q = 1" @proof @have "p dvd q * p" @qed lemma prime_exp: "prime (p ^ n) \ n = 1 \ prime p" by auto2 lemma prime_power_mult: "prime p \ x * y = p ^ k \ \i j. x = p ^ i \ y = p ^ j" @proof @induct k arbitrary x y @with @subgoal "k = Suc k'" @case "p dvd x" @with @obtain x' where "x = p * x'" @have "x * y = p * (x' * y)" @obtain i j where "x' = p ^ i" "y = p ^ j" @have "x = p ^ Suc i" @end @case "p dvd y" @with @obtain y' where "y = p * y'" @have "x * y = p * (x * y')" @obtain i j where "x = p ^ i" "y' = p ^ j" @have "y = p ^ Suc j" @end @endgoal @end @qed subsection \Infinitude of primes\ theorem bigger_prime [resolve]: "\p. prime p \ n < p" @proof @obtain p where "prime p" "p dvd fact n + 1" @case "n \ p" @with @have "(p::nat) dvd fact n" @end @qed theorem primes_infinite: "\ finite {p. prime p}" @proof @obtain b where "prime b" "Max {p. prime p} < b" @qed subsection \Existence and uniqueness of prime factorization\ theorem factorization_exists: "n > 0 \ \M. (\p\#M. prime p) \ n = (\i\#M. i)" @proof @strong_induct n @case "n = 1" @with @have "n = (\i\# {#}. i)" @end @case "prime n" @with @have "n = (\i\# {#n#}. i)" @end @obtain m k where "n = m * k" "1 < m" "m < n" "1 < k" "k < n" @apply_induct_hyp m @obtain M where "(\p\#M. prime p)" "m = (\i\#M. i)" @apply_induct_hyp k @obtain K where "(\p\#K. prime p)" "k = (\i\#K. i)" @have "n = (\i\#(M+K). i)" @qed theorem prime_dvd_multiset [backward1]: "prime p \ p dvd (\i\#M. i) \ \n. n\#M \ p dvd n" @proof @strong_induct M @case "M = {#}" @obtain M' m where "M = M' + {#m#}" @contradiction @apply_induct_hyp M' @qed - + theorem factorization_unique_aux: "\p\#M. prime p \ \p\#N. prime p \ (\i\#M. i) dvd (\i\#N. i) \ M \# N" @proof @strong_induct M arbitrary N @case "M = {#}" @obtain M' m where "M = M' + {#m#}" @have "m dvd (\i\#M. i)" @obtain n where "n \# N" "m dvd n" @obtain N' where "N = N' + {#n#}" @have "m = n" @have "(\i\#M'. i) dvd (\i\#N'. i)" @apply_induct_hyp M' N' @qed setup \add_forward_prfstep_cond @{thm factorization_unique_aux} [with_cond "?M \ ?N"]\ theorem factorization_unique: "\p\#M. prime p \ \p\#N. prime p \ (\i\#M. i) = (\i\#N. i) \ M = N" @proof @have "M \# N" @qed setup \del_prfstep_thm @{thm factorization_unique_aux}\ end diff --git a/thys/Auto2_HOL/HOL/Set_Thms.thy b/thys/Auto2_HOL/HOL/Set_Thms.thy --- a/thys/Auto2_HOL/HOL/Set_Thms.thy +++ b/thys/Auto2_HOL/HOL/Set_Thms.thy @@ -1,197 +1,197 @@ (* File: Set_Thms.thy Author: Bohua Zhan Setup of proof steps related to sets. *) section \Setup for sets and multisets\ theory Set_Thms imports Logic_Thms "HOL-Library.Multiset" begin subsection \Set\ subsubsection \Injective functions\ setup \add_backward_prfstep @{thm injI}\ subsubsection \AC property of intersection and union\ -setup \fold ACUtil.add_ac_data [ +setup \fold Auto2_HOL_Extra_Setup.ACUtil.add_ac_data [ {cfhead = @{cterm inf}, unit = SOME @{cterm inf}, assoc_th = @{thm inf_assoc}, comm_th = @{thm inf_commute}, unitl_th = @{thm inf_top_left}, unitr_th = @{thm inf_top_right}}, {cfhead = @{cterm sup}, unit = SOME @{cterm bot}, assoc_th = @{thm sup_assoc}, comm_th = @{thm sup_commute}, unitl_th = @{thm sup_bot_left}, unitr_th = @{thm sup_bot_right}}] \ subsubsection \Collection and bounded quantification\ setup \add_rewrite_rule @{thm Set.mem_Collect_eq}\ lemma ball_single [rewrite]: "(\x\{x}. P x) = P x" by auto subsubsection \Membership\ -setup \add_rewrite_rule @{thm Set.singleton_iff}\ +setup \add_rewrite_rule @{thm Set.singleton_iff}\ setup \add_forward_prfstep (equiv_forward_th @{thm Set.empty_iff})\ lemma set_membership_distinct [forward]: "x \ s \ y \ s \ x \ y" by auto lemma non_empty_exist_elt [backward]: "U \ {} \ \x. x \ U" by blast lemma non_univ_exist_compl [backward]: "U \ UNIV \ \x. x \ U" by blast setup \add_resolve_prfstep @{thm Set.UNIV_I}\ subsubsection \Insert\ setup \add_backward_prfstep_cond (equiv_backward_th @{thm Set.insert_iff}) [with_cond "?A \ {}"]\ setup \add_forward_prfstep_cond (equiv_forward_th @{thm Set.insert_iff}) [with_score 500, with_cond "?A \ {}"]\ setup \add_forward_prfstep_cond (equiv_forward_th @{thm Set.insert_subset}) [with_cond "?A \ {}"]\ setup \add_backward_prfstep_cond (equiv_backward_th @{thm Set.insert_subset}) [with_score 500, with_cond "?A \ {}"]\ subsubsection \Extensionality\ lemma set_ext [forward]: "\a. a \ S \ a \ T \ S = T" by auto setup \add_backward_prfstep_cond @{thm set_ext} [with_score 500, with_filt (order_filter "S" "T")]\ lemma set_pair_ext [forward]: "\a b. (a, b) \ S \ (a, b) \ T \ S = T" by auto subsubsection \Union\ setup \add_forward_prfstep_cond (equiv_forward_th @{thm Set.Un_iff}) [with_score 500]\ setup \add_backward_prfstep (equiv_backward_th @{thm Set.Un_iff})\ lemma UnD1 [forward]: "c \ A \ B \ c \ A \ c \ B" by auto lemma UnD2 [forward]: "c \ A \ B \ c \ B \ c \ A" by auto lemma UnD1_single [forward]: "c \ {a} \ B \ c \ a \ c \ B" by auto lemma UnD2_single [forward]: "c \ A \ {b} \ c \ b \ c \ A" by auto setup \add_forward_prfstep_cond @{thm Set.UnI1} [with_term "?A \ ?B"]\ setup \add_forward_prfstep_cond @{thm Set.UnI2} [with_term "?A \ ?B"]\ lemma UnI1_single: "a \ {a} \ B" by auto lemma UnI2_single: "b \ A \ {b}" by auto setup \add_forward_prfstep_cond @{thm UnI1_single} [with_term "{?a} \ ?B"]\ setup \add_forward_prfstep_cond @{thm UnI2_single} [with_term "?A \ {?b}"]\ - + lemma union_single_eq [rewrite, backward]: "x \ p \ {x} \ p = p" by auto subsubsection \Intersection\ setup \add_forward_prfstep (equiv_forward_th @{thm Set.Int_iff})\ setup \add_backward_prfstep_cond (equiv_backward_th @{thm Set.Int_iff}) [with_score 500]\ setup \add_rewrite_rule @{thm Set.Int_empty_left}\ setup \add_rewrite_rule @{thm Set.Int_empty_right}\ setup \add_rewrite_rule @{thm Set.Int_absorb}\ lemma set_disjoint_mp [forward, backward2]: "A \ B = {} \ p \ A \ p \ B" by auto lemma set_disjoint_single [rewrite]: "{x} \ B = {} \ x \ B" by simp subsubsection \subset\ setup \add_forward_prfstep @{thm subsetI}\ setup \add_backward_prfstep_cond @{thm subsetI} [with_score 500]\ setup \add_resolve_prfstep @{thm empty_subsetI}\ setup \add_forward_prfstep @{thm subsetD}\ lemma subset_single [rewrite]: "{a} \ B \ a \ B" by simp setup \add_resolve_prfstep @{thm Set.basic_monos(1)}\ setup \add_resolve_prfstep @{thm Set.Un_upper1}\ setup \add_resolve_prfstep @{thm Set.Un_upper2}\ lemma union_is_subset [forward]: "A \ B \ C \ A \ C \ B \ C" by simp setup \add_backward1_prfstep @{thm Set.Un_least}\ setup \add_backward2_prfstep @{thm Set.Un_least}\ lemma subset_union_same1 [backward]: "B \ C \ A \ B \ A \ C" by auto lemma subset_union_same2 [backward]: "A \ B \ A \ C \ B \ C" by auto subsubsection \Diff\ setup \add_forward_prfstep (equiv_forward_th @{thm Set.Diff_iff})\ setup \add_backward_prfstep_cond (equiv_backward_th @{thm Set.Diff_iff}) [with_score 500]\ setup \add_rewrite_rule @{thm Set.empty_Diff}\ lemma mem_diff [rewrite]: "x \ A - B \ x \ A \ x \ B" by simp lemma set_union_minus_same1 [rewrite]: "(A \ B) - B = A - B" by auto lemma set_union_minus_same2 [rewrite]: "(B \ A) - B = A - B" by auto lemma set_union_minus_distinct [rewrite]: "a \ c \ {a} \ (B - {c}) = {a} \ B - {c}" by auto setup \add_forward_prfstep_cond @{thm Set.Diff_subset} [with_term "?A - ?B"]\ lemma union_subtract_elt1 [rewrite]: "x \ B \ ({x} \ B) - {x} = B" by simp lemma union_subtract_elt2 [rewrite]: "x \ B \ (B \ {x}) - {x} = B" by simp lemma subset_sub1 [backward]: "x \ A \ A - {x} \ A" by auto lemma member_notin [forward]: "x \ S - {y} \ x \ y" by simp lemma member_notin_contra: "x \ S \ x \ y \ x \ S - {y}" by simp setup \add_forward_prfstep_cond @{thm member_notin_contra} [with_term "?S - {?y}"]\ subsubsection \Results on finite sets\ setup \add_resolve_prfstep @{thm Finite_Set.finite.emptyI}\ lemma set_finite_single [resolve]: "finite {x}" by simp setup \add_rewrite_rule @{thm Finite_Set.finite_Un}\ lemma Max_ge' [forward]: "finite A \ x > Max A \ \(x \ A)" using Max_ge leD by auto setup \add_backward_prfstep @{thm finite_image_set}\ setup \add_forward_prfstep @{thm finite_atLeastAtMost}\ setup \add_forward_prfstep @{thm rev_finite_subset}\ setup \add_backward1_prfstep @{thm rev_finite_subset}\ subsubsection \Cardinality\ setup \add_rewrite_rule @{thm card.empty}\ lemma card_emptyD [rewrite]: "finite S \ card S = 0 \ S = {}" by simp lemma card_minus1 [rewrite]: "x \ S \ card (S - {x}) = card S - 1" by (simp add: card_Diff_subset) setup \add_forward_prfstep @{thm finite_Diff}\ setup \add_resolve_prfstep @{thm card_mono}\ subsubsection \Image set\ setup \add_rewrite_rule @{thm Set.image_Un}\ setup \add_rewrite_rule @{thm image_set_diff}\ subsection \Multiset\ subsubsection \Basic properties\ lemma mset_member_empty [resolve]: "\p \# {#}" by simp lemma mem_multiset_single [rewrite]: "x \# {#y#} \ x = y" by simp setup \add_backward2_prfstep @{thm subset_mset.antisym}\ setup \add_resolve_prfstep @{thm Multiset.empty_le}\ setup \add_forward_prfstep @{thm mset_subsetD}\ lemma multi_contain_add_self1 [resolve]: "A \# {#x#} + A" by simp lemma multi_contain_add_self2 [resolve]: "A \# A + {#x#}" by simp setup \add_forward_prfstep_cond @{thm Multiset.multi_member_this} [with_term "{#?x#} + ?XS"]\ lemma multi_member_this2: "x \# XS + {#x#}" by simp setup \add_forward_prfstep_cond @{thm multi_member_this2} [with_term "?XS + {#?x#}"]\ setup \add_backward_prfstep @{thm Multiset.subset_mset.add_left_mono}\ setup \add_backward_prfstep @{thm Multiset.subset_mset.add_right_mono}\ subsubsection \Case checking and induction\ lemma multi_nonempty_split' [resolve]: "M \ {#} \ \M' m. M = M' + {#m#}" using multi_nonempty_split by auto lemma multi_member_split' [backward]: "x \# M \ \M'. M = M' + {#x#}" by (metis insert_DiffM2) setup \add_strong_induct_rule @{thm full_multiset_induct}\ subsubsection \Results on mset\ setup \add_rewrite_rule @{thm set_mset_empty}\ setup \add_rewrite_rule @{thm set_mset_single}\ setup \add_rewrite_rule @{thm set_mset_union}\ setup \add_rewrite_rule @{thm image_mset_empty}\ setup \add_rewrite_rule @{thm image_mset_single}\ setup \add_rewrite_rule @{thm image_mset_union}\ setup \add_rewrite_rule @{thm prod_mset_empty}\ setup \add_rewrite_rule @{thm prod_mset_singleton}\ setup \add_rewrite_rule @{thm prod_mset_Un}\ subsubsection \Set interval\ setup \add_rewrite_rule @{thm Set_Interval.ord_class.lessThan_iff}\ setup \add_rewrite_rule @{thm Set_Interval.ord_class.atLeastAtMost_iff}\ end diff --git a/thys/Auto2_HOL/HOL/ac_steps.ML b/thys/Auto2_HOL/HOL/ac_steps.ML --- a/thys/Auto2_HOL/HOL/ac_steps.ML +++ b/thys/Auto2_HOL/HOL/ac_steps.ML @@ -1,230 +1,234 @@ (* File: ac_steps.ML Author: Bohua Zhan Proof steps related to associative-commutative operations. *) signature AC_PROOFSTEPS = sig val max_ac: int Config.T val simp_ac_expr: Proof.context -> ac_info -> box_id * cterm -> (box_id * thm) list val get_ac_head_equiv: Proof.context -> ac_info -> box_id * cterm -> (box_id * thm) list val ac_expand_once: Proof.context -> ac_info -> box_id * cterm -> (box_id * thm) list val ac_expand: Proof.context -> ac_info -> box_id * cterm -> (box_id * thm) list val ac_expand_equiv: proofstep val ac_expand_unit: proofstep val add_ac_proofsteps: theory -> theory end; -structure AC_ProofSteps : AC_PROOFSTEPS = +functor AC_ProofSteps( + structure ACUtil: ACUTIL; + structure Basic_UtilLogic: BASIC_UTIL_LOGIC; + structure ProofStepData: PROOFSTEP_DATA; + structure RewriteTable: REWRITE_TABLE; + structure Update: UPDATE; + ): AC_PROOFSTEPS = struct val max_ac = Attrib.setup_config_int @{binding "max_ac"} (K 20) fun rewrite_on_eqs ac_info eqs ct = let val t = Thm.term_of ct in if ACUtil.head_agrees ac_info t then Conv.binop_conv (rewrite_on_eqs ac_info eqs) ct else case find_first (fn eq => t aconv (Util.lhs_of eq)) eqs of NONE => Conv.all_conv ct | SOME eq_th => eq_th end fun dest_ac_full ac_info ct = if ACUtil.head_agrees ac_info (Thm.term_of ct) then let val (a1, a2) = Util.dest_binop_cargs ct val sort_opt = if ACUtil.has_comm_th ac_info then sort (Term_Ord.term_ord o apply2 Thm.term_of) else I in sort_opt (dest_ac_full ac_info a1 @ dest_ac_full ac_info a2) end else if ACUtil.eq_unit ac_info (Thm.term_of ct) then [] else [ct] (* Simplify each term of the AC expression. *) fun simp_ac_expr ctxt ac_info (id, cu) = let val cus = dest_ac_full ac_info cu in if null cus then [(id, Thm.reflexive cu)] else cus |> map (RewriteTable.simplify_info ctxt) |> BoxID.get_all_merges_info ctxt |> BoxID.merge_box_with_info ctxt id |> map (fn (id', eqs) => (id', rewrite_on_eqs ac_info eqs cu)) end (* Obtain head equivalences of cu, where each term is simplified. *) fun get_ac_head_equiv ctxt ac_info (id, cu) = let fun process_head_equiv (id', eq_th) = let val infos = simp_ac_expr ctxt ac_info (id', Thm.rhs_of eq_th) in map (BoxID.merge_eq_infos ctxt (id', eq_th)) infos end in cu |> RewriteTable.get_head_equiv ctxt |> BoxID.merge_box_with_info ctxt id |> maps process_head_equiv |> filter_out (Thm.is_reflexive o snd) end (* Find ways to modify ct once by rewriting one of the subterms. *) fun ac_expand_once ctxt ac_info (id, ct) = let val cus = dest_ac_full ac_info ct fun get_equiv cu = get_ac_head_equiv ctxt ac_info (id, cu) fun process_info (id', eq) = (id', rewrite_on_eqs ac_info [eq] ct) in map process_info (maps get_equiv cus) end (* Find all ways to write ct, up to a certain limit. *) fun ac_expand ctxt ac_info (id, ct) = let val max_ac = Config.get ctxt max_ac fun ac_equiv_eq_better (id, th) (id', th') = let val seq1 = dest_ac_full ac_info (Thm.rhs_of th) val seq2 = dest_ac_full ac_info (Thm.rhs_of th') in Util.is_subseq (op aconv o apply2 Thm.term_of) (seq1, seq2) andalso BoxID.is_eq_ancestor ctxt id id' end fun has_ac_equiv_eq_better infos info' = exists (fn info => ac_equiv_eq_better info info') infos fun helper (old, new) = case new of [] => old | (id', eq_th) :: rest => if length old + length new > max_ac then old @ take (max_ac - length old) new else let val old' = ((id', eq_th) :: old) val rhs_expand = (ac_expand_once ctxt ac_info (id', Thm.rhs_of eq_th)) |> Util.max_partial ac_equiv_eq_better |> map (BoxID.merge_eq_infos ctxt (id', eq_th)) |> filter_out (has_ac_equiv_eq_better (old' @ rest)) in helper (old', rest @ rhs_expand) end (* Start term *) val start = simp_ac_expr ctxt ac_info (id, ct) in helper ([], start) end fun ac_expand_equiv_fn ctxt item1 item2 = let val thy = Proof_Context.theory_of ctxt val {id = id1, tname = tname1, ...} = item1 val {id = id2, tname = tname2, ...} = item2 val (ct1, ct2) = (the_single tname1, the_single tname2) val (t1, t2) = (Thm.term_of ct1, Thm.term_of ct2) val id = BoxID.merge_boxes ctxt (id1, id2) in case ACUtil.get_head_ac_info thy t1 of NONE => [] | SOME ac_info => if not (ACUtil.head_agrees ac_info t2) then [] else if Term_Ord.term_ord (t2, t1) = LESS then [] else if RewriteTable.is_equiv id ctxt (ct1, ct2) then [] else let val expand1 = ac_expand ctxt ac_info (id, ct1) val expand2 = ac_expand ctxt ac_info (id, ct2) fun get_equiv ((id1, eq_th1), (id2, eq_th2)) = let val ct1 = Thm.rhs_of eq_th1 val ct2 = Thm.rhs_of eq_th2 val ts1 = dest_ac_full ac_info ct1 val ts2 = dest_ac_full ac_info ct2 in if eq_list (op aconv o apply2 Thm.term_of) (ts1, ts2) then let val eq_th1' = ACUtil.normalize_all_ac ac_info ct1 val eq_th2' = ACUtil.normalize_all_ac ac_info ct2 val id' = BoxID.merge_boxes ctxt (id1, id2) val eq = Util.transitive_list [ eq_th1, eq_th1', meta_sym eq_th2', meta_sym eq_th2] in - [(id', to_obj_eq eq)] + [(id', Basic_UtilLogic.to_obj_eq eq)] end else [] end in (maps get_equiv (Util.all_pairs (expand1, expand2))) |> filter (BoxID.has_incr_id o fst) |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) |> map (fn (id, th) => Update.thm_update_sc 1 (id, th)) end end val ac_expand_equiv = {name = "ac_expand_equiv", args = [TypedUniv TY_TERM, TypedUniv TY_TERM], func = TwoStep ac_expand_equiv_fn} fun ac_expand_unit_fn ctxt item = let val thy = Proof_Context.theory_of ctxt val {id, tname, ...} = item val ct = the_single tname val t = Thm.term_of ct in case ACUtil.get_head_ac_info thy t of NONE => [] | SOME ac_info => let val expand = ac_expand ctxt ac_info (id, ct) fun process_expand (id', eq_th) = let val ct = Thm.rhs_of eq_th val ts = dest_ac_full ac_info ct in if length ts <= 1 then let val eq' = ACUtil.normalize_all_ac ac_info ct val eq = Util.transitive_list [eq_th, eq'] in - [(id', to_obj_eq eq)] + [(id', Basic_UtilLogic.to_obj_eq eq)] end else [] end in (maps process_expand expand) |> filter (BoxID.has_incr_id o fst) |> map (fn (id, th) => Update.thm_update_sc 1 (id, th)) end end val ac_expand_unit = {name = "ac_expand_unit", args = [TypedUniv TY_TERM], func = OneStep ac_expand_unit_fn} val add_ac_proofsteps = - fold add_prfstep [ + fold ProofStepData.add_prfstep [ ac_expand_equiv, ac_expand_unit ] end (* structure AC_ProofSteps. *) - -val _ = Theory.setup AC_ProofSteps.add_ac_proofsteps diff --git a/thys/Auto2_HOL/HOL/acdata.ML b/thys/Auto2_HOL/HOL/acdata.ML --- a/thys/Auto2_HOL/HOL/acdata.ML +++ b/thys/Auto2_HOL/HOL/acdata.ML @@ -1,333 +1,338 @@ (* File: acdata.ML Author: Bohua Zhan Dealing with associative-commutative operations. *) (* Data for an AC function. *) type ac_info = { cfhead: cterm, unit: cterm option, assoc_th: thm, (* (a . b) . c = a . (b . c) *) comm_th: thm, (* a . b = b . a *) unitl_th: thm, (* e . a = a *) unitr_th: thm (* a . e = a *) } signature ACUTIL = sig val inst_ac_info: theory -> typ -> ac_info -> ac_info option val head_agrees: ac_info -> term -> bool val eq_unit: ac_info -> term -> bool val add_ac_data: ac_info -> theory -> theory val get_head_ac_info: theory -> term -> ac_info option val has_assoc_th: ac_info -> bool val has_comm_th: ac_info -> bool val has_unit_th: ac_info -> bool val comm_cv: ac_info -> conv val assoc_cv: ac_info -> conv val assoc_sym_cv: ac_info -> conv val swap_cv: ac_info -> conv val swap_r_cv: ac_info -> conv val dest_ac: ac_info -> term -> term list val cdest_ac: ac_info -> cterm -> cterm list val comb_ac_equiv: ac_info -> thm list -> thm val normalize_assoc: ac_info -> conv val move_outmost: ac_info -> term -> conv val normalize_unit: ac_info -> conv val normalize_comm_gen: ac_info -> (term * term -> bool) -> conv val normalize_comm: ac_info -> conv val normalize_au: ac_info -> conv val normalize_all_ac: ac_info -> conv val ac_last_conv: ac_info -> conv -> conv val norm_combine: ac_info -> (term -> bool) -> conv -> conv end; -structure ACUtil : ACUTIL = +functor ACUtil( + structure Basic_UtilBase: BASIC_UTIL_BASE + structure Basic_UtilLogic: BASIC_UTIL_LOGIC + ): ACUTIL = struct (* Register of generators of ac_inst_info. *) structure Data = Theory_Data ( type T = ac_info Symtab.table; val empty = Symtab.empty; val merge = Symtab.merge (fn (info1, info2) => #cfhead info1 aconvc #cfhead info2) ) (* Instantiate an ac_info for a specific type T. *) fun inst_ac_info thy T {assoc_th, comm_th, unitl_th, unitr_th, ...} = let (* Instantiate th to having argument of type T. If not possible, change th to true_th. *) fun inst_th th = - if is_true_th th then true_th else + if Basic_UtilLogic.is_true_th th then Basic_UtilBase.true_th else let (* Extract the first argument of th, then the body type of that argument. *) - val arg_type = th |> prop_of' |> Util.dest_args |> hd + val arg_type = th |> Basic_UtilLogic.prop_of' |> Util.dest_args |> hd |> Term.type_of |> Term.body_type in if arg_type = T then th else let val tenv = Sign.typ_match thy (arg_type, T) Vartab.empty in Util.subst_thm_thy thy (tenv, Vartab.empty) th end - handle Type.TYPE_MATCH => true_th + handle Type.TYPE_MATCH => Basic_UtilBase.true_th end val assoc_th' = inst_th assoc_th val unitl_th' = inst_th unitl_th in - if is_true_th assoc_th' then NONE else - SOME {cfhead = assoc_th' |> cprop_of' |> Thm.dest_arg1 |> Thm.dest_fun2, - unit = if is_true_th unitl_th' then NONE - else SOME (unitl_th' |> cprop_of' |> Thm.dest_arg1 |> Thm.dest_arg1), + if Basic_UtilLogic.is_true_th assoc_th' then NONE else + SOME {cfhead = assoc_th' |> Basic_UtilLogic.cprop_of' |> Thm.dest_arg1 |> Thm.dest_fun2, + unit = if Basic_UtilLogic.is_true_th unitl_th' then NONE + else + SOME (unitl_th' |> Basic_UtilLogic.cprop_of' |> Thm.dest_arg1 |> Thm.dest_arg1), assoc_th = assoc_th', comm_th = inst_th comm_th, unitl_th = unitl_th', unitr_th = inst_th unitr_th} end fun head_agrees {cfhead, ...} t = Util.is_head (Thm.term_of cfhead) t fun eq_unit {unit, ...} t = case unit of NONE => false | SOME ct' => t aconv (Thm.term_of ct') (* Add the given ac_info under the given name. *) fun add_ac_data info thy = let val {assoc_th, ...} = info - val f = assoc_th |> prop_of' |> dest_eq |> snd |> Term.head_of + val f = assoc_th |> Basic_UtilLogic.prop_of' |> Basic_UtilBase.dest_eq |> snd |> Term.head_of in case f of Const (c, _) => let val _ = tracing ("Add ac data for function " ^ c) in Data.map (Symtab.update_new (c, info)) thy end | _ => error "Add AC data: invalid assoc_th" end handle Symtab.DUP _ => error "Add AC data: info already exists." fun get_head_ac_info thy t = case t of Const (c, _) $ _ $ _ => (case Symtab.lookup (Data.get thy) c of NONE => NONE | SOME ac_info => inst_ac_info thy (fastype_of t) ac_info) | _ => NONE -fun has_assoc_th {assoc_th, ...} = not (is_true_th assoc_th) -fun has_comm_th {comm_th, ...} = not (is_true_th comm_th) -fun has_unit_th {unitl_th, ...} = not (is_true_th unitl_th) -fun comm_cv {comm_th, ...} = rewr_obj_eq comm_th -fun assoc_cv {assoc_th, ...} = rewr_obj_eq assoc_th -fun assoc_sym_cv {assoc_th, ...} = rewr_obj_eq (obj_sym assoc_th) +fun has_assoc_th {assoc_th, ...} = not (Basic_UtilLogic.is_true_th assoc_th) +fun has_comm_th {comm_th, ...} = not (Basic_UtilLogic.is_true_th comm_th) +fun has_unit_th {unitl_th, ...} = not (Basic_UtilLogic.is_true_th unitl_th) +fun comm_cv {comm_th, ...} = Basic_UtilLogic.rewr_obj_eq comm_th +fun assoc_cv {assoc_th, ...} = Basic_UtilLogic.rewr_obj_eq assoc_th +fun assoc_sym_cv {assoc_th, ...} = Basic_UtilLogic.rewr_obj_eq (Basic_UtilLogic.obj_sym assoc_th) (* (a . b) . c = (a . c) . b *) fun swap_cv (ac_info as {assoc_th, comm_th, ...}) ct = if head_agrees ac_info (dest_arg1 (Thm.term_of ct)) then - Conv.every_conv [rewr_obj_eq assoc_th, - Conv.arg_conv (rewr_obj_eq comm_th), - rewr_obj_eq (obj_sym assoc_th)] ct + Conv.every_conv [Basic_UtilLogic.rewr_obj_eq assoc_th, + Conv.arg_conv (Basic_UtilLogic.rewr_obj_eq comm_th), + Basic_UtilLogic.rewr_obj_eq (Basic_UtilLogic.obj_sym assoc_th)] ct else - rewr_obj_eq comm_th ct + Basic_UtilLogic.rewr_obj_eq comm_th ct (* a . (b . c) = b . (a . c) *) fun swap_r_cv (ac_info as {assoc_th, comm_th, ...}) ct = if head_agrees ac_info (dest_arg (Thm.term_of ct)) then - Conv.every_conv [rewr_obj_eq (obj_sym assoc_th), - Conv.arg1_conv (rewr_obj_eq comm_th), - rewr_obj_eq assoc_th] ct + Conv.every_conv [Basic_UtilLogic.rewr_obj_eq (Basic_UtilLogic.obj_sym assoc_th), + Conv.arg1_conv (Basic_UtilLogic.rewr_obj_eq comm_th), + Basic_UtilLogic.rewr_obj_eq assoc_th] ct else - rewr_obj_eq comm_th ct + Basic_UtilLogic.rewr_obj_eq comm_th ct (* Destruct t, assuming it is associated to the left. *) fun dest_ac ac_info t = let fun dest t = if head_agrees ac_info t then let val (a1, a2) = Util.dest_binop_args t in a2 :: dest a1 end else [t] in rev (dest t) end fun cdest_ac ac_info ct = let fun dest ct = if head_agrees ac_info (Thm.term_of ct) then let val (a1, a2) = Util.dest_binop_cargs ct in a2 :: dest a1 end else [ct] in rev (dest ct) end (* Given ths: [A1 == B1, ..., An == Bn], get theorem A1...An == B1...Bn. Associate to the left only. *) fun comb_ac_equiv {cfhead, ...} ths = let fun binop_comb th1 th2 = Thm.combination (Thm.combination (Thm.reflexive cfhead) th1) th2 (* Combine in the reverse order. *) fun comb ths = case ths of [] => raise Fail "comb_ac_equiv: empty list" | [th] => th | [th1, th2] => binop_comb th2 th1 | th :: ths' => binop_comb (comb ths') th in comb (rev ths) end (* Normalize association with the given direction. *) fun normalize_assoc ac_info ct = if not (has_assoc_th ac_info) then Conv.all_conv ct else let val {assoc_th, ...} = ac_info (* First rewrite into form (...) * a, then rewrite the remaining parts. *) fun normalize ct = if head_agrees ac_info (Thm.term_of ct) then - Conv.every_conv [Conv.repeat_conv (rewr_obj_eq (obj_sym assoc_th)), - Conv.arg1_conv normalize] ct + Conv.every_conv [ + Conv.repeat_conv (Basic_UtilLogic.rewr_obj_eq (Basic_UtilLogic.obj_sym assoc_th)), + Conv.arg1_conv normalize] ct else Conv.all_conv ct in normalize ct end (* Move the given u within ct to the rightmost position. Assume associate to the left. *) fun move_outmost (ac_info as {comm_th, ...}) u ct = if not (has_assoc_th ac_info andalso has_comm_th ac_info) then raise Fail "move_outmost: commutativity is not available." else if u aconv (Thm.term_of ct) then Conv.all_conv ct else if not (head_agrees ac_info (Thm.term_of ct)) then raise Fail "move_outmost: u not found in ct." else let val (a, b) = Util.dest_binop_args (Thm.term_of ct) in if u aconv b then Conv.all_conv ct else if head_agrees ac_info a then ((Conv.arg1_conv (move_outmost ac_info u)) then_conv (swap_cv ac_info)) ct else if u aconv a then - rewr_obj_eq comm_th ct + Basic_UtilLogic.rewr_obj_eq comm_th ct else raise Fail "move_outmost: u not found in ct." end (* In a product of a_1, a_2, ..., remove any a_i that is a unit. *) fun normalize_unit (ac_info as {unitl_th, unitr_th, ...}) ct = if not (has_unit_th ac_info) then Conv.all_conv ct else let fun normalize ct = if head_agrees ac_info (Thm.term_of ct) then Conv.every_conv [Conv.binop_conv normalize, - Conv.try_conv (rewr_obj_eq unitl_th), - Conv.try_conv (rewr_obj_eq unitr_th)] ct + Conv.try_conv (Basic_UtilLogic.rewr_obj_eq unitl_th), + Conv.try_conv (Basic_UtilLogic.rewr_obj_eq unitr_th)] ct else Conv.all_conv ct in normalize ct end (* Rearrange subterms of ct according to the given term ordering. Returns theorem ct == ct'. *) fun normalize_comm_gen (ac_info as {comm_th, ...}) termless ct = if not (has_comm_th ac_info) then Conv.all_conv ct else let (* If there are two terms a.b, swap if a > b. If there are at least three terms, in the left associate case this is (a.b).c, swap b and c if b > c. If there is a swap, recursively call swap_last until the original outside term is swapped into position. *) fun swap_last ct = if head_agrees ac_info (Thm.term_of ct) then let val (a1, a2) = Util.dest_binop_args (Thm.term_of ct) in if head_agrees ac_info a1 then (* Structure of t is a1 . a2 = (_ . b2) . a2. *) if termless (a2, dest_arg a1) then ((swap_cv ac_info) then_conv (Conv.arg1_conv swap_last)) ct else Conv.all_conv ct else (* Structure of t is a1 . a2. *) - if termless (a2, a1) then rewr_obj_eq comm_th ct + if termless (a2, a1) then Basic_UtilLogic.rewr_obj_eq comm_th ct else Conv.all_conv ct end else Conv.all_conv ct (* Full ordering. Recursively perform full ordering on all but the outermost, then swap outermost into position. *) fun normalize ct = if head_agrees ac_info (Thm.term_of ct) then ((Conv.arg1_conv normalize) then_conv swap_last) ct else Conv.all_conv ct in normalize ct end fun normalize_comm ac_info = normalize_comm_gen ac_info (fn (s, t) => Term_Ord.term_ord (s, t) = LESS) (* Normalize all except comm. *) fun normalize_au ac_info = Conv.every_conv [normalize_unit ac_info, normalize_assoc ac_info] (* Normalize everything. *) fun normalize_all_ac ac_info = Conv.every_conv [normalize_au ac_info, normalize_comm ac_info] (* Rewrite the last term in ct using cv. Assume associative to left. *) fun ac_last_conv ac_info cv ct = if head_agrees ac_info (Thm.term_of ct) then Conv.arg_conv cv ct else cv ct (* Given ct in the form x_1 * ... * x_n, where some sequence of x_i satisfies predicate pred. Combine these x_i into a single term using the binary combinator cv. *) fun norm_combine ac_info pred cv ct = let val t = Thm.term_of ct in if head_agrees ac_info t then let val (a, b) = Util.dest_binop_args t in if pred b then if pred a then cv ct else if head_agrees ac_info a andalso pred (dest_arg a) then Conv.every_conv [assoc_cv ac_info, Conv.arg_conv cv, norm_combine ac_info pred cv] ct else Conv.all_conv ct else Conv.arg1_conv (norm_combine ac_info pred cv) ct end else Conv.all_conv ct end end (* structure ACUtil. *) diff --git a/thys/Auto2_HOL/HOL/acdata_test.ML b/thys/Auto2_HOL/HOL/acdata_test.ML --- a/thys/Auto2_HOL/HOL/acdata_test.ML +++ b/thys/Auto2_HOL/HOL/acdata_test.ML @@ -1,82 +1,82 @@ (* File: acdata_test.ML Author: Bohua Zhan Unit test for acdata.ML. *) local val thy = @{theory} val ctxt = @{context} in val test_comb_ac_equiv = let fun err n = "test_comb_ac_equiv: " ^ (string_of_int n) fun test n (t1, t2) = let - val ac_info = the (ACUtil.get_head_ac_info thy t1) + val ac_info = the (Auto2_HOL_Extra_Setup.ACUtil.get_head_ac_info thy t1) handle Option.Option => raise Fail "test_comb_ac_equiv: ac_info" - val ts1 = ACUtil.dest_ac ac_info t1 - val ts2 = ACUtil.dest_ac ac_info t2 + val ts1 = Auto2_HOL_Extra_Setup.ACUtil.dest_ac ac_info t1 + val ts2 = Auto2_HOL_Extra_Setup.ACUtil.dest_ac ac_info t2 fun eq (t1', t2') = Thm.assume (Thm.global_cterm_of thy (Logic.mk_equals (t1', t2'))) val eqs = map eq (ts1 ~~ ts2) - val th = ACUtil.comb_ac_equiv ac_info eqs + val th = Auto2_HOL_Extra_Setup.ACUtil.comb_ac_equiv ac_info eqs in if Thm.prop_of th aconv Logic.mk_equals (t1, t2) then () else raise Fail (err n) end val _ = test 0 (@{term "(a::nat) + b + c"}, @{term "(d::nat) + e + f"}) in () end (* Generic function for testing conv with ac_info argument. *) fun test_ac_conv ctxt cv err_str (str1, str2) = let val t1 = Proof_Context.read_term_pattern ctxt str1 in - case ACUtil.get_head_ac_info thy t1 of + case Auto2_HOL_Extra_Setup.ACUtil.get_head_ac_info thy t1 of NONE => let val _ = trace_t ctxt "t1:" t1 in raise Fail "test_ac_conv: ac_info" end | SOME ac_info => Util.test_conv ctxt (cv ac_info) err_str (str1, str2) end val test_normalize_assoc = let val test_data = [ ("(a::nat) + (b + d) + (c + e)", "(a::nat) + b + d + c + e"), ("((a::nat) + 0) + (b + 0)", "(a::nat) + 0 + b + 0") ] in - map (test_ac_conv ctxt ACUtil.normalize_assoc "test_normalize_assoc") test_data + map (test_ac_conv ctxt Auto2_HOL_Extra_Setup.ACUtil.normalize_assoc "test_normalize_assoc") test_data end val test_move_outmost = let val err_str = "test_move_outmost" fun test (stru, (str1, str2)) = let val (t1, u) = (Proof_Context.read_term_pattern ctxt str1, Proof_Context.read_term_pattern ctxt stru) val ac_info = - the (ACUtil.get_head_ac_info thy t1) + the (Auto2_HOL_Extra_Setup.ACUtil.get_head_ac_info thy t1) handle Option.Option => raise Fail (err_str ^ ": ac_info") in - Util.test_conv ctxt (ACUtil.move_outmost ac_info u) err_str (str1, str2) + Util.test_conv ctxt (Auto2_HOL_Extra_Setup.ACUtil.move_outmost ac_info u) err_str (str1, str2) end val test_data = [ ("a::nat", ("(a::nat) + b", "(b::nat) + a")), ("a::nat", ("(a::nat) + b + c", "(b::nat) + c + a")), ("a::nat", ("(b::nat) + a", "(b::nat) + a")), ("a::nat", ("(b::nat) + a + c", "(b::nat) + c + a")) ] in map test test_data end end; (* local *) diff --git a/thys/Auto2_HOL/HOL/arith.ML b/thys/Auto2_HOL/HOL/arith.ML --- a/thys/Auto2_HOL/HOL/arith.ML +++ b/thys/Auto2_HOL/HOL/arith.ML @@ -1,169 +1,169 @@ (* File: arith.ML Author: Bohua Zhan Arithmetic proof steps. *) signature NAT_UTIL = sig val lookup_numc: Type.tyenv * Envir.tenv -> int -> int val lookup_numc0: Type.tyenv * Envir.tenv -> int val lookup_numc1: Type.tyenv * Envir.tenv -> int val lookup_numc2: Type.tyenv * Envir.tenv -> int val mk_nat: int -> term val mk_int: int -> term val nat0: term val cnat0: cterm val mk_less: term * term -> term val mk_le: term * term -> term val nat_le_th: int -> int -> thm val nat_less_th: int -> int -> thm val nat_neq_th: int -> int -> thm val nat_fold_reduce: term -> term val nat_fold_conv: conv val plus_ac_on_typ: theory -> typ -> ac_info val times_ac_on_typ: theory -> typ -> ac_info val add_arith_ac_data: theory -> theory val add_arith_proofsteps: theory -> theory end; structure Nat_Util : NAT_UTIL = struct fun lookup_numc inst n = UtilArith.dest_numc (lookup_instn inst ("NUMC", n)) fun lookup_numc0 inst = lookup_numc inst 0 fun lookup_numc1 inst = lookup_numc inst 1 fun lookup_numc2 inst = lookup_numc inst 2 fun mk_nat n = HOLogic.mk_number natT n fun mk_int n = HOLogic.mk_number intT n val nat0 = mk_nat 0 val cnat0 = @{cterm "0::nat"} local val ctxt = @{context} in fun mk_less (m, n) = Const (@{const_name less}, natT --> natT --> boolT) $ m $ n fun mk_le (m, n) = Const (@{const_name less_eq}, natT --> natT --> boolT) $ m $ n (* Obtain the theorem m <= n. *) fun nat_le_th m n = if m > n then raise Fail "nat_le_th: input" else UtilArith.prove_by_arith ctxt [] (mk_le (mk_nat m, mk_nat n)) (* Obtain the theorem m < n. *) fun nat_less_th m n = if m >= n then raise Fail "nat_less_th: input" else UtilArith.prove_by_arith ctxt [] (mk_less (mk_nat m, mk_nat n)) (* Obtain the theorem m ~= n. *) fun nat_neq_th m n = if m = n orelse m < 0 orelse n < 0 then raise Fail "nat_neq_th: input" else UtilArith.prove_by_arith ctxt [] (mk_not (mk_eq (mk_nat m, mk_nat n))) fun nat_fold_reduce t = if fastype_of t <> natT then t else let val (f, (n1, n2)) = t |> Util.dest_binop |> apsnd (apply2 UtilArith.dest_numc) in case f of Const (@{const_name plus}, _) => mk_nat (n1 + n2) | Const (@{const_name minus}, _) => mk_nat (Int.max (0, n1 - n2)) | Const (@{const_name times}, _) => mk_nat (n1 * n2) | _ => t end handle Fail "dest_binop" => t | Fail "dest_numc" => t fun nat_fold_conv ct = let val t = Thm.term_of ct val t' = nat_fold_reduce t in if t aconv t' then Conv.all_conv ct else to_meta_eq (UtilArith.prove_by_arith ctxt [] (mk_eq (t, t'))) end end (* local ctxt = @{context}. *) val plus_ac = {cfhead = @{cterm plus}, unit = SOME @{cterm 0}, assoc_th = @{thm add_ac(1)}, comm_th = @{thm add_ac(2)}, unitl_th = @{thm add_0}, unitr_th = @{thm add_0_right}} val times_ac = {cfhead = @{cterm times}, unit = SOME @{cterm 1}, assoc_th = @{thm mult_ac(1)}, comm_th = @{thm mult_ac(2)}, unitl_th = @{thm mult_1}, unitr_th = @{thm mult_1_right}} val gcd_ac = {cfhead = @{cterm gcd}, unit = SOME @{cterm 0}, assoc_th = @{thm gcd.assoc}, comm_th = @{thm gcd.commute}, unitl_th = @{thm gcd_0_left_nat}, unitr_th = @{thm gcd_0_nat}} val add_arith_ac_data = - fold ACUtil.add_ac_data [plus_ac, times_ac, gcd_ac] + fold Auto2_HOL_Extra_Setup.ACUtil.add_ac_data [plus_ac, times_ac, gcd_ac] fun plus_ac_on_typ thy T = - the (ACUtil.inst_ac_info thy T plus_ac) + the (Auto2_HOL_Extra_Setup.ACUtil.inst_ac_info thy T plus_ac) handle Option.Option => raise Fail "plus_ac_on_typ: cannot inst ac_info." fun times_ac_on_typ thy T = - the (ACUtil.inst_ac_info thy T times_ac) + the (Auto2_HOL_Extra_Setup.ACUtil.inst_ac_info thy T times_ac) handle Option.Option => raise Fail "times_ac_on_typ: cannot inst ac_info." val add_arith_proofsteps = fold add_prfstep_custom [ (* Resolve equality facts with constants. *) ("compare_consts", [WithFact @{term_pat "(?NUMC1::nat) = ?NUMC2"}, Filter (fn _ => fn (_, inst) => lookup_numc1 inst <> lookup_numc2 inst)], fn ((id, _), ths) => fn _ => fn ctxt => - [Update.thm_update (id, UtilArith.contra_by_arith ctxt ths)]), + [Auto2_Setup.Update.thm_update (id, UtilArith.contra_by_arith ctxt ths)]), ("compare_consts_le", [WithFact @{term_pat "(?NUMC1::nat) <= ?NUMC2"}, Filter (fn _ => fn (_, inst) => lookup_numc1 inst > lookup_numc2 inst)], fn ((id, _), ths) => fn _ => fn ctxt => - [Update.thm_update (id, UtilArith.contra_by_arith ctxt ths)]), + [Auto2_Setup.Update.thm_update (id, UtilArith.contra_by_arith ctxt ths)]), ("compare_consts_less", [WithFact @{term_pat "(?NUMC1::nat) < ?NUMC2"}, Filter (fn _ => fn (_, inst) => lookup_numc1 inst >= lookup_numc2 inst)], fn ((id, _), ths) => fn _ => fn ctxt => - [Update.thm_update (id, UtilArith.contra_by_arith ctxt ths)]) + [Auto2_Setup.Update.thm_update (id, UtilArith.contra_by_arith ctxt ths)]) ] #> fold add_prfstep_conv [ ("eval_plus_consts", [WithTerm @{term_pat "(?NUMC1::nat) + ?NUMC2"}, Filter (fn _ => fn (_, inst) => lookup_numc1 inst > 0 andalso lookup_numc2 inst > 0)], nat_fold_conv), ("eval_mult_consts", [WithTerm @{term_pat "(?NUMC1::nat) * ?NUMC2"}, Filter (fn _ => fn (_, inst) => lookup_numc1 inst <> 1 andalso lookup_numc2 inst <> 1)], nat_fold_conv), ("eval_minus_consts", [WithTerm @{term_pat "(?NUMC1::nat) - ?NUMC2"}, Filter (fn _ => fn (_, inst) => lookup_numc2 inst >= 1)], nat_fold_conv)] end (* structure Nat_Util. *) val mk_nat = Nat_Util.mk_nat val mk_int = Nat_Util.mk_int val plus_ac_on_typ = Nat_Util.plus_ac_on_typ val times_ac_on_typ = Nat_Util.times_ac_on_typ val _ = Theory.setup Nat_Util.add_arith_ac_data val _ = Theory.setup Nat_Util.add_arith_proofsteps diff --git a/thys/Auto2_HOL/HOL/auto2_hol.ML b/thys/Auto2_HOL/HOL/auto2_hol.ML deleted file mode 100644 --- a/thys/Auto2_HOL/HOL/auto2_hol.ML +++ /dev/null @@ -1,132 +0,0 @@ -(* - File: auto2_hol.ML - Author: Bohua Zhan - - Setup of auto2 for HOL. -*) - -structure UtilBase : UTIL_BASE = -struct - -(* Types *) - -val boolT = @{typ bool} -val mk_setT = HOLogic.mk_setT - -(* Equality *) - -fun dest_eq t = - case t of - Const (@{const_name HOL.eq}, _) $ lhs $ rhs => (lhs, rhs) - | _ => raise Fail "dest_eq" - -fun cdest_eq ct = - case Thm.term_of ct of - Const (@{const_name HOL.eq}, _) $ _ $ _ => (Thm.dest_arg1 ct, Thm.dest_arg ct) - | _ => raise Fail "dest_eq" - -fun mk_eq (t, u) = - let - val T = fastype_of t - in - Const (@{const_name HOL.eq}, T --> T --> boolT) $ t $ u - end - -fun is_eq_term t = - let - val _ = assert (fastype_of t = boolT) "is_eq_term: wrong type" - in - case t of Const (@{const_name HOL.eq}, _) $ _ $ _ => true - | _ => false - end - -(* Terms *) - -val bTrue = @{term True} -val bFalse = @{term False} -val Trueprop_name = @{const_name HOL.Trueprop} -val Not_name = @{const_name HOL.Not} -val Conj_name = @{const_name HOL.conj} -val Disj_name = @{const_name HOL.disj} -val Imp_name = @{const_name HOL.implies} -val All_name = @{const_name HOL.All} -val Ex_name = @{const_name HOL.Ex} - -(* If expressions are treated differently. In a term "if a then b else - c", only terms in "a" are considered in the proof state. - *) -fun is_if t = - case t of - Const (@{const_name If}, _) $ _ $ _ $ _ => true - | _ => false - -val cTrueprop = @{cterm Trueprop} -val cNot = @{cterm Not} -val cConj = @{cterm conj} -val cDisj = @{cterm disj} - -(* Theorems for equality *) -val to_meta_eq_cv = Conv.rewr_conv @{thm to_meta_eq} -val to_obj_eq_cv = Conv.rewr_conv @{thm atomize_eq} -val to_obj_eq_iff = apply_to_thm (Util.concl_conv to_obj_eq_cv) -val obj_sym_cv = Conv.rewr_conv @{thm obj_sym} - -(* Theorems *) -val true_th = @{thm TrueI} -val iffD_th = @{thm iffD} -val nn_create_th = @{thm nn_create} -val nn_cancel_th = @{thm HOL.nnf_simps(6)} -val to_contra_form_th = @{thm to_contra_form} -val to_contra_form_th' = @{thm to_contra_form'} -val atomize_imp_th = @{thm atomize_imp} -val atomize_all_th = @{thm atomize_all} -val conjunct1_th = @{thm conjunct1} -val conjunct2_th = @{thm conjunct2} -val conjI_th = @{thm conjI} -val or_intro1_th = @{thm or_intro1} -val or_intro2_th = @{thm or_intro2} -val iffD1_th = @{thm iffD1} -val iffD2_th = @{thm iffD2} -val inv_back_th = @{thm inv_backward} -val sym_th = @{thm sym} -val exE_th' = @{thm exE'} -val eq_True_th = @{thm HOL.eqTrueI} -val eq_True_inv_th = @{thm HOL.eqTrueE} -val disj_True1_th = @{thm HOL.simp_thms(30)} -val disj_True2_th = @{thm HOL.simp_thms(29)} -val ex_vardef_th = @{thm HOL.simp_thms(37)} -val imp_conv_disj_th = @{thm imp_conv_disj} -val de_Morgan_conj_th = @{thm de_Morgan_conj} -val de_Morgan_disj_th = @{thm de_Morgan_disj} -val not_ex_th = @{thm HOL.not_ex} -val not_all_th = @{thm HOL.not_all} -val not_imp_th = @{thm HOL.not_imp} -val or_cancel1_th = @{thm or_cancel1} -val or_cancel2_th = @{thm or_cancel2} -val swap_all_disj_th = @{thm swap_all_disj} -val swap_ex_conj_th = @{thm swap_ex_conj} -val all_trivial_th = @{thm HOL.simp_thms(35)} -val case_split_th = @{thm HOL.case_split} - -val atomize_conjL_th = @{thm HOL_Base.atomize_conjL} -val backward_conv_th = @{thm backward_conv} -val backward1_conv_th = @{thm backward1_conv} -val backward2_conv_th = @{thm backward2_conv} -val resolve_conv_th = @{thm resolve_conv} -val contra_triv_th = @{thm contra_triv} - -val conj_assoc_th = @{thm conj_assoc} -val conj_commute_th = @{thm conj_commute} -val disj_assoc_th = @{thm disj_assoc} -val disj_commute_th = @{thm disj_commute} - -val Mem_name = "Set.member" -val Ball_name = "Set.Ball" -val Bex_name = "Set.Bex" -val Bex_def_th = @{thm Bex_def'} -val Ball_def_th = @{thm Ball_def'} - -end (* structure Base *) - -structure Basic_UtilBase: BASIC_UTIL_BASE = UtilBase -open Basic_UtilBase diff --git a/thys/Auto2_HOL/HOL/auto2_hol_extra_setup.ML b/thys/Auto2_HOL/HOL/auto2_hol_extra_setup.ML new file mode 100644 --- /dev/null +++ b/thys/Auto2_HOL/HOL/auto2_hol_extra_setup.ML @@ -0,0 +1,58 @@ +(* + File: auto2_hol_extra_setup.ML + Author: Kevin Kappelmann + + Extra setup for auto2 in HOL. +*) + +signature AUTO2_HOL_EXTRA_SETUP = +sig + structure Unfolding_Keyword: UNFOLDING_KEYWORD + structure Induct_ProofSteps_Keywords: INDUCT_PROOFSTEPS_KEYWORDS + structure ACUtil: ACUTIL + structure AC_ProofSteps: AC_PROOFSTEPS + structure Unfolding: UNFOLDING + structure Induct_ProofSteps: INDUCT_PROOFSTEPS + structure Extra_HOL: EXTRA_HOL +end + +functor Auto2_HOL_Extra_Setup( + structure Auto2_Setup: AUTO2_SETUP; + structure Unfolding_Keyword: UNFOLDING_KEYWORD; + structure Induct_ProofSteps_Keywords: INDUCT_PROOFSTEPS_KEYWORDS; + ): AUTO2_HOL_EXTRA_SETUP = +struct + +structure Unfolding_Keyword = Unfolding_Keyword +structure Induct_ProofSteps_Keywords = Induct_ProofSteps_Keywords +structure ACUtil = ACUtil( + structure Basic_UtilBase = Auto2_Setup.UtilBase; + structure Basic_UtilLogic = Auto2_Setup.UtilLogic; +) +structure AC_ProofSteps = AC_ProofSteps( + structure ACUtil = ACUtil; + structure Basic_UtilLogic = Auto2_Setup.UtilLogic; + structure ProofStepData = Auto2_Setup.ProofStepData; + structure RewriteTable = Auto2_Setup.RewriteTable; + structure Update = Auto2_Setup.Update; +) +val _ = Theory.setup AC_ProofSteps.add_ac_proofsteps +structure Unfolding = Unfolding( + structure Auto2_Outer = Auto2_Setup.Auto2_Outer; + structure Basic_UtilLogic = Auto2_Setup.UtilLogic; + structure Unfolding_Keyword = Unfolding_Keyword; +) +structure Induct_ProofSteps = Induct_ProofSteps( + structure Auto2_Outer = Auto2_Setup.Auto2_Outer; + structure Induct_ProofSteps_Keywords = Induct_ProofSteps_Keywords; + structure UtilBase = Auto2_Setup.UtilBase; + structure UtilLogic = Auto2_Setup.UtilLogic; +) +structure Extra_HOL = Extra_HOL( + structure Basic_UtilBase = Auto2_Setup.UtilBase; + structure Basic_UtilLogic = Auto2_Setup.UtilLogic; + structure ProofStep = Auto2_Setup.ProofStep; + structure ProofStepData = Auto2_Setup.ProofStepData; +) + +end \ No newline at end of file diff --git a/thys/Auto2_HOL/HOL/auto2_hol_util_base.ML b/thys/Auto2_HOL/HOL/auto2_hol_util_base.ML new file mode 100644 --- /dev/null +++ b/thys/Auto2_HOL/HOL/auto2_hol_util_base.ML @@ -0,0 +1,131 @@ +(* + File: auto2_hol_util_base.ML + Author: Bohua Zhan + + Setup of auto2 utility for HOL. +*) + +structure Auto2_UtilBase : UTIL_BASE = +struct + +(* Types *) + +val boolT = @{typ bool} +val mk_setT = HOLogic.mk_setT + +(* Equality *) + +fun dest_eq t = + case t of + Const (@{const_name HOL.eq}, _) $ lhs $ rhs => (lhs, rhs) + | _ => raise Fail "dest_eq" + +fun cdest_eq ct = + case Thm.term_of ct of + Const (@{const_name HOL.eq}, _) $ _ $ _ => (Thm.dest_arg1 ct, Thm.dest_arg ct) + | _ => raise Fail "dest_eq" + +fun mk_eq (t, u) = + let + val T = fastype_of t + in + Const (@{const_name HOL.eq}, T --> T --> boolT) $ t $ u + end + +fun is_eq_term t = + let + val _ = assert (fastype_of t = boolT) "is_eq_term: wrong type" + in + case t of Const (@{const_name HOL.eq}, _) $ _ $ _ => true + | _ => false + end + +(* Terms *) + +val bTrue = @{term True} +val bFalse = @{term False} +val Trueprop_name = @{const_name HOL.Trueprop} +val Not_name = @{const_name HOL.Not} +val Conj_name = @{const_name HOL.conj} +val Disj_name = @{const_name HOL.disj} +val Imp_name = @{const_name HOL.implies} +val All_name = @{const_name HOL.All} +val Ex_name = @{const_name HOL.Ex} + +(* If expressions are treated differently. In a term "if a then b else + c", only terms in "a" are considered in the proof state. + *) +fun is_if t = + case t of + Const (@{const_name If}, _) $ _ $ _ $ _ => true + | _ => false + +val cTrueprop = @{cterm Trueprop} +val cNot = @{cterm Not} +val cConj = @{cterm conj} +val cDisj = @{cterm disj} + +(* Theorems for equality *) +val to_meta_eq_cv = Conv.rewr_conv @{thm to_meta_eq} +val to_obj_eq_cv = Conv.rewr_conv @{thm atomize_eq} +val to_obj_eq_iff = apply_to_thm (Util.concl_conv to_obj_eq_cv) +val obj_sym_cv = Conv.rewr_conv @{thm obj_sym} + +(* Theorems *) +val true_th = @{thm TrueI} +val iffD_th = @{thm iffD} +val nn_create_th = @{thm nn_create} +val nn_cancel_th = @{thm HOL.nnf_simps(6)} +val to_contra_form_th = @{thm to_contra_form} +val to_contra_form_th' = @{thm to_contra_form'} +val atomize_imp_th = @{thm atomize_imp} +val atomize_all_th = @{thm atomize_all} +val conjunct1_th = @{thm conjunct1} +val conjunct2_th = @{thm conjunct2} +val conjI_th = @{thm conjI} +val or_intro1_th = @{thm or_intro1} +val or_intro2_th = @{thm or_intro2} +val iffD1_th = @{thm iffD1} +val iffD2_th = @{thm iffD2} +val inv_back_th = @{thm inv_backward} +val sym_th = @{thm sym} +val exE_th' = @{thm exE'} +val eq_True_th = @{thm HOL.eqTrueI} +val eq_True_inv_th = @{thm HOL.eqTrueE} +val disj_True1_th = @{thm HOL.simp_thms(30)} +val disj_True2_th = @{thm HOL.simp_thms(29)} +val ex_vardef_th = @{thm HOL.simp_thms(37)} +val imp_conv_disj_th = @{thm imp_conv_disj} +val de_Morgan_conj_th = @{thm de_Morgan_conj} +val de_Morgan_disj_th = @{thm de_Morgan_disj} +val not_ex_th = @{thm HOL.not_ex} +val not_all_th = @{thm HOL.not_all} +val not_imp_th = @{thm HOL.not_imp} +val or_cancel1_th = @{thm or_cancel1} +val or_cancel2_th = @{thm or_cancel2} +val swap_all_disj_th = @{thm swap_all_disj} +val swap_ex_conj_th = @{thm swap_ex_conj} +val all_trivial_th = @{thm HOL.simp_thms(35)} +val case_split_th = @{thm HOL.case_split} + +val atomize_conjL_th = @{thm HOL_Base.atomize_conjL} +val backward_conv_th = @{thm backward_conv} +val backward1_conv_th = @{thm backward1_conv} +val backward2_conv_th = @{thm backward2_conv} +val resolve_conv_th = @{thm resolve_conv} +val contra_triv_th = @{thm contra_triv} + +val conj_assoc_th = @{thm conj_assoc} +val conj_commute_th = @{thm conj_commute} +val disj_assoc_th = @{thm disj_assoc} +val disj_commute_th = @{thm disj_commute} + +val Mem_name = "Set.member" +val Ball_name = "Set.Ball" +val Bex_name = "Set.Bex" +val Bex_def_th = @{thm Bex_def'} +val Ball_def_th = @{thm Ball_def'} + +end (* structure Base *) + +structure Auto2_Basic_UtilBase : BASIC_UTIL_BASE = Auto2_UtilBase diff --git a/thys/Auto2_HOL/HOL/extra_hol.ML b/thys/Auto2_HOL/HOL/extra_hol.ML --- a/thys/Auto2_HOL/HOL/extra_hol.ML +++ b/thys/Auto2_HOL/HOL/extra_hol.ML @@ -1,91 +1,94 @@ (* File: extra_hol.ML Author: Bohua Zhan Extra setup for HOL. *) signature EXTRA_HOL = sig val add_forward_arg1_prfstep_cond: thm -> pre_prfstep_descriptor list -> theory -> theory val add_forward_arg1_prfstep: thm -> theory -> theory val add_forward_arg_prfstep_cond: thm -> pre_prfstep_descriptor list -> theory -> theory val add_forward_arg_prfstep: thm -> theory -> theory val add_rewrite_arg_rule_cond: thm -> pre_prfstep_descriptor list -> theory -> theory val add_rewrite_arg_rule: thm -> theory -> theory val add_simple_datatype: string -> theory -> theory val del_simple_datatype: string -> theory -> theory end; -structure Extra_HOL : EXTRA_HOL = +functor Extra_HOL( + structure Basic_UtilBase: BASIC_UTIL_BASE; + structure Basic_UtilLogic: BASIC_UTIL_LOGIC; + structure ProofStep: PROOFSTEP; + structure ProofStepData: PROOFSTEP_DATA; + ): EXTRA_HOL = struct fun add_forward_arg1_prfstep_cond th conds thy = let - val concl = th |> concl_of' |> strip_conj |> hd + val concl = th |> Basic_UtilLogic.concl_of' |> Basic_UtilLogic.strip_conj |> hd in - thy |> add_forward_prfstep_cond - th ([K (WithTerm (dest_arg1 concl))] @ conds) + thy |> ProofStepData.add_forward_prfstep_cond + th ([K (ProofStep.WithTerm (dest_arg1 concl))] @ conds) end fun add_forward_arg1_prfstep th = add_forward_arg1_prfstep_cond th [] fun add_forward_arg_prfstep_cond th conds thy = let - val concl = th |> concl_of' |> strip_conj |> hd + val concl = th |> Basic_UtilLogic.concl_of' |> Basic_UtilLogic.strip_conj |> hd in - thy |> add_forward_prfstep_cond - th ([K (WithTerm (dest_arg concl))] @ conds) + thy |> ProofStepData.add_forward_prfstep_cond + th ([K (ProofStep.WithTerm (dest_arg concl))] @ conds) end fun add_forward_arg_prfstep th = add_forward_arg_prfstep_cond th [] fun add_rewrite_arg_rule_cond th conds thy = let - val concl = th |> concl_of' |> strip_conj |> hd - val _ = assert (is_eq_term concl) "rewrite_arg" - val (lhs, _) = dest_eq concl + val concl = th |> Basic_UtilLogic.concl_of' |> Basic_UtilLogic.strip_conj |> hd + val _ = assert (Basic_UtilBase.is_eq_term concl) "rewrite_arg" + val (lhs, _) = Basic_UtilBase.dest_eq concl in - thy |> add_forward_prfstep_cond - th ([K (WithTerm (dest_arg lhs))] @ conds) + thy |> ProofStepData.add_forward_prfstep_cond + th ([K (ProofStep.WithTerm (dest_arg lhs))] @ conds) end fun add_rewrite_arg_rule th = add_rewrite_arg_rule_cond th [] fun add_simple_datatype s thy = let val collapse_th = Global_Theory.get_thm thy (s ^ ".collapse") val case_th = Global_Theory.get_thm thy (s ^ ".case") val sel_th = Global_Theory.get_thms thy (s ^ ".sel") val simp_th = hd (Global_Theory.get_thms thy (s ^ ".simps")) - val var = collapse_th |> prop_of' |> dest_arg - val (f, args) = collapse_th |> prop_of' |> dest_arg1 |> Term.strip_comb + val var = collapse_th |> Basic_UtilLogic.prop_of' |> dest_arg + val (f, args) = collapse_th |> Basic_UtilLogic.prop_of' |> dest_arg1 |> Term.strip_comb val vars = map (fn (n, T) => Var (("x",n),T)) (tag_list 1 (map fastype_of args)) val rhs = Term.list_comb (f, vars) - val neq = get_neg (mk_eq (var, rhs)) - val filt = [with_filt (neq_filter neq)] + val neq = Basic_UtilLogic.get_neg (Basic_UtilBase.mk_eq (var, rhs)) + val filt = [ProofStepData.with_filt (ProofStep.neq_filter neq)] in - thy |> add_rewrite_rule_back_cond collapse_th filt - |> add_rewrite_rule case_th - |> fold add_rewrite_rule sel_th - |> add_forward_prfstep (equiv_forward_th simp_th) + thy |> ProofStepData.add_rewrite_rule_back_cond collapse_th filt + |> ProofStepData.add_rewrite_rule case_th + |> fold ProofStepData.add_rewrite_rule sel_th + |> ProofStepData.add_forward_prfstep (Basic_UtilLogic.equiv_forward_th simp_th) end fun del_simple_datatype s thy = let val collapse_th = Global_Theory.get_thm thy (s ^ ".collapse") val case_th = Global_Theory.get_thm thy (s ^ ".case") val sel_th = Global_Theory.get_thms thy (s ^ ".sel") val simp_th = hd (Global_Theory.get_thms thy (s ^ ".simps")) in - thy |> fold del_prfstep_thm (collapse_th :: case_th :: simp_th :: sel_th) + thy |> fold ProofStepData.del_prfstep_thm (collapse_th :: case_th :: simp_th :: sel_th) end end (* structure Extra_HOL *) - -open Extra_HOL diff --git a/thys/Auto2_HOL/HOL/induct_outer.ML b/thys/Auto2_HOL/HOL/induct_outer.ML --- a/thys/Auto2_HOL/HOL/induct_outer.ML +++ b/thys/Auto2_HOL/HOL/induct_outer.ML @@ -1,538 +1,549 @@ (* File: induct_outer.ML Author: Bohua Zhan Proof language for induction. *) signature INDUCT_PROOFSTEPS = sig val add_induct_data: string -> term * thm -> theory -> theory val add_typed_induct_data: string -> typ * thm -> theory -> theory val get_typed_ind_th: theory -> string -> typ -> thm val get_term_ind_th: theory -> string -> term -> thm val check_strong_ind_prop: term -> term list * term val add_strong_induct_rule: thm -> theory -> theory val add_case_induct_rule: thm -> theory -> theory val add_prop_induct_rule: thm -> theory -> theory val add_var_induct_rule: thm -> theory -> theory val add_cases_rule: thm -> theory -> theory val add_fun_induct_rule: term * thm -> theory -> theory val strong_induct_cmd: string * string list -> Proof.state -> Proof.state val apply_induct_hyp_cmd: string list -> Proof.state -> Proof.state val case_induct_cmd: string -> Proof.state -> Proof.state val prop_induct_cmd: string * string option -> Proof.state -> Proof.state val induct_cmd: string -> string * string option * string list * string option -> Proof.state -> Proof.state val is_simple_fun_induct: thm -> bool val fun_induct_cmd: string * string list * string option -> Proof.state -> Proof.state end; -structure Induct_ProofSteps : INDUCT_PROOFSTEPS = +signature INDUCT_PROOFSTEPS_KEYWORDS = +sig + val apply_induct_hyp: string * Position.T + val case_induct: string * Position.T + val cases: string * Position.T + val fun_induct: string * Position.T + val induct: string * Position.T + val prop_induct: string * Position.T + val strong_induct: string * Position.T + val arbitrary: string parser + val with': string parser +end; + +functor Induct_ProofSteps( + structure Auto2_Outer: AUTO2_OUTER; + structure Induct_ProofSteps_Keywords: INDUCT_PROOFSTEPS_KEYWORDS; + structure UtilBase: UTIL_BASE; + structure UtilLogic: UTIL_LOGIC; + ) : INDUCT_PROOFSTEPS = struct structure Data = Theory_Data ( type T = ((term * thm) list) Symtab.table; val empty = Symtab.empty; val merge = Symtab.merge_list (eq_fst (op =)) ) fun add_induct_data str (t, ind_th) = Data.map (Symtab.map_default (str, []) (cons (t, ind_th))) fun add_typed_induct_data str (ty, ind_th) = add_induct_data str (Term.dummy_pattern ty, ind_th) fun get_typed_ind_th thy ind_type ty = let val typ_can_match = can (fn t' => Sign.typ_match thy (type_of t', ty) Vartab.empty) in case Symtab.lookup (Data.get thy) ind_type of NONE => raise Fail (ind_type ^ ": cannot find theorem.") | SOME lst => case find_first (fn (t', _) => typ_can_match t') lst of NONE => raise Fail (ind_type ^ ": cannot find theorem.") | SOME (_, ind_th) => ind_th end fun get_term_ind_th thy ind_type t = let val data = Symtab.lookup_list (Data.get thy) ind_type fun match_data (pat, th) = let val inst = Pattern.first_order_match thy (pat, t) fo_init in SOME (Util.subst_thm_thy thy inst th) end handle Pattern.MATCH => NONE in case get_first match_data data of NONE => raise Fail (ind_type ^ ": cannot find theorem.") | SOME ind_th => ind_th end (* Check a strong induction theorem ind_th is of the right form, and extract the induction variables and substitution. *) fun check_strong_ind_prop ind_prop = let fun err str = "Strong induction: " ^ str val (cond_ind, concl) = - ind_prop |> Logic.dest_implies |> apply2 dest_Trueprop + ind_prop |> Logic.dest_implies |> apply2 UtilLogic.dest_Trueprop (* concl must be of form ?P [?vars]. *) val err_concl = err "concl of ind_th must be ?P [?vars]." val (P, pat_vars) = Term.strip_comb concl handle TERM _ => error err_concl val _ = assert (is_Var P andalso forall is_Var pat_vars andalso (dest_Var P |> fst |> fst) = "P") err_concl (* cond_ind must be of form !n. P' n --> ?P n. Return the substitution pattern P'. *) val err_ind_hyp = err "cond_ind of ind_th must be !n. P' --> ?P vars." fun dest_one_all var body = case body of Const (c, _) $ Abs (_, _, t) => if c = UtilBase.All_name then subst_bound (var, t) else error err_ind_hyp | _ => error err_ind_hyp val (pat_subst, P_vars) = - cond_ind |> fold dest_one_all pat_vars |> dest_imp + cond_ind |> fold dest_one_all pat_vars |> UtilLogic.dest_imp val _ = assert (P_vars aconv concl) err_ind_hyp in (pat_vars, pat_subst) end fun add_strong_induct_rule ind_th thy = let val name = Util.name_of_thm ind_th val ctxt = Proof_Context.init_global thy val ind_th' = apply_to_thm (UtilLogic.to_obj_conv_on_horn ctxt) ind_th val (pat_var, pat_subst) = check_strong_ind_prop (Thm.prop_of ind_th') |> apfst the_single handle List.Empty => error "Strong induction: more than one var." val ty_var = type_of pat_var val _ = writeln (name ^ "\nSubstitution: " ^ (Util.string_of_terms_global thy [pat_var, pat_subst])) in thy |> add_typed_induct_data "strong_induct" (ty_var, ind_th') end fun add_case_induct_rule ind_th thy = let - val init_assum = ind_th |> Thm.prems_of |> hd |> dest_Trueprop + val init_assum = ind_th |> Thm.prems_of |> hd |> UtilLogic.dest_Trueprop in thy |> add_induct_data "case_induct" (init_assum, ind_th) end fun add_prop_induct_rule ind_th thy = let - val init_assum = ind_th |> Thm.prems_of |> hd |> dest_Trueprop + val init_assum = ind_th |> Thm.prems_of |> hd |> UtilLogic.dest_Trueprop in thy |> add_induct_data "prop_induct" (init_assum, ind_th) end fun add_var_induct_rule ind_th thy = let - val (P, n) = ind_th |> concl_of' |> Term.dest_comb + val (P, n) = ind_th |> UtilLogic.concl_of' |> Term.dest_comb val _ = assert (Term.is_Var P andalso Term.is_Var n) "add_var_induct_rule: concl of ind_th must be ?P ?var" in thy |> add_typed_induct_data "var_induct" (type_of n, ind_th) end fun add_cases_rule ind_th thy = let - val (P, n) = ind_th |> concl_of' |> Term.dest_comb + val (P, n) = ind_th |> UtilLogic.concl_of' |> Term.dest_comb val _ = assert (Term.is_Var P andalso Term.is_Var n) "add_cases_rule: concl of ind_th must be ?P ?var" in thy |> add_typed_induct_data "cases" (type_of n, ind_th) end fun add_fun_induct_rule (t, ind_th) thy = thy |> add_induct_data "fun_induct" (t, ind_th) (* Obtain the induction statement. *) fun get_induct_stmt ctxt (filt_A, ind_vars, stmt, arbitrary) = case stmt of NONE => let val (_, (As, C)) = ctxt |> Auto2_State.get_subgoal |> Util.strip_meta_horn - val obj_As = As |> map dest_Trueprop |> filter filt_A - val obj_C = dest_Trueprop C + val obj_As = As |> map UtilLogic.dest_Trueprop |> filter filt_A + val obj_C = UtilLogic.dest_Trueprop C in (UtilLogic.list_obj_horn (arbitrary, (obj_As, obj_C))) |> fold Util.lambda_abstract (rev ind_vars) end | SOME s => (UtilLogic.list_obj_horn (arbitrary, ([], Syntax.read_term ctxt s))) |> fold Util.lambda_abstract (rev ind_vars) fun apply_simple_induct_th ind_th vars arbitraries prem_only state = let val {context = ctxt, ...} = Proof.goal state val prop = Auto2_State.get_selected ctxt val (vars', _) = prop |> Thm.prems_of |> the_single |> Util.strip_meta_horn val ind_th = ind_th |> apply_to_thm (Conv.binop_conv (UtilLogic.to_meta_conv ctxt)) val assum = hd (Drule.cprems_of ind_th) val ind_th = ind_th |> Util.send_first_to_hyps |> fold Thm.forall_elim (map (Thm.cterm_of ctxt) arbitraries) |> fold Thm.forall_intr (map (Thm.cterm_of ctxt) vars') |> Thm.implies_intr assum val t' = case Thm.prop_of ind_th of imp $ A $ B => imp $ Util.rename_abs_term vars A $ B | _ => raise Fail "strong_induct_cmd" val ind_th = ind_th |> Thm.renamed_prop t' val prop = prop |> Auto2_Outer.refine_subgoal_th ind_th in if prem_only then let val (_, (As, _)) = prop |> Thm.prems_of |> the_single |> Util.strip_meta_horn - val stmt = dest_Trueprop (hd As) + val stmt = UtilLogic.dest_Trueprop (hd As) in state |> Proof.map_contexts (Auto2_State.map_head_th (K prop)) |> Proof.map_contexts (Auto2_State.set_induct_stmt stmt) |> Proof.map_contexts (Auto2_State.add_prem_only stmt) end else state |> Proof.map_contexts (Auto2_State.map_head_th (K prop)) end fun strong_induct_cmd (s, t) state = let val {context = ctxt, ...} = Proof.goal state val thy = Proof_Context.theory_of ctxt val var = Syntax.read_term ctxt s val arbitraries = map (Syntax.read_term ctxt) t val P = get_induct_stmt ctxt (K true, [var], NONE, arbitraries) val ind_th = get_typed_ind_th thy "strong_induct" (type_of var) - val (var_P, var_n) = ind_th |> concl_of' |> Term.dest_comb + val (var_P, var_n) = ind_th |> UtilLogic.concl_of' |> Term.dest_comb val inst = fold (Pattern.match thy) [(var_P, P), (var_n, var)] fo_init val ind_th = Util.subst_thm ctxt inst ind_th in state |> apply_simple_induct_th ind_th [var] arbitraries true end val arbitrary = - Scan.option (@{keyword "arbitrary"} |-- Scan.repeat Parse.term) + Scan.option (Induct_ProofSteps_Keywords.arbitrary |-- Scan.repeat Parse.term) val _ = - Outer_Syntax.command @{command_keyword "@strong_induct"} + Outer_Syntax.command Induct_ProofSteps_Keywords.strong_induct "apply strong induction" ((Parse.term -- arbitrary) >> (fn (s, t) => Toplevel.proof (fn state => strong_induct_cmd (s, these t) state))) fun apply_induct_hyp_cmd s state = let val {context = ctxt, ...} = Proof.goal state val ts = Syntax.read_terms ctxt s val induct_stmt = Auto2_State.get_last_induct_stmt ctxt - val stmt = induct_stmt |> the |> mk_Trueprop |> Thm.cterm_of ctxt + val stmt = induct_stmt |> the |> UtilLogic.mk_Trueprop |> Thm.cterm_of ctxt handle Option.Option => raise Fail "apply_induct_hyp: no induct_stmt" val prop = Auto2_State.get_selected ctxt val (_, (As, _)) = prop |> Thm.prems_of |> the_single |> Util.strip_meta_horn val _ = assert (member (op aconv) As (Thm.term_of stmt)) "apply_induct_hyp: induct_stmt not found among As." val cAs = map (Thm.cterm_of ctxt) As val th = stmt |> Thm.assume |> apply_to_thm (UtilLogic.to_meta_conv ctxt) |> fold Thm.forall_elim (map (Thm.cterm_of ctxt) ts) |> apply_to_thm (Util.normalize_meta_all_imp ctxt) val prems = th |> Thm.prems_of |> map (fn t => Logic.list_implies (As, t)) |> map (Thm.cterm_of ctxt) val prems_th = (map (Auto2_Outer.auto2_solve ctxt) prems) |> map Util.send_all_to_hyps val concl = th |> fold Thm.elim_implies prems_th |> fold Thm.implies_intr (rev cAs) val _ = writeln ("Obtained " ^ Syntax.string_of_term ctxt (Thm.concl_of concl)) in state |> Proof.map_contexts ( Auto2_State.map_head_th (Auto2_Outer.have_after_qed ctxt concl)) end val _ = - Outer_Syntax.command @{command_keyword "@apply_induct_hyp"} + Outer_Syntax.command Induct_ProofSteps_Keywords.apply_induct_hyp "apply induction hypothesis" ((Scan.repeat Parse.term) >> (fn s => Toplevel.proof (fn state => apply_induct_hyp_cmd s state))) fun solve_goals ind_th pats_opt filt_As state = let val {context = ctxt, ...} = Proof.goal state val (_, (As, _)) = ctxt |> Auto2_State.get_subgoal |> Util.strip_meta_horn val use_As = filter filt_As As val cAs = map (Thm.cterm_of ctxt) As val ind_goals = ind_th |> Thm.prems_of |> map (fn t => Logic.list_implies (use_As, t)) |> map (Thm.cterm_of ctxt) |> map (UtilLogic.to_meta_conv ctxt) in case pats_opt of NONE => let (* Solve the right side, obtain the left side. *) fun solve_eq eq = Thm.equal_elim (meta_sym eq) (Auto2_Outer.auto2_solve ctxt (Thm.rhs_of eq)) val ths = ind_goals |> map solve_eq |> map Util.send_all_to_hyps val ind_concl = ind_th |> fold Thm.elim_implies ths |> fold Thm.implies_intr (rev cAs) val after_qed = Auto2_Outer.have_after_qed ctxt ind_concl in state |> Proof.map_contexts (Auto2_State.map_head_th after_qed) end | SOME pats => let (* Create new block with the subgoals *) fun after_qed ths prop = let val ths' = (ind_goals ~~ ths) |> map (fn (eq, th) => Thm.equal_elim (meta_sym eq) th) |> map Util.send_all_to_hyps val ind_concl = ind_th |> fold Thm.elim_implies ths' |> fold Thm.implies_intr (rev cAs) in Auto2_Outer.have_after_qed ctxt ind_concl prop end val _ = writeln ("Patterns: " ^ Util.string_of_terms ctxt pats) val new_frame = Auto2_State.multiple_frame ( pats ~~ map Thm.rhs_of ind_goals, SOME ([], after_qed)) in state |> Proof.map_contexts (Auto2_State.push_head new_frame) end end fun case_induct_cmd s state = let val {context = ctxt, ...} = Proof.goal state val thy = Proof_Context.theory_of ctxt val start = Syntax.read_term ctxt s val ind_th = get_term_ind_th thy "case_induct" start (* Obtain list of assumptions *) val (_, (_, C)) = ctxt |> Auto2_State.get_subgoal |> Util.strip_meta_horn (* Instantiate the induction theorem *) - val var_P = concl_of' ind_th - val inst = Pattern.match thy (var_P, dest_Trueprop C) fo_init + val var_P = UtilLogic.concl_of' ind_th + val inst = Pattern.match thy (var_P, UtilLogic.dest_Trueprop C) fo_init val ind_th = Util.subst_thm_thy thy inst ind_th in state |> solve_goals ind_th NONE (K true) end val _ = - Outer_Syntax.command @{command_keyword "@case_induct"} "apply induction" + Outer_Syntax.command Induct_ProofSteps_Keywords.case_induct "apply induction" (Parse.term >> (fn s => Toplevel.proof (fn state => case_induct_cmd s state))) val for_stmt = Scan.option (@{keyword "for"} |-- Parse.term) fun prop_induct_cmd (s, t) state = let val {context = ctxt, ...} = Proof.goal state val thy = Proof_Context.theory_of ctxt val start = Syntax.read_term ctxt s val ind_th = get_term_ind_th thy "prop_induct" start - val (var_P, args) = ind_th |> concl_of' |> Term.strip_comb + val (var_P, args) = ind_th |> UtilLogic.concl_of' |> Term.strip_comb - val start_As = strip_conj start + val start_As = UtilLogic.strip_conj start val filt_A = (fn t => not (member (op aconv) start_As t)) val P = get_induct_stmt ctxt (filt_A, args, t, []) val _ = writeln ("Induct statement: " ^ Syntax.string_of_term ctxt P) val inst = Pattern.match thy (var_P, P) fo_init (* Instantiate the induction theorem *) val ind_th = Util.subst_thm_thy thy inst ind_th in state |> solve_goals ind_th NONE (K true) end val _ = - Outer_Syntax.command @{command_keyword "@prop_induct"} "apply induction" + Outer_Syntax.command Induct_ProofSteps_Keywords.prop_induct "apply induction" ((Parse.term -- for_stmt) >> (fn (s, t) => Toplevel.proof (fn state => prop_induct_cmd (s, t) state))) (* Given an induction subgoal of the form !!x_i. A_i ==> C, retrieve the list of induction patterns. *) fun retrieve_pat ind_vars t = let val (vars, (_, C)) = Util.strip_meta_horn t fun free_to_var t = let val (x, T) = Term.dest_Free t in Var ((x,0), T) end val pat_vars = map free_to_var vars - val args = C |> dest_Trueprop |> Util.dest_args + val args = C |> UtilLogic.dest_Trueprop |> Util.dest_args |> map (Term.subst_atomic (vars ~~ pat_vars)) in - HOLogic.mk_tuple (map mk_eq (ind_vars ~~ args)) + HOLogic.mk_tuple (map UtilBase.mk_eq (ind_vars ~~ args)) end fun induct_cmd ind_ty_str (s, t, u, v) state = let val {context = ctxt, ...} = Proof.goal state val thy = Proof_Context.theory_of ctxt val var = Syntax.read_term ctxt s val arbitraries = map (Syntax.read_term ctxt) u val filt_A = Util.occurs_frees (var :: arbitraries) val P = get_induct_stmt ctxt (filt_A, [var], t, arbitraries) val ind_th = get_typed_ind_th thy ind_ty_str (type_of var) (* Instantiate the induction theorem *) - val concl = concl_of' ind_th + val concl = UtilLogic.concl_of' ind_th val (var_P, var_n) = Term.dest_comb concl val inst = fold (Pattern.match thy) [(var_P, P), (var_n, var)] fo_init val ind_th' = Util.subst_thm_thy thy inst ind_th val pats = case v of NONE => NONE | _ => SOME (map (retrieve_pat [var]) (Thm.prems_of ind_th)) in state |> solve_goals ind_th' pats (not o filt_A) end val _ = - Outer_Syntax.command @{command_keyword "@induct"} "apply induction" - (Parse.term -- for_stmt -- arbitrary -- Scan.option @{keyword "@with"} >> + Outer_Syntax.command Induct_ProofSteps_Keywords.induct "apply induction" + (Parse.term -- for_stmt -- arbitrary -- Scan.option Induct_ProofSteps_Keywords.with' >> (fn (((s, t), u), v) => Toplevel.proof ( fn state => induct_cmd "var_induct" (s, t, these u, v) state))) val _ = - Outer_Syntax.command @{command_keyword "@cases"} "apply induction" - (Parse.term -- Scan.option @{keyword "@with"} >> + Outer_Syntax.command Induct_ProofSteps_Keywords.cases "apply induction" + (Parse.term -- Scan.option Induct_ProofSteps_Keywords.with' >> (fn (s, v) => Toplevel.proof ( fn state => induct_cmd "cases" (s, NONE, [], v) state))) fun get_fun_induct_th thy t = let val ind_th = get_term_ind_th thy "fun_induct" (Term.head_of t) handle Fail _ => Global_Theory.get_thm thy (Util.get_head_name t ^ ".induct") handle ERROR _ => raise Fail "fun_induct: cannot find theorem." val (_, args) = Term.strip_comb t - val (_, pat_args) = ind_th |> concl_of' |> Term.strip_comb + val (_, pat_args) = ind_th |> UtilLogic.concl_of' |> Term.strip_comb val inst = Util.first_order_match_list thy (pat_args ~~ args) fo_init in Util.subst_thm_thy thy inst ind_th end fun is_simple_fun_induct ind_th = let val prems = Thm.prems_of ind_th in if length prems > 1 then false else let val (var, (_, C)) = Util.strip_meta_horn (the_single prems) - val (_, args) = Term.strip_comb (dest_Trueprop C) + val (_, args) = Term.strip_comb (UtilLogic.dest_Trueprop C) in eq_list (op aconv) (var, args) end end fun fun_induct_cmd (s, t, u) state = let val {context = ctxt, ...} = Proof.goal state val thy = Proof_Context.theory_of ctxt val expr = Syntax.read_term ctxt s val arbitraries = map (Syntax.read_term ctxt) t val ind_th = get_fun_induct_th thy expr - val (var_P, vars) = ind_th |> concl_of' |> Term.strip_comb + val (var_P, vars) = ind_th |> UtilLogic.concl_of' |> Term.strip_comb in if is_simple_fun_induct ind_th then let val _ = assert (is_none u) "fun_induct: simple induction." (* Instantiate the induction theorem *) val P = get_induct_stmt ctxt (K true, vars, NONE, arbitraries) val inst = Pattern.match thy (var_P, P) fo_init val ind_th = Util.subst_thm_thy thy inst ind_th in state |> apply_simple_induct_th ind_th vars arbitraries false end else let (* Instantiate the induction theorem *) val filt_A = Util.occurs_frees (vars @ arbitraries) val P = get_induct_stmt ctxt (filt_A, vars, NONE, arbitraries) val inst = Pattern.match thy (var_P, P) fo_init val ind_th' = ind_th |> Util.subst_thm_thy thy inst val prems = Thm.prems_of ind_th val pats = case u of NONE => NONE | SOME _ => SOME (map (retrieve_pat vars) prems) in state |> solve_goals ind_th' pats (not o filt_A) end end val _ = - Outer_Syntax.command @{command_keyword "@fun_induct"} "apply induction" - (Parse.term -- arbitrary -- Scan.option @{keyword "@with"} >> + Outer_Syntax.command Induct_ProofSteps_Keywords.fun_induct "apply induction" + (Parse.term -- arbitrary -- Scan.option Induct_ProofSteps_Keywords.with' >> (fn ((s, t), u) => Toplevel.proof (fn state => fun_induct_cmd (s, these t, u) state))) end (* structure Induct_ProofSteps. *) - -val add_strong_induct_rule = Induct_ProofSteps.add_strong_induct_rule -val add_case_induct_rule = Induct_ProofSteps.add_case_induct_rule -val add_prop_induct_rule = Induct_ProofSteps.add_prop_induct_rule -val add_var_induct_rule = Induct_ProofSteps.add_var_induct_rule -val add_fun_induct_rule = Induct_ProofSteps.add_fun_induct_rule -val add_cases_rule = Induct_ProofSteps.add_cases_rule diff --git a/thys/Auto2_HOL/HOL/list_ac.ML b/thys/Auto2_HOL/HOL/list_ac.ML --- a/thys/Auto2_HOL/HOL/list_ac.ML +++ b/thys/Auto2_HOL/HOL/list_ac.ML @@ -1,286 +1,286 @@ (* File: list_ac.ML Author: Bohua Zhan Special normalization procedure for list append. In addition to normalizing the AC function append, also normalize a # xs to [a] @ xs. *) signature LIST_AC = sig datatype heads = LIST_APPEND | LIST_CONS | LIST_NIL | LIST_OTHER val append_const: typ -> term val case_head: typ -> term -> heads val is_list_head: typ -> term -> bool val get_list_ty: term -> typ option val rewrite_on_eqs: typ -> thm list -> cterm -> thm val dest_list_full: Proof.context -> typ -> cterm -> cterm list val simp_list_expr: Proof.context -> typ -> box_id * cterm -> (box_id * thm) list val get_list_head_equiv: Proof.context -> typ -> box_id * cterm -> (box_id * thm) list val list_expand_once: Proof.context -> typ -> box_id * cterm -> (box_id * thm) list val list_expand: Proof.context -> typ -> box_id * cterm -> (box_id * thm) list val normalize_list_assoc: typ -> conv val normalize_list: typ -> conv val list_expand_equiv: proofstep val list_expand_unit: proofstep val add_list_proofsteps: theory -> theory end; structure List_AC : LIST_AC = struct datatype heads = LIST_APPEND | LIST_CONS | LIST_NIL | LIST_OTHER fun append_const T = let val lT = HOLogic.listT T in Const (@{const_name append}, lT --> lT --> lT) end fun case_head T t = if Util.is_head (append_const T) t then LIST_APPEND else if Util.is_head (HOLogic.cons_const T) t andalso not (dest_arg t aconv HOLogic.nil_const T) then LIST_CONS else if t aconv HOLogic.nil_const T then LIST_NIL else LIST_OTHER fun is_list_head T t = case case_head T t of LIST_APPEND => true | LIST_CONS => true | _ => false fun get_list_ty t = case fastype_of t of Type ("List.list", [T]) => SOME T | _ => NONE fun rewrite_on_eqs T eqs ct = case case_head T (Thm.term_of ct) of LIST_APPEND => Conv.binop_conv (rewrite_on_eqs T eqs) ct | LIST_CONS => Conv.binop_conv (rewrite_on_eqs T eqs) ct | LIST_NIL => Conv.all_conv ct | _ => (case find_first (fn eq => (Thm.term_of ct) aconv (Util.lhs_of eq)) eqs of NONE => Conv.all_conv ct | SOME eq_th => eq_th) fun dest_list_full ctxt T ct = case case_head T (Thm.term_of ct) of LIST_APPEND => dest_list_full ctxt T (Thm.dest_arg1 ct) @ dest_list_full ctxt T (Thm.dest_arg ct) | LIST_CONS => Thm.cterm_of ctxt ( HOLogic.mk_list T [dest_arg1 (Thm.term_of ct)]) :: dest_list_full ctxt T (Thm.dest_arg ct) | LIST_NIL => [] | _ => [ct] fun dest_subs ctxt T ct = case case_head T (Thm.term_of ct) of LIST_APPEND => dest_subs ctxt T (Thm.dest_arg1 ct) @ dest_subs ctxt T (Thm.dest_arg ct) | LIST_CONS => Thm.dest_arg1 ct :: dest_subs ctxt T (Thm.dest_arg ct) | LIST_NIL => [] | _ => [ct] fun simp_list_expr ctxt T (id, cu) = let val cus = dest_subs ctxt T cu in if null cus then [(id, Thm.reflexive cu)] else - cus |> map (RewriteTable.simplify_info ctxt) + cus |> map (Auto2_Setup.RewriteTable.simplify_info ctxt) |> BoxID.get_all_merges_info ctxt |> BoxID.merge_box_with_info ctxt id |> map (fn (id', eqs) => (id', rewrite_on_eqs T eqs cu)) end (* Obtain head equivalences of cu, where each term is simplified. *) fun get_list_head_equiv ctxt T (id, cu) = let fun process_head_equiv (id', eq_th) = let val infos = simp_list_expr ctxt T (id', Thm.rhs_of eq_th) in map (BoxID.merge_eq_infos ctxt (id', eq_th)) infos end in - cu |> RewriteTable.get_head_equiv ctxt + cu |> Auto2_Setup.RewriteTable.get_head_equiv ctxt |> BoxID.merge_box_with_info ctxt id |> maps process_head_equiv |> filter_out (Thm.is_reflexive o snd) end (* Find ways to modify ct once by rewriting one of the subterms. *) fun list_expand_once ctxt T (id, ct) = let val cus = dest_subs ctxt T ct fun get_equiv cu = get_list_head_equiv ctxt T (id, cu) fun process_info (id', eq) = (id', rewrite_on_eqs T [eq] ct) in map process_info (maps get_equiv cus) end (* Find all ways to write ct, up to a certain limit. *) fun list_expand ctxt T (id, ct) = let - val max_ac = Config.get ctxt AC_ProofSteps.max_ac + val max_ac = Config.get ctxt Auto2_HOL_Extra_Setup.AC_ProofSteps.max_ac fun list_equiv_eq_better (id, th) (id', th') = let val seq1 = dest_list_full ctxt T (Thm.rhs_of th) val seq2 = dest_list_full ctxt T (Thm.rhs_of th') in Util.is_subseq (op aconv o apply2 Thm.term_of) (seq1, seq2) andalso BoxID.is_eq_ancestor ctxt id id' end fun has_list_equiv_eq_better infos info' = exists (fn info => list_equiv_eq_better info info') infos fun helper (old, new) = case new of [] => old | (id', eq_th) :: rest => if length old + length new > max_ac then old @ take (max_ac - length old) new else let val old' = ((id', eq_th) :: old) val rhs_expand = (list_expand_once ctxt T (id', Thm.rhs_of eq_th)) |> Util.max_partial list_equiv_eq_better |> map (BoxID.merge_eq_infos ctxt (id', eq_th)) |> filter_out (has_list_equiv_eq_better (old' @ rest)) in helper (old', rest @ rhs_expand) end (* Start term *) val start = simp_list_expr ctxt T (id, ct) in helper ([], start) end fun normalize_list_assoc T ct = case case_head T (Thm.term_of ct) of LIST_APPEND => (case case_head T (dest_arg1 (Thm.term_of ct)) of LIST_APPEND => Conv.every_conv [rewr_obj_eq @{thm List.append_assoc}, Conv.arg_conv (normalize_list_assoc T)] ct | _ => Conv.all_conv ct) | _ => Conv.all_conv ct fun normalize_list T ct = case case_head T (Thm.term_of ct) of LIST_APPEND => Conv.every_conv [ Conv.binop_conv (normalize_list T), Conv.try_conv (rewr_obj_eq @{thm List.append.append_Nil}), Conv.try_conv (rewr_obj_eq @{thm List.append_Nil2}), normalize_list_assoc T] ct | LIST_CONS => Conv.every_conv [rewr_obj_eq @{thm cons_to_append}, normalize_list T] ct | _ => Conv.all_conv ct fun list_expand_equiv_fn ctxt item1 item2 = let val {id = id1, tname = tname1, ...} = item1 val {id = id2, tname = tname2, ...} = item2 val (ct1, ct2) = (the_single tname1, the_single tname2) val (t1, t2) = (Thm.term_of ct1, Thm.term_of ct2) val id = BoxID.merge_boxes ctxt (id1, id2) val T = the (get_list_ty t1) in if Term_Ord.term_ord (t2, t1) = LESS then [] - else if RewriteTable.is_equiv id ctxt (ct1, ct2) then [] + else if Auto2_Setup.RewriteTable.is_equiv id ctxt (ct1, ct2) then [] else if not (is_list_head T t1) orelse not (is_list_head T t2) then [] else let val expand1 = list_expand ctxt T (id, ct1) val expand2 = list_expand ctxt T (id, ct2) fun get_equiv ((id1, eq_th1), (id2, eq_th2)) = let val ct1 = Thm.rhs_of eq_th1 val ct2 = Thm.rhs_of eq_th2 val ts1 = dest_list_full ctxt T ct1 val ts2 = dest_list_full ctxt T ct2 in if eq_list (op aconv o apply2 Thm.term_of) (ts1, ts2) then let val eq_th1' = normalize_list T ct1 val eq_th2' = normalize_list T ct2 val id' = BoxID.merge_boxes ctxt (id1, id2) val eq = Util.transitive_list [ eq_th1, eq_th1', meta_sym eq_th2', meta_sym eq_th2] in [(id', to_obj_eq eq)] end else [] end in (maps get_equiv (Util.all_pairs (expand1, expand2))) |> filter (BoxID.has_incr_id o fst) |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) - |> map (fn (id, th) => Update.thm_update_sc 1 (id, th)) + |> map (fn (id, th) => Auto2_Setup.Update.thm_update_sc 1 (id, th)) end end val list_expand_equiv = {name = "list_expand_equiv", args = [TypedMatch (TY_TERM, @{term_pat "?A::?'a list"}), TypedMatch (TY_TERM, @{term_pat "?B::?'a list"})], func = TwoStep list_expand_equiv_fn} fun list_expand_unit_fn ctxt item = let val {id, tname, ...} = item val ct = the_single tname val t = Thm.term_of ct val T = the (get_list_ty t) in if not (is_list_head T t) then [] else let val expand = list_expand ctxt T (id, ct) fun process_expand (id', eq_th) = let val ct = Thm.rhs_of eq_th val ts = dest_list_full ctxt T ct in if length ts <= 1 then let val eq' = normalize_list T ct val eq = Util.transitive_list [eq_th, eq'] in [(id', to_obj_eq eq)] end else [] end in (maps process_expand expand) |> filter (BoxID.has_incr_id o fst) - |> map (fn (id, th) => Update.thm_update_sc 1 (id, th)) + |> map (fn (id, th) => Auto2_Setup.Update.thm_update_sc 1 (id, th)) end end val list_expand_unit = {name = "list_expand_unit", args = [TypedMatch (TY_TERM, @{term_pat "?A::?'a list"})], func = OneStep list_expand_unit_fn} val add_list_proofsteps = fold add_prfstep [ list_expand_equiv, list_expand_unit ] end (* structure List_AC. *) val _ = Theory.setup List_AC.add_list_proofsteps diff --git a/thys/Auto2_HOL/HOL/logic_steps_test.ML b/thys/Auto2_HOL/HOL/logic_steps_test.ML --- a/thys/Auto2_HOL/HOL/logic_steps_test.ML +++ b/thys/Auto2_HOL/HOL/logic_steps_test.ML @@ -1,257 +1,261 @@ (* File: logic_steps_test.ML Author: Bohua Zhan Unit test for logic_steps.ML. *) local val ctxt = @{context} val thy = @{theory} in val test_norm_all_disj = let fun test (str1, str2) = let val (t1, t2) = (Proof_Context.read_term_pattern ctxt str1, Proof_Context.read_term_pattern ctxt str2) val ct1 = Thm.cterm_of ctxt t1 - val (vars, ts) = Logic_ProofSteps.strip_all_disj t1 - val res_th = Logic_ProofSteps.norm_all_disj ctxt ct1 + val (vars, ts) = Auto2_Setup.Logic_ProofSteps.strip_all_disj t1 + val res_th = Auto2_Setup.Logic_ProofSteps.norm_all_disj ctxt ct1 in - if not (Logic_ProofSteps.mk_all_disj (vars, ts) aconv t2) then + if not (Auto2_Setup.Logic_ProofSteps.mk_all_disj (vars, ts) aconv t2) then let val _ = trace_t ctxt "Input" t1 - val _ = trace_t ctxt "Output" (Logic_ProofSteps.mk_all_disj (vars, ts)) + val _ = trace_t ctxt "Output" + (Auto2_Setup.Logic_ProofSteps.mk_all_disj (vars, ts)) val _ = trace_t ctxt "Expected" t2 in raise Fail "test_strip_all_disj" end else if not (Util.rhs_of res_th aconv t2) then let val _ = trace_t ctxt "Input" t1 val _ = trace_t ctxt "Output" (Util.rhs_of res_th) val _ = trace_t ctxt "Expected" t2 in raise Fail "test_norm_all_disj" end else () end val test_data = [ ("A::bool", "A::bool"), ("A | B", "A | B"), ("A | B | C | D", "A | B | C | D"), ("((A | B) | C) | D", "A | B | C | D"), ("(A | B | C) | (D | E) | F", "A | B | C | D | E | F"), ("~~A", "A::bool"), ("~~((A | B) | C)", "A | B | C"), ("~(A & B)", "~A | ~B"), ("~(A & B & C & D)", "~A | ~B | ~C | ~D"), ("~(((A & B) & C) & D)", "~A | ~B | ~C | ~D"), ("A --> B", "~A | B"), ("A --> B --> C --> D", "~A | ~B | ~C | D"), ("A & B --> C", "~A | ~B | C"), ("A --> B | C", "~A | B | C"), ("p & (q --> r) --> s", "~p | ~(q --> r) | s"), ("~ (p & (q --> r)) | s", "~p | ~(q --> r) | s"), ("!x. A x", "!x. A x"), ("!x y z. P x y z", "!x y z. P x y z"), ("!x:S. A x", "!x. x ~: S | A x"), ("!x:S. !y:S. !z:S. P x y z", "!x y z. x ~: S | y ~: S | z ~: S | P x y z"), ("!x. A x --> B x --> C x", "!x. ~A x | ~B x | C x"), ("!x. A x --> (!y. B x y --> (!z. C x y z))", "!x y z. ~A x | ~B x y | C x y z"), ("!x:S. A x --> (!y:S. B x y --> (!z:S. C x y z))", "!x y z. x ~: S | ~A x | y ~: S | ~B x y | z ~: S | C x y z"), ("!x. (!y. A y) --> B x", "!x. ~(!y. A y) | B x"), ("~(EX x. P x)", "!x. ~P x"), ("~(EX x:S. P x)", "!x. x ~: S | ~P x"), ("~(EX x. P x & Q x)", "!x. ~P x | ~Q x"), ("~(EX x:S. P x & Q x)", "!x. x ~: S | ~P x | ~Q x"), ("~(EX x y. P x y)", "!x y. ~P x y"), ("~(EX x:S. EX y:S. P x y)", "!x y. x ~: S | y ~: S | ~P x y"), ("~(EX x. P x & (EX y. Q x y))", "!x y. ~P x | ~Q x y"), ("p & (q --> r) --> s", "~p | ~(q --> r) | s"), ("~ (p & (q --> r)) | s", "~p | ~(q --> r) | s"), ("!x. ?P x", "!x. ?P x"), ("(!x. P x) | (!y. Q y)", "!x y. P x | Q y"), ("!x. (P::'a=>bool) y", "(P::'a=>bool) y"), ("!x. P (y::'a) --> Q y", "~P (y::'a) | Q y") ] in map test test_data end fun free_to_var t = case t of Free (x, T) => Var ((x, 0), T) | _ => raise Fail "free_to_var" val test_disj_prop_match = let fun test (str1, str2) = let val (t, u) = (Proof_Context.read_term_pattern ctxt str1, Syntax.read_term ctxt str2) - val (var_u, us) = Logic_ProofSteps.strip_all_disj u + val (var_u, us) = Auto2_Setup.Logic_ProofSteps.strip_all_disj u val var_u' = map free_to_var var_u val us' = map (subst_atomic (var_u ~~ var_u')) us val (cvar_u, cus) = (map Term.dest_Var var_u', map (Thm.cterm_of ctxt) us') - val (var_t, ts) = (Logic_ProofSteps.strip_all_disj t) |> Logic_ProofSteps.replace_disj_vars ctxt - val res = Logic_ProofSteps.disj_prop_match ctxt ([], fo_init) (t, (var_t, ts), (cvar_u, cus)) - val u' = Logic_ProofSteps.mk_all_disj (var_u, us) + val (var_t, ts) = (Auto2_Setup.Logic_ProofSteps.strip_all_disj t) |> Auto2_Setup.Logic_ProofSteps.replace_disj_vars ctxt + val res = Auto2_Setup.Logic_ProofSteps.disj_prop_match ctxt ([], fo_init) (t, (var_t, ts), (cvar_u, cus)) + val u' = Auto2_Setup.Logic_ProofSteps.mk_all_disj (var_u, us) in case res of [(([], instsp), eq)] => if Util.lhs_of eq aconv (Util.subst_term_norm instsp t) andalso Util.rhs_of eq aconv u' then () else let val _ = trace_t ctxt "t" t val _ = trace_t ctxt "u" u val _ = trace_thm ctxt "Output" eq in raise Fail "disj_prop_match: wrong equation" end | _ => let val _ = trace_t ctxt "t" t val _ = trace_t ctxt "u" u val _ = tracing ("Number of results: " ^ (string_of_int (length res))) in raise Fail "disj_prop_match: wrong number of equations" end end val test_data = [ ("A::bool", "A::bool"), ("A | B | C", "A | B | C"), ("(A | B) | C", "A | B | C"), ("A --> B --> C", "~A | ~B | C"), ("B --> A --> C", "C | ~B | ~A"), ("A & (B --> C) --> D", "A & (B --> C) --> D"), ("p & (q --> r) --> s", "~ (p & (q --> r)) | s"), ("!x. P x | Q x", "!x. P x | Q x"), ("!x::nat. Q x | P x", "!x::nat. P x | Q x"), ("!x::nat. P x --> Q x", "!x::nat. Q x | ~P x"), ("!x y. P x y | Q x y", "!x y. Q x y | P x y"), ("!a. P a | Q a", "!x. P x | Q x"), ("~(EX x. P x)", "~(EX x. P x)"), ("~(EX x. P x)", "!x. ~P x"), ("!x. ~P x", "~(EX x. P x)"), ("~(EX x:S. P x)", "~(EX x:S. P x)"), ("~(EX x:S. P x)", "!x:S. ~P x"), ("~(EX x:S. EX y:S. P x y)", "!x:S. !y:S. ~P x y"), ("!x. P x ?a", "!x. P x a"), ("!x:?S. P x", "!x:S. P x"), ("~(EX x. P x ?a)", "!x. ~P x a"), ("~(EX x:?S. P x)", "!x:S. ~P x"), ("!x. ?P x", "!x. P x"), ("!x y. ?P x | ?Q y", "!x y. P x | Q y"), ("~(EX x. EX y. ?P x & ?Q y)", "~(EX x. EX y. P x & Q y)"), ("!x y. ?P x | ?Q y", "!x y. P y | Q x"), ("!x. P y --> Q y", "P y --> Q y") ] in map test test_data end val test_disj_prop_no_match = let fun test (str1, str2) = let val (t, u) = (Proof_Context.read_term_pattern ctxt str1, Syntax.read_term ctxt str2) - val (var_u, us) = Logic_ProofSteps.strip_all_disj u + val (var_u, us) = Auto2_Setup.Logic_ProofSteps.strip_all_disj u val var_u' = map free_to_var var_u val us' = map (subst_atomic (var_u ~~ var_u')) us val (cvar_u, cus) = (map Term.dest_Var var_u', map (Thm.cterm_of ctxt) us') - val (var_t, ts) = (Logic_ProofSteps.strip_all_disj t) |> Logic_ProofSteps.replace_disj_vars ctxt - val res = Logic_ProofSteps.disj_prop_match ctxt ([], fo_init) (t, (var_t, ts), (cvar_u, cus)) + val (var_t, ts) = (Auto2_Setup.Logic_ProofSteps.strip_all_disj t) |> Auto2_Setup.Logic_ProofSteps.replace_disj_vars ctxt + val res = Auto2_Setup.Logic_ProofSteps.disj_prop_match ctxt ([], fo_init) (t, (var_t, ts), (cvar_u, cus)) in if null res then () else let val _ = trace_t ctxt "t" t val _ = trace_t ctxt "u" u in raise Fail "disj_prop_no_match: expected no match" end end val test_data = [ ("!x y. ?P x | ?Q y", "!x y. P x y | Q x y"), ("!x::'a. P (a::'a) | Q x", "!a::'a. P (a::'a) | Q a") ] in map test test_data end val test_match_update = let val x = Free ("x", natT) val y = Free ("y", natT) val A = Free ("A", boolT) val B = Free ("B", natT --> boolT) val D = Free ("D", boolT) val ctxt' = ctxt |> fold Variable.declare_term [x, y, A, B, D] - val ctxt'' = ctxt' |> RewriteTable.add_rewrite ([], assume_eq thy (x, y)) |> snd + val ctxt'' = + ctxt' + |> Auto2_Setup.RewriteTable.add_rewrite ([], assume_eq thy (x, y)) + |> snd fun disj_to_ritem th = let val subs = strip_disj (prop_of' th) in - Fact (TY_DISJ, bFalse :: disj :: subs, th) + Fact (Auto2_Setup.Logic_ProofSteps.TY_DISJ, bFalse :: disj :: subs, th) end fun read_prop str = str |> Proof_Context.read_term_pattern ctxt' |> mk_Trueprop fun updt_to_thm updt = case updt of - AddItems {raw_items, ...} => map BoxItem.get_thm_raw raw_items + AddItems {raw_items, ...} => map Auto2_Setup.BoxItem.get_thm_raw raw_items | ResolveBox {th, ...} => [th] | _ => [] fun test_updt ctxt (disj_str, t_str, res_strs) = let val disj_th = Util.assume_thm ctxt' (read_prop disj_str) val th = Util.assume_thm ctxt' (read_prop t_str) val res_ts = map read_prop res_strs - val mk_box_item = BoxItem.mk_box_item ctxt' + val mk_box_item = Auto2_Setup.BoxItem.mk_box_item ctxt' val items = [mk_box_item (0, [], 0, disj_to_ritem disj_th), - mk_box_item (0, [0], 0, Update.thm_to_ritem th)] - val updts = ProofStep.apply_prfstep ctxt items Logic_ProofSteps.match_update_prfstep + mk_box_item (0, [0], 0, Auto2_Setup.Update.thm_to_ritem th)] + val updts = Auto2_Setup.ProofStep.apply_prfstep ctxt items Auto2_Setup.Logic_ProofSteps.match_update_prfstep val res_ts' = maps updt_to_thm updts |> map Thm.prop_of in if eq_set (op aconv) (res_ts, res_ts') then () else let val _ = trace_t ctxt' "Input disj:" (prop_of' disj_th) val _ = trace_t ctxt' "Input t:" (prop_of' th) val _ = trace_tlist ctxt' "Expected:" res_ts val _ = trace_tlist ctxt' "Actual:" res_ts' in raise Fail "test_match_update" end end val data1 = [("B x | B y | ~A", "A", ["B x | B y"]), ("B x | B y | ~A", "~B x", ["B y | ~A"]), ("B x | B y | ~A", "~B y", [])] val _ = map (test_updt ctxt') data1 val data2 = [("B x | B y | ~A", "~B x", ["B y | ~A"]), ("B x | B y | ~A", "D", []), ("B x | B y", "~B x", ["False"])] val _ = map (test_updt ctxt'') data2 in () end end diff --git a/thys/Auto2_HOL/HOL/matcher_test.ML b/thys/Auto2_HOL/HOL/matcher_test.ML --- a/thys/Auto2_HOL/HOL/matcher_test.ML +++ b/thys/Auto2_HOL/HOL/matcher_test.ML @@ -1,269 +1,269 @@ (* File: matcher_test.ML Author: Bohua Zhan Unit test of matcher.ML. *) local val init = @{context} (* Conversion between (id, th) used in rewrite.ML and (id, t) used in test cases. *) fun to_term_info (id, th) = (id, Util.rhs_of th) (* Comparison of list of (box_id, term) pairs. *) fun eq_info_list (l1, l2) = (length l1 = length l2) andalso eq_set (eq_pair (op =) (op aconv)) (l1, l2) fun print_term_info ctxt (id, t) = "(" ^ (BoxID.string_of_box_id id) ^ ", " ^ (Syntax.string_of_term ctxt t) ^ ")" fun print_term_infos ctxt lst = commas (map (print_term_info ctxt) lst) (* info is expected, in box_id * term form, info' is returned value, in box_id * thm form (to be printed). *) fun assert_eq_info (info, info') txt ctxt = if eq_info_list (info, map to_term_info info') then ctxt else let val _ = tracing ("Expected: " ^ print_term_infos ctxt info ^ "\n" ^ "Actual: " ^ print_infos ctxt info') in raise Fail txt end fun add_prim_id (prim_id, id) ctxt = let val (prim_id', ctxt') = BoxID.add_prim_id id ctxt val _ = assert (prim_id = prim_id') "add_prim_id" in ctxt' end fun declare_term str ctxt = Proof_Context.augment (Syntax.read_term ctxt str) ctxt fun declare_pat str ctxt = Proof_Context.augment (Proof_Context.read_term_pattern ctxt str) ctxt fun add_rewrite id (str1, str2) ctxt = let val (t1, t2) = (Syntax.read_term ctxt str1, Syntax.read_term ctxt str2) val thy = Proof_Context.theory_of ctxt val th = assume_eq thy (t1, t2) in - ctxt |> RewriteTable.add_rewrite (id, th) |> snd + ctxt |> Auto2_Setup.RewriteTable.add_rewrite (id, th) |> snd end (* Check th is actually t(env) == u. *) fun check_info (t, u) txt ctxt ((_, inst), th) = let val t' = Util.subst_term_norm inst t val (lhs, rhs) = Logic.dest_equals (Thm.prop_of th) in if lhs aconv t' andalso rhs aconv u then () else let val _ = tracing ("inst does not match th.\nt', u' = " ^ (Util.string_of_terms ctxt [t', u]) ^ "\nth = " ^ (Syntax.string_of_term ctxt (Thm.prop_of th))) in raise Fail txt end end fun assert_match_gen matcher ((str_t, str_u), info_str) ctxt = let val (t, u) = (Proof_Context.read_term_pattern ctxt str_t, Syntax.read_term ctxt str_u) val _ = trace_tlist ctxt "Matching" [t, u] val info = map (apsnd (Syntax.read_term ctxt)) info_str val info' = matcher ctxt [] (t, Thm.cterm_of ctxt u) ([], fo_init) val _ = map (check_info (t, u) "assert_match" ctxt) info' val info'' = map (fn ((id, _), th) => (id, meta_sym th)) info' in assert_eq_info (info, info'') "assert_match" ctxt end -val assert_match = assert_match_gen Matcher.match +val assert_match = assert_match_gen Auto2_Setup.Matcher.match fun assert_match_list (pairs_str, info_str) ctxt = let val pairs = map (fn (is_head, (str_t, str_u)) => (is_head, (Proof_Context.read_term_pattern ctxt str_t, Thm.cterm_of ctxt (Syntax.read_term ctxt str_u)))) pairs_str val info = map (apsnd (map (Syntax.read_term ctxt))) info_str - val info' = Matcher.rewrite_match_list ctxt pairs ([], fo_init) + val info' = Auto2_Setup.Matcher.rewrite_match_list ctxt pairs ([], fo_init) fun check_info_list (instsp, ths) = let fun check_pair ((t, cu), th) = check_info (t, Thm.term_of cu) "assert_match_list" ctxt (instsp, th) in map check_pair (map snd pairs ~~ ths) end val _ = map check_info_list info' val info'' = map (fn ((id, _), ths) => (id, map meta_sym ths)) info' val eq_info_list' = eq_set (eq_pair (op =) (eq_list (op aconv))) fun to_term_info' (id, ths) = (id, map Util.rhs_of ths) in if eq_info_list' (info, map to_term_info' info'') then ctxt else let val _ = tracing ("got " ^ print_infos' ctxt info'') in raise Fail "assert_match_list" end end fun done str _ = let val _ = tracing ("Finished " ^ str) in () end in (* Basic rewriting. *) val test_rewrite = init |> add_prim_id (1, []) |> add_prim_id (2, []) |> declare_term "[a::nat, b, c]" |> declare_term "f::(nat => nat => nat)" |> declare_pat "[?m::nat, ?n]" |> declare_pat "?A::('a::{plus})" |> add_rewrite [1] ("a", "b") |> add_rewrite [2] ("a", "c") |> fold assert_match [(("a", "b"), [([1], "a")]), (("c", "a"), [([2], "c")]), (("?n + ?n", "a + b"), [([1], "a + a")]), (("?A + ?A", "c + a"), [([2], "c + c")]), (("?A + ?A + ?A", "a + b + c"), [([1, 2], "a + a + a")]), (("f ?m ?n", "f a b"), [([], "f a b")]), (("f ?m ?m", "f c a"), [([2], "f c c")])] |> done "test_rewrite" (* Rewriting from atomic to function application. *) val test_rewrite2 = init |> add_prim_id (1, []) |> add_prim_id (2, []) |> declare_term "[a::nat, b, c]" |> declare_pat "[?n::nat]" |> add_rewrite [1] ("a", "b + c") |> add_rewrite [2] ("b", "c") |> assert_match (("?n + ?n", "a"), [([1, 2], "b + b")]) |> done "test_rewrite2" (* Abstractions. *) val test_abstraction = init |> declare_term "[m::nat, n]" |> declare_term "P::(nat => nat => nat)" |> declare_pat "?A::(nat => nat => 'a::{plus})" |> fold assert_match [(("%x y. ?A x y", "%m n. P m n"), [([], "%m n. P m n")]), (("%x y. ?A x y + ?A y x", "%m n. P m n + P n m"), [([], "%m n. P m n + P n m")]), (("%x y. ?A x y", "%x y. (x::nat) + y + y"), [([], "%x y. (x::nat) + y + y")])] |> done "test_abstraction" (* Abstractions with rewriting. *) val test_abstraction_rewrite = init |> declare_term "[m::nat, n, p]" |> declare_term "P::(nat => nat => nat)" |> declare_term "f::(nat => nat => 'a::{plus})" |> declare_pat "?A::(nat => ?'a::{plus})" |> add_rewrite [] ("m", "n") |> add_rewrite [] ("p", "0::nat") |> assert_match (("%x. ?A x + ?A x", "%x. P x m + P x n"), [([], "%x. P x m + P x m")]) |> assert_match (("%x. f x 0", "%x. f x p"), [([], "%x. f x 0")]) |> done "test_abstraction_rewrite" (* Test matching of higher order patterns. *) val test_higher_order = init |> declare_term "f::(nat => nat)" |> declare_pat "?f::(nat => nat)" |> assert_match (("%n. (?f n, ?f (n + 1))", "%n. (f n, f (n + 1))"), [([], "%n. (f n, f (n + 1))")]) |> assert_match (("%n. (?f (n + 1), ?f n)", "%n. (f (n + 1), f n)"), [([], "%n. (f (n + 1), f n)")]) |> done "test_higher_order" (* Test handling of schematic type variables. *) val test_match_type = init |> assert_match (("image_mset ?f {#}", "{#i. i :# {#}#}"), [([], "{#i. i :# {#}#}")]) |> assert_match (("image_mset ?f {#}", "{#(i::nat). i :# {#}#}"), [([], "{#(i::nat). i :# {#}#}")]) (* sorted [] has concrete type, but not []. *) |> assert_match (("sorted []", "sorted []"), [([], "sorted []")]) |> done "test_match_type" (* Test special schematic variable ?NUMC. *) val test_numc = init |> declare_term "[k::int, m]" |> add_rewrite [] ("k", "1::int") |> assert_match (("?NUMC::int", "k"), [([], "1::int")]) |> assert_match (("?NUMC", "k"), [([], "1::int")]) |> assert_match (("?NUMC", "m"), []) |> done "test_numc" (* Test match list. *) val test_match_list = init |> declare_term "[k::nat, l]" |> declare_pat "?a::nat" |> add_rewrite [] ("k", "l") |> assert_match_list ([(false, ("?a", "k")), (false, ("?a", "l"))], [([], ["k", "k"])]) |> assert_match_list ([(false, ("?a", "k")), (true, ("?a", "l"))], []) |> assert_match_list ([(true, ("?a", "k")), (false, ("?a", "l"))], [([], ["k", "k"])]) |> assert_match_list ([(true, ("?a + b", "k + b")), (true, ("?a + b", "l + b"))], [([], ["k + b", "k + b"])]) |> done "test_match_list" (* Test of pre-matcher. *) fun assert_pre_match_gen res at_head (str_t, str_u) ctxt = let val (t, cu) = (Proof_Context.read_term_pattern ctxt str_t, Thm.cterm_of ctxt (Syntax.read_term ctxt str_u)) - val pre_match_fn = if at_head then Matcher.pre_match_head - else Matcher.pre_match + val pre_match_fn = if at_head then Auto2_Setup.Matcher.pre_match_head + else Auto2_Setup.Matcher.pre_match in if pre_match_fn ctxt (t, cu) = res then ctxt else let val _ = trace_tlist ctxt "Pattern, term: " [t, Thm.term_of cu] in raise Fail ("assert_pre_match (expected " ^ (Util.string_of_bool res) ^ ")") end end val assert_pre_match = assert_pre_match_gen true false val assert_not_pre_match = assert_pre_match_gen false false val test_pre_match = init |> declare_term "[m::nat, n]" |> declare_term "[a::'a]" |> declare_pat "?n::nat" |> declare_pat "?a::?'a" |> assert_pre_match ("?a", "n") |> assert_pre_match ("?n", "n") |> assert_not_pre_match ("?n", "a") |> assert_not_pre_match ("m", "n") |> done "test_pre_match" val test_pre_match_quant = init |> declare_term "[P::(nat => bool), Q]" |> declare_term "[S::(nat => nat => bool), T]" |> declare_term "[x::nat, y, z]" |> assert_pre_match ("ALL x. P x", "ALL y. P y") |> assert_not_pre_match ("ALL x. P x", "ALL x. Q x") |> assert_pre_match ("ALL x y. S x y", "ALL y z. S y z") |> assert_not_pre_match ("ALL x y. S x y", "ALL x y. S y x") |> assert_not_pre_match ("ALL x y. S x y", "ALL x y. T x y") |> assert_pre_match ("ALL x y. ?S x y", "ALL x y. x < y") |> assert_pre_match ("ALL x. x + y = z & ?P x", "ALL x. x + y = z & P x") |> done "test_pre_match_quant" end diff --git a/thys/Auto2_HOL/HOL/nat_sub.ML b/thys/Auto2_HOL/HOL/nat_sub.ML --- a/thys/Auto2_HOL/HOL/nat_sub.ML +++ b/thys/Auto2_HOL/HOL/nat_sub.ML @@ -1,661 +1,663 @@ (* File: nat_sub.ML Author: Bohua Zhan Normalization of expressions involving subtraction on natural numbers. *) signature NAT_SUB = sig val fheads: term list val norm_plus1: wfconv val norm_plus: wfconv val norm_minus': wfconv val move_outmost: term -> wfconv val cancel_terms: wfconv val norm_minus: wfconv type monomial = cterm list * int val reduce_monomial_list: monomial list -> monomial list val add_polynomial_list: monomial list * monomial list -> monomial list val norm_minus_ct: cterm -> monomial list val norm_ring_term: cterm -> term val get_sub_head_equiv: Proof.context -> box_id * cterm -> (box_id * (wfterm * thm)) list val nat_sub_expand_once: Proof.context -> box_id * wfterm -> (box_id * (wfterm * thm)) list val nat_sub_expand: Proof.context -> box_id * cterm -> (box_id * (wfterm * thm)) list val nat_sub_expand_equiv: proofstep val nat_sub_expand_unit: proofstep val add_nat_sub_proofsteps: theory -> theory end; structure NatSub : NAT_SUB = struct +structure WfTerm = Auto2_Setup.WfTerm + val fheads = [@{term "(+) :: (nat => nat => nat)"}, @{term "(-) :: (nat => nat => nat)"}, @{term "(*) :: (nat => nat => nat)"}] val plus_info = Nat_Util.plus_ac_on_typ @{theory} natT val is_numc = UtilArith.is_numc val dest_numc = UtilArith.dest_numc val is_plus = UtilArith.is_plus val is_minus = UtilArith.is_minus val is_times = UtilArith.is_times val is_zero = UtilArith.is_zero val is_one = UtilArith.is_one val wf_rewr_obj_eq = WfTerm.rewr_obj_eq fheads val nat0 = Nat_Util.nat0 val mult_1 = @{thm mult_1} (* 1 * a = a *) val mult_1_right = @{thm mult_1_right} (* a * 1 = a *) val mult_0 = @{thm mult_0} (* 0 * a = 0 *) val mult_0_right = @{thm mult_0_right} (* a * 0 = 0 *) val add_0 = @{thm add_0} (* 0 + a = a *) val add_0_right = @{thm add_0_right} (* a + 0 = a *) val add_assoc = wf_rewr_obj_eq @{thm add_ac(1)} val add_assoc_sym = wf_rewr_obj_eq (obj_sym @{thm add_ac(1)}) val add_comm = wf_rewr_obj_eq @{thm add_ac(2)} val swap_add = WfTerm.every_conv [add_assoc, WfTerm.arg_conv add_comm, add_assoc_sym] val times_assoc = wf_rewr_obj_eq @{thm mult_ac(1)} val times_assoc_sym = wf_rewr_obj_eq (obj_sym @{thm mult_ac(1)}) val times_comm = wf_rewr_obj_eq @{thm mult_ac(2)} val swap_times = WfTerm.every_conv [times_assoc, WfTerm.arg_conv times_comm, times_assoc_sym] val distrib_r_th = @{thm nat_distrib(1)} val distrib_l_th = @{thm nat_distrib(2)} val nat_fold_wfconv = WfTerm.conv_of Nat_Util.nat_fold_conv val cancel_0_left = WfTerm.try_conv (wf_rewr_obj_eq @{thm add_0_left}) val append_0_left = wf_rewr_obj_eq (obj_sym @{thm add_0_left}) (* When comparing atoms, constants are greater than non-constants. *) fun compare_atom (t1, t2) = if is_numc t1 andalso is_numc t2 then EQUAL else if is_numc t1 then GREATER else if is_numc t2 then LESS else Term_Ord.term_ord (t1, t2) (* Multiply a monomial with an atom. *) fun norm_mult_atom wft = let val t = WfTerm.term_of wft val (arg1, arg2) = Util.dest_binop_args t in if is_one arg1 then wf_rewr_obj_eq mult_1 wft else if is_one arg2 then wf_rewr_obj_eq mult_1_right wft else if is_zero arg1 then wf_rewr_obj_eq mult_0 wft else if is_zero arg2 then wf_rewr_obj_eq mult_0_right wft else if is_times arg1 then case compare_atom (dest_arg arg1, arg2) of GREATER => WfTerm.every_conv [swap_times, WfTerm.arg1_conv norm_mult_atom] wft | EQUAL => if is_numc (dest_arg arg1) andalso is_numc arg2 then WfTerm.every_conv [times_assoc, WfTerm.arg_conv nat_fold_wfconv] wft else WfTerm.all_conv wft | _ => WfTerm.all_conv wft else case compare_atom (arg1, arg2) of GREATER => times_comm wft | EQUAL => if is_numc arg1 andalso is_numc arg2 then nat_fold_wfconv wft else WfTerm.all_conv wft | _ => WfTerm.all_conv wft end (* Multiply two monomials. *) fun norm_mult_monomial wft = let val t = WfTerm.term_of wft val (_, arg2) = Util.dest_binop_args t in if is_times arg2 then WfTerm.every_conv [times_assoc_sym, WfTerm.arg1_conv norm_mult_monomial, norm_mult_atom] wft else norm_mult_atom wft end (* Destruct t into the form arg * coeff, where coeff is a constant. *) fun dest_monomial t = if is_times t andalso is_numc (dest_arg t) then (dest_arg1 t, UtilArith.dest_numc (dest_arg t)) else if is_numc t then (Nat_Util.mk_nat 1, UtilArith.dest_numc t) else (t, 1) (* Normalize ct into the form arg * coeff. Example: a * 4 == a * 4, a == a * 1, 4 == 1 * 4. *) fun norm_monomial wft = let val t = WfTerm.term_of wft in if is_times t andalso is_numc (dest_arg t) then WfTerm.all_conv wft else if is_numc t then wf_rewr_obj_eq (obj_sym @{thm mult_1}) wft else wf_rewr_obj_eq (obj_sym @{thm mult_1_right}) wft end (* Compare two monomials by removing coefficient. *) fun compare_monomial (t1, t2) = let val (arg1, _) = dest_monomial t1 val (arg2, _) = dest_monomial t2 in if is_one arg1 andalso is_one arg2 then EQUAL else if is_one arg1 then GREATER else if is_one arg2 then LESS else Term_Ord.term_ord (arg1, arg2) end (* Combine two monomials. *) fun combine_monomial wft = WfTerm.every_conv [WfTerm.binop_conv norm_monomial, wf_rewr_obj_eq (obj_sym distrib_l_th), WfTerm.arg_conv nat_fold_wfconv] wft (* Normalize a + b, where a is a sum and b is an atom. If b is a constant, add it to the end of a. Otherwise, insert b into sorted position. *) fun norm_plus1 wft = let val (a, b) = wft |> WfTerm.term_of |> Util.dest_binop_args in if is_zero a then wf_rewr_obj_eq add_0 wft else if is_zero b then wf_rewr_obj_eq add_0_right wft else if is_plus a then case compare_monomial (dest_arg a, b) of LESS => WfTerm.all_conv wft | EQUAL => if is_numc (dest_arg a) andalso is_numc b then WfTerm.every_conv [add_assoc, WfTerm.arg_conv nat_fold_wfconv] wft else WfTerm.every_conv [add_assoc, WfTerm.arg_conv combine_monomial] wft | GREATER => WfTerm.every_conv [swap_add, WfTerm.arg1_conv norm_plus1] wft else case compare_monomial (a, b) of LESS => WfTerm.all_conv wft | EQUAL => if is_numc a andalso is_numc b then nat_fold_wfconv wft else combine_monomial wft | GREATER => wf_rewr_obj_eq @{thm add_ac(2)} wft end (* Normalize a + b, where a and b are both sums. *) fun norm_plus wft = let val (_, b) = wft |> WfTerm.term_of |> Util.dest_binop_args in if is_plus b then WfTerm.every_conv [ add_assoc_sym, WfTerm.arg1_conv norm_plus, norm_plus1] wft else norm_plus1 wft end (* Assume wft is a sum containing u. Move u into the outermost position. *) fun move_outmost u wft = if u aconv (WfTerm.term_of wft) then WfTerm.all_conv wft else if not (is_plus (WfTerm.term_of wft)) then raise Fail "move_outmost: u not found in wft." else let val (a, b) = wft |> WfTerm.term_of |> Util.dest_binop_args in if u aconv b then WfTerm.all_conv wft else if is_plus a then WfTerm.every_conv [WfTerm.arg1_conv (move_outmost u), swap_add] wft else if u aconv a then add_comm wft else raise Fail "move_outmost: u not found in wft." end (* Cancel terms with the same argument from the two sides. *) fun cancel_terms wft = let val (a, b) = wft |> WfTerm.term_of |> Util.dest_binop_args - val (ts1, ts2) = apply2 (ACUtil.dest_ac plus_info) (a, b) + val (ts1, ts2) = apply2 (Auto2_HOL_Extra_Setup.ACUtil.dest_ac plus_info) (a, b) fun find_same_arg (t1, t2) = case compare_monomial (t1, t2) of EQUAL => if t1 aconv nat0 orelse t2 aconv nat0 then NONE else SOME (t1, t2) | _ => NONE fun prepare_left t1 = if length ts1 = 1 then WfTerm.arg1_conv append_0_left else WfTerm.arg1_conv (move_outmost t1) fun prepare_right t2 = if length ts2 = 1 then WfTerm.arg_conv append_0_left else WfTerm.arg_conv (move_outmost t2) fun apply_th (t1, t2) wft = let val (_, n1) = dest_monomial t1 val (_, n2) = dest_monomial t2 val fold_cv = WfTerm.every_conv [ WfTerm.arg_conv nat_fold_wfconv, WfTerm.try_conv (wf_rewr_obj_eq mult_1), WfTerm.try_conv (wf_rewr_obj_eq mult_1_right)] val le_th = if n1 >= n2 then Nat_Util.nat_le_th n2 n1 else Nat_Util.nat_le_th n1 n2 in if n1 = n2 then wf_rewr_obj_eq @{thm nat_sub_combine} wft else if n1 > n2 then WfTerm.every_conv [ WfTerm.binop_conv (WfTerm.arg_conv norm_monomial), wf_rewr_obj_eq (le_th RS @{thm nat_sub_combine2}), WfTerm.arg1_conv (WfTerm.arg_conv fold_cv), WfTerm.arg1_conv norm_plus] wft else WfTerm.every_conv [ WfTerm.binop_conv (WfTerm.arg_conv norm_monomial), wf_rewr_obj_eq (le_th RS @{thm nat_sub_combine3}), WfTerm.arg_conv (WfTerm.arg_conv fold_cv), WfTerm.arg_conv norm_plus] wft end in case get_first find_same_arg (Util.all_pairs (ts1, ts2)) of NONE => WfTerm.all_conv wft | SOME (t1, t2) => WfTerm.every_conv [prepare_left t1, prepare_right t2, apply_th (t1, t2), WfTerm.arg1_conv cancel_0_left, WfTerm.arg_conv cancel_0_left, cancel_terms] wft end fun norm_mult_poly_monomial wft = let val t = WfTerm.term_of wft val (arg1, _) = Util.dest_binop_args t in if is_plus arg1 then WfTerm.every_conv [wf_rewr_obj_eq distrib_r_th, WfTerm.arg1_conv norm_mult_poly_monomial, WfTerm.arg_conv norm_mult_monomial, norm_plus] wft else norm_mult_monomial wft end fun norm_mult_polynomials wft = let val t = WfTerm.term_of wft val (_, arg2) = Util.dest_binop_args t in if is_plus arg2 then WfTerm.every_conv [wf_rewr_obj_eq distrib_l_th, WfTerm.arg1_conv norm_mult_polynomials, WfTerm.arg_conv norm_mult_poly_monomial, norm_plus] wft else norm_mult_poly_monomial wft end (* First step, put wft into the form a - b, where a and b are normalized sums. *) fun norm_minus' wft = let val t = WfTerm.term_of wft in if is_plus t then (* (a - b) + (c - d) = (a + c) - (b + d) *) WfTerm.every_conv [WfTerm.binop_conv norm_minus', wf_rewr_obj_eq @{thm nat_sub1}, WfTerm.binop_conv norm_plus, cancel_terms] wft else if is_minus t then (* (a - b) - (c - d) = (a + d) - (b + c) *) WfTerm.every_conv [WfTerm.binop_conv norm_minus', wf_rewr_obj_eq @{thm nat_sub2}, WfTerm.binop_conv norm_plus, cancel_terms] wft else if is_times t then (* (a - b) * (c - d) = (ac + bd) - (ad + bc) *) WfTerm.every_conv [ WfTerm.binop_conv norm_minus', wf_rewr_obj_eq @{thm nat_sub3}, WfTerm.binop_conv (WfTerm.binop_conv norm_mult_polynomials), WfTerm.binop_conv norm_plus, cancel_terms] wft else (* Normalize a into a - 0. *) wf_rewr_obj_eq @{thm nat_sub_norm} wft end fun norm_minus wft = WfTerm.every_conv [ norm_minus', WfTerm.try_conv (wf_rewr_obj_eq @{thm diff_zero})] wft (* Fast computation of the expected normalization. Return a list of monomials. *) type monomial = cterm list * int fun mk_plus (t1, t2) = Const (@{const_name plus}, natT --> natT --> natT) $ t1 $ t2 fun list_plus ts = let fun list_rev ts = case ts of [] => Nat_Util.mk_nat 0 | [t] => t | t :: ts' => mk_plus (list_rev ts', t) in list_rev (rev ts) end fun mk_times (t1, t2) = Const (@{const_name times}, natT --> natT --> natT) $ t1 $ t2 fun list_times ts = let fun list_rev ts = case ts of [] => Nat_Util.mk_nat 1 | [t] => t | t :: ts' => mk_times (list_rev ts', t) in list_rev (rev ts) end (* Compare two monomials *) fun compare_monomial_list ((l1, _), (l2, _)) = if null l1 andalso null l2 then EQUAL else if null l1 then GREATER else if null l2 then LESS else Term_Ord.term_ord (list_times (map Thm.term_of l1), list_times (map Thm.term_of l2)) (* Reduce a list of monomials: combine monomials of the same body. *) fun reduce_monomial_list ls = if null ls then [] else let val ((l1, c1), rest) = (hd ls, reduce_monomial_list (tl ls)) in case rest of [] => [(l1, c1)] | [(_, 0)] => [(l1, c1)] | (l2, c2) :: rest' => if eq_list (op aconvc) (l1, l2) then if c1 + c2 = 0 then rest' else (l1, c1 + c2) :: rest' else (l1, c1) :: (l2, c2) :: rest' end (* Multiply two monomials. *) fun mult_monomial ((l1, c1), (l2, c2)) = ((l1 @ l2) |> sort (compare_atom o apply2 Thm.term_of), c1 * c2) (* Multiply two such lists: take the pairwise product, sort within each monomial, then sort the list of monomials. *) fun mult_polynomial_term (ls1, ls2) = (Util.all_pairs (ls1, ls2)) |> map mult_monomial |> sort compare_monomial_list |> reduce_monomial_list fun add_polynomial_list (ls1, ls2) = (ls1 @ ls2) |> sort compare_monomial_list |> reduce_monomial_list fun negate_polynomial_list ls = map (fn (ls, n) => (ls, ~n)) ls fun norm_minus_ct ct = let val t = Thm.term_of ct in if is_plus t then add_polynomial_list (norm_minus_ct (Thm.dest_arg1 ct), norm_minus_ct (Thm.dest_arg ct)) else if is_minus t then add_polynomial_list ( norm_minus_ct (Thm.dest_arg1 ct), negate_polynomial_list (norm_minus_ct (Thm.dest_arg ct))) else if is_times t then mult_polynomial_term (norm_minus_ct (Thm.dest_arg1 ct), norm_minus_ct (Thm.dest_arg ct)) else if is_numc t then [([], dest_numc t)] else [([ct], 1)] end fun subterms_of ct = ct |> norm_minus_ct |> map fst |> flat fun to_monomial (l, c) = if null l then mk_nat c else if c = 1 then list_times (map Thm.term_of l) else mk_times (list_times (map Thm.term_of l), mk_nat c) fun norm_ring_term ct = let val dest_ring = norm_minus_ct ct val (ps, ms) = filter_split (fn (_, n) => n >= 0) dest_ring in if length ms > 0 then Const (@{const_name minus}, natT --> natT --> natT) $ list_plus (map to_monomial ps) $ list_plus (map to_monomial (negate_polynomial_list ms)) else list_plus (map to_monomial ps) end (* Obtain head equivalences of cu, normalized to a - b form, and where each term is simplified. *) fun get_sub_head_equiv ctxt (id, ct) = let fun process_wf_head_equiv (id', (wft, eq_th)) = let val cts = subterms_of (Thm.rhs_of eq_th) - val simps = WellformData.simplify ctxt fheads cts (id', wft) + val simps = Auto2_Setup.WellformData.simplify ctxt fheads cts (id', wft) fun process_simp (id'', (wft', eq_th')) = (BoxID.merge_boxes ctxt (id', id''), (wft', Util.transitive_list [eq_th, eq_th'])) in map process_simp simps end in - (id, ct) |> WellformData.get_head_equiv ctxt fheads + (id, ct) |> Auto2_Setup.WellformData.get_head_equiv ctxt fheads |> maps process_wf_head_equiv |> filter_out (Thm.is_reflexive o snd o snd) end (* Given wft in normalized form a - b, expand one of the subterms once. Calls the corresponding function in ac_steps on a and b, then combine the results. *) fun nat_sub_expand_once ctxt (id, wft) = let val ct = WfTerm.cterm_of wft val subt = subterms_of ct fun get_equiv cu = get_sub_head_equiv ctxt (id, cu) fun process_info (id', wf_eq) = (id', WfTerm.rewrite_on_eqs fheads [wf_eq] wft) in map process_info (maps get_equiv subt) end (* Find all ways to write ct, up to a certain limit. *) fun nat_sub_expand ctxt (id, ct) = let - val max_ac = Config.get ctxt AC_ProofSteps.max_ac + val max_ac = Config.get ctxt Auto2_HOL_Extra_Setup.AC_ProofSteps.max_ac fun ac_equiv_eq_better (id, (_, th)) (id', (_, th')) = let val seq1 = subterms_of (Thm.rhs_of th) val seq2 = subterms_of (Thm.rhs_of th') in Util.is_subseq (op aconvc) (seq1, seq2) andalso BoxID.is_eq_ancestor ctxt id id' end fun has_ac_equiv_eq_better infos info' = exists (fn info => ac_equiv_eq_better info info') infos fun helper (old, new) = case new of [] => old | (id', (wft, eq_th)) :: rest => if length old + length new > max_ac then old @ take (max_ac - length old) new else let val old' = ((id', (wft, eq_th)) :: old) fun merge_info (id'', (wft', eq_th')) = (BoxID.merge_boxes ctxt (id', id''), (wft', Util.transitive_list [eq_th, eq_th'])) val rhs_expand = (nat_sub_expand_once ctxt (id', wft)) |> Util.max_partial ac_equiv_eq_better |> map merge_info |> filter_out (has_ac_equiv_eq_better (old' @ rest)) in helper (old', rest @ rhs_expand) end (* Start term *) val ts = subterms_of ct - val start = (WellformData.cterm_to_wfterm ctxt fheads (id, ct)) - |> maps (WellformData.simplify ctxt fheads ts) + val start = (Auto2_Setup.WellformData.cterm_to_wfterm ctxt fheads (id, ct)) + |> maps (Auto2_Setup.WellformData.simplify ctxt fheads ts) in helper ([], start) end fun is_nat_sub_form t = if is_plus t orelse is_minus t orelse is_times t then fastype_of (dest_arg t) = natT else false fun nat_sub_expand_equiv_fn ctxt item1 item2 = let val {id = id1, tname = tname1, ...} = item1 val {id = id2, tname = tname2, ...} = item2 val (ct1, ct2) = (the_single tname1, the_single tname2) val (t1, t2) = (Thm.term_of ct1, Thm.term_of ct2) val id = BoxID.merge_boxes ctxt (id1, id2) in if not (is_nat_sub_form t1) orelse not (is_nat_sub_form t2) then [] else if Term_Ord.term_ord (t2, t1) = LESS then [] - else if RewriteTable.is_equiv id ctxt (ct1, ct2) then [] + else if Auto2_Setup.RewriteTable.is_equiv id ctxt (ct1, ct2) then [] else let val expand1 = nat_sub_expand ctxt (id, ct1) val expand2 = nat_sub_expand ctxt (id, ct2) fun get_equiv ((id1, (wft1, eq_th1)), (id2, (wft2, eq_th2))) = let val ct1 = Thm.rhs_of eq_th1 val ct2 = Thm.rhs_of eq_th2 val ts1 = norm_minus_ct ct1 val ts2 = norm_minus_ct ct2 in if eq_list (eq_pair (eq_list (op aconvc)) (op =)) (ts1, ts2) then let val (wft1', eq1) = norm_minus wft1 val (wft2', eq2) = norm_minus wft2 val _ = assert (WfTerm.term_of wft1' aconv WfTerm.term_of wft2') "nat_sub_expand_equiv_fn" val id' = BoxID.merge_boxes ctxt (id1, id2) val eq = Util.transitive_list [ eq_th1, eq1, meta_sym eq2, meta_sym eq_th2] in [(id', to_obj_eq eq)] end else [] end in (maps get_equiv (Util.all_pairs (expand1, expand2))) |> filter (BoxID.has_incr_id o fst) |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) - |> map (fn (id, th) => Update.thm_update_sc 1 (id, th)) + |> map (fn (id, th) => Auto2_Setup.Update.thm_update_sc 1 (id, th)) end end val nat_sub_expand_equiv = {name = "nat_sub_expand_equiv", args = [TypedMatch (TY_TERM, @{term_pat "?A::nat"}), TypedMatch (TY_TERM, @{term_pat "?B::nat"})], func = TwoStep nat_sub_expand_equiv_fn} fun nat_sub_expand_unit_fn ctxt item = let val {id, tname, ...} = item val ct = the_single tname val t = Thm.term_of ct in if not (is_nat_sub_form t) then [] else let val expand = nat_sub_expand ctxt (id, ct) fun process_expand (id', (wft, eq_th)) = let val ct = Thm.rhs_of eq_th val ts = norm_minus_ct ct in if length ts = 0 orelse (length ts = 1 andalso snd (the_single ts) = 1) then let val (_, eq') = norm_minus wft val eq = Util.transitive_list [eq_th, eq'] in [(id', to_obj_eq eq)] end else [] end in (maps process_expand expand) |> filter (BoxID.has_incr_id o fst) - |> map (fn (id, th) => Update.thm_update_sc 1 (id, th)) + |> map (fn (id, th) => Auto2_Setup.Update.thm_update_sc 1 (id, th)) end end val nat_sub_expand_unit = {name = "nat_sub_expand_unit", args = [TypedMatch (TY_TERM, @{term_pat "?A::nat"})], func = OneStep nat_sub_expand_unit_fn} val add_nat_sub_proofsteps = fold add_prfstep [ nat_sub_expand_equiv, nat_sub_expand_unit ] end (* NatSub *) val _ = Theory.setup NatSub.add_nat_sub_proofsteps diff --git a/thys/Auto2_HOL/HOL/nat_sub_test.ML b/thys/Auto2_HOL/HOL/nat_sub_test.ML --- a/thys/Auto2_HOL/HOL/nat_sub_test.ML +++ b/thys/Auto2_HOL/HOL/nat_sub_test.ML @@ -1,103 +1,103 @@ (* File: nat_sub_test.ML Author: Bohua Zhan Unit test for nat_sub.ML. *) local val ts = map (fn x => Free (x, natT)) ["a", "b", "c", "m", "n"] val ctxt = fold Util.declare_free_term ts @{context} in fun test_term ctxt f err_str (str1, str2) = let val (t1, t2) = (Proof_Context.read_term_pattern ctxt str1, Proof_Context.read_term_pattern ctxt str2) val t2' = f (Thm.cterm_of ctxt t1) in if t2 aconv t2' then () else let val _ = trace_t ctxt "Input:" t1 val _ = trace_t ctxt "Expected:" t2 val _ = trace_t ctxt "Actual:" t2' in raise Fail err_str end end val test = let val test_data = [ (* No repeated terms. *) ("a", "a"), ("a + b", "a + b"), ("a - b", "a - b"), ("a - b + c", "a + c - b"), ("a - b - c", "a - (b + c)"), ("a - b + (c - d)", "a + c - (b + d)"), ("a - b - (c - d)", "a + d - (b + c)"), ("a + b + c - d - e - f", "a + b + c - (d + e + f)"), ("a - b + c - d + e - f", "a + c + e - (b + d + f)"), (* Numerical constants (on one side only). *) ("0::nat", "0::nat"), ("2::nat", "2::nat"), ("a + 2 + 3", "a + 5"), ("a - b - 2 - 3", "a - (b + 5)"), ("2 + 3 + a", "a + 5"), ("a - 2 - b - 3", "a - (b + 5)"), ("0 + a - 0 - b", "a - b"), (* Cancellation needed. *) ("a - a", "(0::nat)"), ("a - 0", "a"), ("a + b - a", "b"), ("a - (a - a)", "a"), ("a + b - a - c", "b - c"), ("a + b - a - b", "0::nat"), ("a + b + c - a - b", "c"), ("c + (b + (b + a - b) - b)", "a + c"), (* Cancellation of constants needed. *) ("a + 5 - b - 3", "a + 2 - b"), ("a + 3 - b - 5", "a - (b + 2)"), ("a + 5 - b - 5", "a - b"), ("a - 5 - (b - 3)", "a - (b + 2)"), ("a - 3 - (b - 5)", "(a + 2) - b"), ("a + 5 - 3", "a + 2"), ("a - 5 + 3", "a - 2"), ("5 - a - 3", "2 - a"), ("3 - a - 3", "0 - a"), ("(5::nat) - 3", "2::nat"), ("(3::nat) - 3", "0::nat"), (* Monomial *) ("a * 3", "a * 3"), ("a * b + b * a", "a * b * 2"), (* Cancellation between terms *) ("a * 3 + a * 2", "a * 5"), ("a * 3 - a * 2", "a"), ("a * 3 + b - a * 2", "a + b"), ("a * 2 + b - a * 3", "b - a"), ("a * 2 + b * 3 - a * 3 - b * 2", "b - a"), ("a * 3 - a * 2 - a", "0::nat"), (* Distributivity *) ("(a + 2) * b", "b * 2 + a * b"), ("(a + 2) * 2", "a * 2 + 4"), ("(a - 2) * b", "a * b - b * 2"), ("(a - 2) * 2", "a * 2 - 4"), ("(a + 1) * (a - 1)", "a * a - 1"), ("(a + 3) * (a - 2)", "a + a * a - 6"), ("(a - 1) * (a - 1)", "a * a + 1 - a * 2") ] in - map (WfTerm.test_wfconv ctxt NatSub.fheads NatSub.norm_minus "test") test_data @ + map (Auto2_Setup.WfTerm.test_wfconv ctxt NatSub.fheads NatSub.norm_minus "test") test_data @ map (test_term ctxt NatSub.norm_ring_term "test_t") test_data end end diff --git a/thys/Auto2_HOL/HOL/normalize_test.ML b/thys/Auto2_HOL/HOL/normalize_test.ML --- a/thys/Auto2_HOL/HOL/normalize_test.ML +++ b/thys/Auto2_HOL/HOL/normalize_test.ML @@ -1,90 +1,90 @@ (* File: normalize_test.ML Author: Bohua Zhan Unit test for normalizer.ML. *) local val ts = map (fn x => Free (x, boolT)) ["A", "B", "C"] val ctxt = fold Util.declare_free_term ts @{context} in val test_normalize = let fun test (str, strs) = let val t = Syntax.read_term ctxt str val ts = map (Syntax.read_term ctxt) strs val ritem = - Update.thm_to_ritem (Util.assume_thm ctxt (mk_Trueprop t)) - val ritems' = Normalizer.normalize ctxt ritem - val ts' = map (dest_Trueprop o Thm.prop_of o BoxItem.get_thm_raw) + Auto2_Setup.Update.thm_to_ritem (Util.assume_thm ctxt (mk_Trueprop t)) + val ritems' = Auto2_Setup.Normalizer.normalize ctxt ritem + val ts' = map (dest_Trueprop o Thm.prop_of o Auto2_Setup.BoxItem.get_thm_raw) ritems' in if length ts = length ts' andalso eq_set (op aconv) (ts, ts') then () else let val _ = trace_t ctxt "Input:" t val _ = trace_tlist ctxt "Expected:" ts val _ = trace_tlist ctxt "Actual:" ts' in raise Fail "test_normalize" end end val test_data = [ ("A & B & C", ["A", "B", "C"]), ("~ (A | B | C)", ["~ A", "~ B", "~ C"]), ("~ ~ (~ A & (~ ~ B))", ["~ A", "B"]) ] in map test test_data end val test_use_vardefs = let fun test (s1, s2) = let val (t1, t2) = the_pair (Syntax.read_terms ctxt [s1, s2]) val t2 = if fastype_of t2 = propT then t2 else mk_Trueprop t2 val th1 = t1 |> mk_Trueprop |> Thm.cterm_of ctxt |> Thm.assume - |> apply_to_thm (UtilLogic.to_meta_conv ctxt) + |> apply_to_thm (Auto2_Setup.UtilLogic.to_meta_conv ctxt) |> Util.forall_elim_sch - val (_, th2) = Normalizer.meta_use_vardefs th1 + val (_, th2) = Auto2_Setup.Normalizer.meta_use_vardefs th1 in if Thm.prop_of th2 aconv t2 then () else let val _ = trace_t ctxt "Input:" t1 val _ = trace_t ctxt "Expected:" t2 val _ = trace_t ctxt "Actual:" (Thm.prop_of th2) in raise Fail "test_use_vardefs" end end val test_data = [ ("!s. s = f x --> P s", "P (f x)"), ("!s t. s = f x --> t = g x --> P s t", "P (f x) (g x)"), ("!s t. x < y --> s = f x --> t = g y --> P s t", "x < y ==> P (f x) (g y)"), ("!s. ~s = f x | P s", "P (f x)"), ("!s t. ~s = f x | ~t = g x | P s t", "P (f x) (g x)"), ("!s t. ~x < y | ~s = f x | ~t = g y | P s t", "~x < y | P (f x) (g y)"), ("!s. P s | ~s = f x", "P (f x)"), ("!a b. (a,b) = c --> P a b", "P (fst c) (snd c)"), ("!a b. (a,b) ~= c | P a b", "P (fst c) (snd c)"), ("!a b c. (a,(b,c)) = d --> P a b c", "P (fst d) (fst (snd d)) (snd (snd d))"), ("!a b c. ((a,b),c) = d --> P a b c", "P (fst (fst d)) (snd (fst d)) (snd d)") ] in map test test_data end end (* local *) diff --git a/thys/Auto2_HOL/HOL/order.ML b/thys/Auto2_HOL/HOL/order.ML --- a/thys/Auto2_HOL/HOL/order.ML +++ b/thys/Auto2_HOL/HOL/order.ML @@ -1,976 +1,976 @@ (* File: order.ML Author: Bohua Zhan Ordering on natural numbers. We implement something like a decision procedure for difference logic using proof steps. *) val TY_NAT_ORDER = "NAT_ORDER" signature NAT_ORDER = sig val is_order: term -> bool val is_plus_const: term -> bool val is_minus_const: term -> bool val is_standard_ineq: term -> bool val dest_ineq: term -> term * term * int val dest_ineq_th: thm -> term * term * int val fold_double_plus: conv val norm_ineq_th': thm -> thm val convert_const_x: thm -> thm val convert_const_y: thm -> thm val norm_ineq_th: thm -> thm val norm_ineq_minus_th': thm -> thm option val norm_ineq_minus_th: thm -> thm option datatype order_type = LESS | LESS_LMINUS | LESS_LPLUS | LESS_RMINUS | LESS_RPLUS | LE | LE_LMINUS | LE_LPLUS | LE_RMINUS | LE_RPLUS type order_info val to_normal_th: order_type -> thm -> thm val th_to_ritem: thm -> raw_item val th_to_normed_ritems: thm -> raw_item list - val nat_order_normalizer: Normalizer.normalizer + val nat_order_normalizer: Auto2_Setup.Normalizer.normalizer val nat_eq_diff_prfstep: proofstep val get_nat_order_info: box_item -> order_info val nat_order_typed_matcher: item_matcher val transitive: proofstep val transitive_resolve: proofstep val single_resolve: proofstep val single_resolve_zero: proofstep val nat_order_match: term -> box_item -> Proof.context -> id_inst -> id_inst_th list val nat_order_matcher: item_matcher val nat_order_noteq_matcher: item_matcher val nat_order_single_match: term -> box_item -> Proof.context -> id_inst -> id_inst_th list val nat_order_single_matcher: item_matcher val shadow_nat_order_prfstep: proofstep val shadow_nat_order_single: proofstep val string_of_nat_order: Proof.context -> term * term * int * thm -> string val output_nat_order: Proof.context -> term list * thm -> string val shadow_nat_order: Proof.context -> box_id -> term list * cterm list -> bool val add_nat_order_proofsteps: theory -> theory end; structure Nat_Order : NAT_ORDER = struct val is_numc = UtilArith.is_numc val dest_numc = UtilArith.dest_numc val nat0 = Nat_Util.nat0 val nat_less_th = Nat_Util.nat_less_th val nat_fold_conv0_right = Conv.try_conv (rewr_obj_eq @{thm Nat.add_0_right}) val nat_fold_conv0_left = Conv.try_conv (rewr_obj_eq @{thm Nat.plus_nat.add_0}) (* Whether the given term is < or <= on natural numbers. *) fun is_less t = case t of Const (@{const_name less}, _) $ _ $ _ => true | _ => false fun is_less_eq t = case t of Const (@{const_name less_eq}, _) $ _ $ _ => true | _ => false fun is_order t = is_less t orelse is_less_eq t datatype order_type = LESS_LMINUS | LESS_LPLUS | LESS_RMINUS | LESS_RPLUS | LESS | LE_LMINUS | LE_LPLUS | LE_RMINUS | LE_RPLUS | LE (* Check whether t is in the form a + n, where n is a constant. *) fun is_plus_const t = UtilArith.is_plus t andalso UtilArith.is_numc (dest_arg t) (* Check whether t is in the form a - n, where n is a constant. *) fun is_minus_const t = UtilArith.is_minus t andalso UtilArith.is_numc (dest_arg t) (* Check whether t is an inequality in the standard form. This means t is either a + n <= b with n > 0, or a <= b + n with n >= 0. *) fun is_standard_ineq t = if is_less_eq t then let val (lhs, rhs) = Util.dest_binop_args t in fastype_of lhs = natT andalso ((is_plus_const lhs andalso not (is_plus_const rhs) andalso UtilArith.dest_numc (dest_arg lhs) > 0) orelse (is_plus_const rhs andalso not (is_plus_const lhs))) end else false (* Assume t is in the form x + n <= y or x <= y + n, deconstruct into the triple (x, y, d), where d = -n in the first case and n in the second case. *) fun dest_ineq t = let val _ = assert (is_standard_ineq t) "dest_ineq" val (lhs, rhs) = Util.dest_binop_args t in if is_plus_const lhs then (dest_arg1 lhs, rhs, ~ (UtilArith.dest_numc (dest_arg lhs))) else (lhs, dest_arg1 rhs, UtilArith.dest_numc (dest_arg rhs)) end fun dest_ineq_th th = - if UtilLogic.is_Trueprop (Thm.prop_of th) andalso + if Auto2_Setup.UtilLogic.is_Trueprop (Thm.prop_of th) andalso is_standard_ineq (prop_of' th) then dest_ineq (prop_of' th) else let val _ = trace_thm_global "th:" th in raise Fail "dest_ineq_th" end (* In expression x + n < y, fold n. *) val fold_const_left = apply_to_thm' (Conv.arg1_conv (Conv.arg_conv Nat_Util.nat_fold_conv)) (* In expression x < y + n, fold n. *) val fold_const_right = apply_to_thm' (Conv.arg_conv (Conv.arg_conv Nat_Util.nat_fold_conv)) (* In expression (x + c1) + c2, fold c1 + c2. *) fun fold_double_plus ct = let val t = Thm.term_of ct in if is_plus_const t andalso is_plus_const (dest_arg1 t) then Conv.every_conv [rewr_obj_eq @{thm add_ac(1)}, Conv.arg_conv Nat_Util.nat_fold_conv, fold_double_plus] ct else Conv.all_conv ct end (* Normalize inequality to standard form. This function is able to process any theorem of the form a fold_const_right else ([Nat_Util.nat_le_th n2 n1, th] MRS @{thm reduce_le_plus_consts'}) |> fold_const_left end else if is_plus_const lhs then if UtilArith.dest_numc (dest_arg lhs) = 0 then th RS @{thm norm_le_lplus0} else th else if is_plus_const rhs then th else if is_minus_const lhs then th RS @{thm norm_le_lminus} else if is_minus_const rhs then th RS @{thm norm_le_rminus} else th RS @{thm norm_le} else (* is_less t *) if is_plus_const lhs andalso is_plus_const rhs then let val (lhs, rhs) = Util.dest_binop_args t val (n1, n2) = apply2 (UtilArith.dest_numc o dest_arg) (lhs, rhs) in if n1 <= n2 then (th RS @{thm reduce_less_plus_consts}) |> fold_const_right |> apply_to_thm' try_fold0 |> norm_ineq_th' else ([Nat_Util.nat_le_th n2 n1, th] MRS @{thm reduce_less_plus_consts'}) |> fold_const_left |> norm_ineq_th' end else if is_plus_const lhs then (th RS @{thm norm_less_lplus}) |> fold_const_left else if is_plus_const rhs then (th RS @{thm norm_less_rplus}) |> fold_const_right else if is_minus_const lhs then (th RS @{thm norm_less_lminus}) |> fold_const_right else if is_minus_const rhs then (th RS @{thm norm_less_rminus}) |> fold_const_left else th RS @{thm norm_less} end (* Handle the case of x fold_const_left else if d - xn < 0 then (th RS @{thm cv_const4}) |> fold_const_left else (th RS @{thm cv_const5}) |> fold_const_right end else th end (* Handle the case of x = 0 then (th RS @{thm cv_const2}) |> fold_const_right else if d < 0 andalso yn + d < 0 then ([Nat_Util.nat_less_th yn (~d), th] MRS @{thm cv_const3}) |> fold_const_left else (th RS @{thm cv_const6}) |> fold_const_right end else th end (* Overall normalization function. *) fun norm_ineq_th th = let val th' = th |> apply_to_thm' (Conv.arg1_conv fold_double_plus) |> apply_to_thm' (Conv.arg_conv fold_double_plus) |> norm_ineq_th' |> convert_const_x |> convert_const_y in if is_standard_ineq (prop_of' th') then th' else let val _ = trace_thm_global "th':" th' in raise Fail "norm_ineq_th: invalid output." end end (* Normalization function that does not consider the minus case. *) fun norm_ineq_minus_th' th = let val t = prop_of' th val _ = assert (is_less_eq t orelse is_less t) "norm_ineq_minus_th" val (lhs, rhs) = Util.dest_binop_args t in if is_less_eq t then if is_minus_const lhs andalso not (is_plus_const rhs) then SOME (th RS @{thm norm_le}) else if is_minus_const rhs andalso not (is_plus_const lhs) then SOME (th RS @{thm norm_le}) else NONE else (* is_less t *) if is_minus_const lhs andalso not (is_plus_const rhs) then SOME (th RS @{thm norm_less}) else if is_minus_const rhs andalso not (is_plus_const lhs) then SOME (th RS @{thm norm_less}) else NONE end fun norm_ineq_minus_th th = case norm_ineq_minus_th' th of NONE => NONE | SOME th' => let val th'' = th' |> convert_const_x |> convert_const_y in if is_standard_ineq (prop_of' th'') then SOME th'' else let val _ = trace_thm_global "th'':" th'' in raise Fail "norm_ineq_minus_th: invalid output." end end (* The first three values (x, y, n) specify a NAT_ORDER item, and the last value provides theorem justifying it. If n >= 0, then the theorem is x <= y + n. Otherwise, the theorem is x + (~n) <= y. *) type order_info = cterm * cterm * int * thm (* Conversion from a theorem to its normal form. *) fun to_normal_th order_ty th = case order_ty of LESS_LMINUS => (th RS @{thm norm_less_lminus}) |> fold_const_right | LESS_LPLUS => (th RS @{thm norm_less_lplus}) |> fold_const_left | LESS_RMINUS => (th RS @{thm norm_less_rminus}) |> fold_const_left | LESS_RPLUS => (th RS @{thm norm_less_rplus}) |> fold_const_right | LESS => th RS @{thm norm_less} | LE_LMINUS => th RS @{thm norm_le_lminus} | LE_RMINUS => th RS @{thm norm_le_rminus} | LE => th RS @{thm norm_le} | LE_LPLUS => if UtilArith.dest_numc (dest_arg (dest_arg1 (prop_of' th))) = 0 then th RS @{thm norm_le_lplus0} else th | _ => th (* eq_x is a meta equality. Use it to rewrite x in an order info. *) fun rewrite_info_x eq_x (cx, cy, diff, th) = let val _ = assert (Thm.lhs_of eq_x aconvc cx) "rewrite_info_x: invalid equality." val th' = if diff >= 0 then (* rewrite x in x <= y + n. *) th |> apply_to_thm' (Conv.arg1_conv (Conv.rewr_conv eq_x)) else (* rewrite x in x + n <= y. *) th |> apply_to_thm' (Conv.arg1_conv ( Conv.arg1_conv (Conv.rewr_conv eq_x))) in (Thm.rhs_of eq_x, cy, diff, th') end (* eq_y is a meta equality. Use it to rewrite y in an order info. *) fun rewrite_info_y eq_y (cx, cy, diff, th) = let val _ = assert (Thm.lhs_of eq_y aconvc cy) "rewrite_info_y: invalid equality." val th' = if diff >= 0 then (* rewrite y in x <= y + n. *) th |> apply_to_thm' (Conv.arg_conv ( Conv.arg1_conv (Conv.rewr_conv eq_y))) else (* rewrite y in x + n <= y. *) th |> apply_to_thm' (Conv.arg_conv (Conv.rewr_conv eq_y)) in (cx, Thm.rhs_of eq_y, diff, th') end (* Conversion from a normalized order theorem to an raw item. *) fun th_to_ritem th = let val (x, y, n) = dest_ineq_th th in Fact (TY_NAT_ORDER, [x, y, mk_int n], th) end (* Overall normalization function. Include both strategies (with and without considering the minus case). *) fun th_to_normed_ritems th = let val th' = norm_ineq_th th val minus_th = the_list (norm_ineq_minus_th th) in map th_to_ritem (th' :: minus_th) end fun nat_order_normalizer _ ritem = case ritem of Handler _ => [ritem] | Fact (ty_str, _, th) => if ty_str <> TY_PROP then [ritem] else if is_order (prop_of' th) andalso fastype_of (dest_arg (prop_of' th)) = natT then th_to_normed_ritems th else [ritem] (* Given an equality such as x + n = y, produce inequalities x + n <= y and x + n >= y. Applies also when one side of the inequality is a constant. *) val nat_eq_diff_prfstep = - ProofStep.prfstep_custom + Auto2_Setup.ProofStep.prfstep_custom "nat_eq_diff" [WithItem (TY_EQ, @{term_pat "(?x::nat) = ?y"})] (fn ((id, _), _) => fn items => fn _ => let val {prop, ...} = the_single items val (lhs, rhs) = prop |> prop_of' |> Util.dest_binop_args in if is_plus_const lhs orelse is_plus_const rhs orelse is_numc lhs orelse is_numc rhs then let val prop' = to_meta_eq prop val ths = (prop' RS @{thm eq_to_ineqs}) - |> UtilLogic.split_conj_th + |> Auto2_Setup.UtilLogic.split_conj_th in [AddItems {id = id, sc = NONE, raw_items = maps th_to_normed_ritems ths}] end else [] end) (* Obtain quadruple x, y, diff, and theorem from item. *) fun get_nat_order_info {tname, prop, ...} = let val (cx, cy, cdiff_t) = the_triple tname in (cx, cy, dest_numc (Thm.term_of cdiff_t), prop) end (* Matching function for retrieving a NAT_ORDER item. The pattern should be a triple (?x, ?y, ?m). *) val nat_order_typed_matcher = let fun pre_match pat {tname, ...} ctxt = - Matcher.pre_match_head + Auto2_Setup.Matcher.pre_match_head ctxt (pat, Thm.cterm_of ctxt (HOLogic.mk_tuple (map Thm.term_of tname))) fun match pat (item as {tname = cts, ...}) ctxt (id, inst) = let val pats = HOLogic.strip_tuple pat - val insts' = Matcher.rewrite_match_list + val insts' = Auto2_Setup.Matcher.rewrite_match_list ctxt (map (pair false) (pats ~~ cts)) (id, inst) fun process_inst (inst, ths) = let (* eq_x: pat_x(env) == x, eq_y: pat_y(env) = y *) val (eq_x, eq_y, eq_n) = the_triple ths in if Thm.is_reflexive eq_n then let val (_, _, _, prop) = item |> get_nat_order_info |> rewrite_info_x (meta_sym eq_x) |> rewrite_info_y (meta_sym eq_y) in [(inst, prop)] end else [] end in maps process_inst insts' end in {pre_match = pre_match, match = match} end (* Helper function for transitivity of two inequalities. *) fun order_trans th1 th2 = let val (_, y1, d1) = dest_ineq_th th1 val (x2, _, d2) = dest_ineq_th th2 val _ = assert (y1 aconv x2) "order_trans" in if d1 < 0 andalso d2 < 0 then ([th1, th2] MRS @{thm trans1}) |> fold_const_left else if d1 >= 0 andalso d2 >= 0 then ([th1, th2] MRS @{thm trans2}) |> fold_const_right else if d1 < 0 andalso d2 >= 0 then if d2 >= (~d1) then ([th1, th2] MRS @{thm trans3}) |> fold_const_right else (* d2 < (~d1) *) ([Nat_Util.nat_less_th d2 (~d1), th1, th2] MRS @{thm trans4}) |> fold_const_left else (* d1 >= 0 andalso d2 < 0 *) if d1 >= (~d2) then ([th1, th2] MRS @{thm trans5}) |> fold_const_right else ([Nat_Util.nat_less_th d1 (~d2), th1, th2] MRS @{thm trans6}) |> fold_const_left end (* Apply transitivity to two NAT_ORDER infos. *) val transitive = - ProofStep.prfstep_custom + Auto2_Setup.ProofStep.prfstep_custom "nat_order_transitive" [WithItem (TY_NAT_ORDER, @{term_pat "(?x, ?y, ?m)"}), WithItem (TY_NAT_ORDER, @{term_pat "(?y, ?z, ?n)"}), Filter (neq_filter @{term_pat "?x ~= ?y"}), Filter (neq_filter @{term_pat "?y ~= ?z"}), Filter (neq_filter @{term_pat "?x ~= ?z"})] (fn ((id, _), ths) => fn _ => fn _ => let val (th1, th2) = the_pair ths val item' = th_to_ritem (order_trans th1 th2) in [AddItems {id = id, sc = NONE, raw_items = [item']}] end) (* Helper function for obtaining a contradiction from two inequalities. *) fun order_trans_resolve th1 th2 = let val (_, _, d1) = dest_ineq_th th1 val (_, _, d2) = dest_ineq_th th2 in if d1 < 0 andalso d2 < 0 then [nat_less_th 0 (~d1), th1, th2] MRS @{thm trans_resolve1} else if d1 < 0 andalso d2 >= 0 then [nat_less_th d2 (~d1), th1, th2] MRS @{thm trans_resolve2} else (* d1 >= 0 andalso d2 < 0 *) [nat_less_th d1 (~d2), th2, th1] MRS @{thm trans_resolve2} end (* Try to derive a contradiction from two NAT_ORDER items. *) val transitive_resolve = - ProofStep.prfstep_custom + Auto2_Setup.ProofStep.prfstep_custom "nat_order_transitive_resolve" [WithItem (TY_NAT_ORDER, @{term_pat "(?x, ?y, ?m)"}), WithItem (TY_NAT_ORDER, @{term_pat "(?y, ?x, ?n)"})] (fn ((id, inst), ths) => fn _ => fn _ => let val m = dest_numc (lookup_inst inst "m") val n = dest_numc (lookup_inst inst "n") val (th1, th2) = the_pair ths (* Fold x <= y + 0 to x <= y. *) val fold0 = apply_to_thm' (Conv.arg_conv nat_fold_conv0_right) val try_fold0l = Conv.try_conv nat_fold_conv0_left val res_ths = if m + n < 0 then [order_trans_resolve th1 th2] else if m = 0 andalso n = 0 then [(map fold0 [th1, th2]) MRS @{thm Orderings.order_antisym}] else if m + n = 0 then (* Fold 0 at left in case x or y is zero. *) [([th1, th2] MRS @{thm Orderings.order_antisym}) |> apply_to_thm' (Conv.arg_conv try_fold0l) |> apply_to_thm' (Conv.arg1_conv try_fold0l)] else [] in - map (fn th => Update.thm_update (id, th)) res_ths + map (fn th => Auto2_Setup.Update.thm_update (id, th)) res_ths end) (* Try to derive a contradiction from a single NAT_ORDER item. There are two types of resolves: when two sides are equal, and when the right side is zero (both with negative diff). *) val single_resolve = - ProofStep.prfstep_custom + Auto2_Setup.ProofStep.prfstep_custom "nat_order_single_resolve" [WithItem (TY_NAT_ORDER, @{term_pat "(?x, ?x, ?n)"})] (fn ((id, inst), ths) => fn _ => fn _ => let val n = dest_numc (lookup_inst inst "n") val th = the_single ths in if n < 0 then - [Update.thm_update ( + [Auto2_Setup.Update.thm_update ( id, [nat_less_th 0 (~n), th] MRS @{thm single_resolve})] else [] end) val single_resolve_zero = - ProofStep.prfstep_custom + Auto2_Setup.ProofStep.prfstep_custom "nat_order_single_resolve_zero" [WithItem (TY_NAT_ORDER, @{term_pat "(?x, 0, ?n)"})] (fn ((id, inst), ths) => fn _ => fn _ => let val n = dest_numc (lookup_inst inst "n") val th = the_single ths in if n < 0 then - [Update.thm_update (id, [nat_less_th 0 (~n), th] + [Auto2_Setup.Update.thm_update (id, [nat_less_th 0 (~n), th] MRS @{thm single_resolve_const})] else [] end) (* Returns true if ty is either natT or a schematic type of sort linorder. *) fun valid_type ctxt ty = case ty of TVar _ => Sign.of_sort (Proof_Context.theory_of ctxt) (ty, ["Orderings.linorder"]) | _ => (ty = natT) (* Given t in inequality form, return whether the type of the argument is appropriate for nat_order_match: that is, either natT or a schematic type of sort linorder. *) fun valid_arg_type ctxt t = valid_type ctxt (fastype_of (dest_arg t)) fun is_order_pat ctxt pat = (is_order pat andalso valid_arg_type ctxt pat) orelse (is_order (get_neg pat) andalso valid_arg_type ctxt (get_neg pat)) val less_nat = @{term "(<)::nat => nat => bool"} val le_nat = @{term "(<=)::nat => nat => bool"} (* Assuming t is an order. Return the simplified form of the negation of t. *) fun get_neg_order t = if is_neg t then dest_not t else if is_less t then le_nat $ dest_arg t $ dest_arg1 t else if is_less_eq t then less_nat $ dest_arg t $ dest_arg1 t else raise Fail "get_neg_order" fun is_plus_const_gen t = is_plus_const t orelse (is_numc t andalso not (t aconv nat0)) fun plus_const_of_gen t = if is_plus_const t then dest_arg1 t else nat0 fun rewr_side eq_th ct = if Thm.lhs_of eq_th aconvc ct then Conv.rewr_conv eq_th ct else if is_numc (Thm.term_of ct) then Conv.every_conv [rewr_obj_eq (obj_sym @{thm Nat.plus_nat.add_0}), Conv.arg1_conv (Conv.rewr_conv eq_th)] ct else Conv.arg1_conv (Conv.rewr_conv eq_th) ct (* Return a pair of terms to be matched. *) fun analyze_pattern t = let val (lhs, rhs) = Util.dest_binop_args t in if is_plus_const_gen lhs then (plus_const_of_gen lhs, rhs, if is_less t then LESS_LPLUS else LE_LPLUS) else if is_plus_const_gen rhs then (lhs, plus_const_of_gen rhs, if is_less t then LESS_RPLUS else LE_RPLUS) else if is_minus_const lhs then (dest_arg1 lhs, rhs, if is_less t then LESS_LMINUS else LE_LMINUS) else if is_minus_const rhs then (lhs, dest_arg1 rhs, if is_less t then LESS_RMINUS else LE_RMINUS) else (lhs, rhs, if is_less t then LESS else LE) end fun analyze_pattern2 t = let val (lhs, rhs) = Util.dest_binop_args t in if is_minus_const lhs andalso not (is_plus_const_gen rhs) then SOME (lhs, rhs, if is_less t then LESS else LE) else if is_minus_const rhs andalso not (is_plus_const_gen lhs) then SOME (lhs, rhs, if is_less t then LESS else LE) else NONE end (* Matching function. *) fun find_contradiction_item pat item ctxt (id, inst) = let val (x, y, d1, th1) = get_nat_order_info item fun process_pattern (pat_l, pat_r, order_ty) = let val pairs = [(false, (pat_l, y)), (false, (pat_r, x))] in map (pair order_ty) - (Matcher.rewrite_match_list ctxt pairs (id, inst)) + (Auto2_Setup.Matcher.rewrite_match_list ctxt pairs (id, inst)) end val insts = process_pattern (analyze_pattern pat) val insts2 = case analyze_pattern2 pat of NONE => [] | SOME pattern => process_pattern pattern fun process_inst (order_ty, ((id', inst'), eq_ths)) = let val (eq_th1, eq_th2) = the_pair eq_ths val assum = pat |> Util.subst_term_norm inst' |> mk_Trueprop |> Thm.cterm_of ctxt val th2 = assum |> Thm.assume |> apply_to_thm' (Conv.arg1_conv (rewr_side eq_th1)) |> apply_to_thm' (Conv.arg_conv (rewr_side eq_th2)) |> to_normal_th order_ty val (_, _, d2) = dest_ineq_th th2 in if d1 + d2 < 0 then [((id', inst'), (order_trans_resolve th1 th2) |> Thm.implies_intr assum - |> apply_to_thm UtilLogic.rewrite_from_contra_form)] + |> apply_to_thm Auto2_Setup.UtilLogic.rewrite_from_contra_form)] else [] end in maps process_inst (insts @ insts2) end fun nat_order_match pat item ctxt (id, inst) = if not (is_order_pat ctxt pat) then [] else let val is_neg = is_neg pat val neg_pat = get_neg_order pat val res = find_contradiction_item neg_pat item ctxt (id, inst) in if is_neg then res else res |> map (apsnd (apply_to_thm' UtilArith.neg_ineq_cv)) end fun nat_order_pre_match pat item ctxt = if not (is_order_pat ctxt pat) then false else let val neg_pat = get_neg_order pat val (x, y, _, _) = get_nat_order_info item fun process_pattern (pat_l, pat_r, _) = - Matcher.pre_match ctxt (pat_l, y) andalso - Matcher.pre_match ctxt (pat_r, x) + Auto2_Setup.Matcher.pre_match ctxt (pat_l, y) andalso + Auto2_Setup.Matcher.pre_match ctxt (pat_r, x) val match1 = process_pattern (analyze_pattern neg_pat) val match2 = case analyze_pattern2 neg_pat of NONE => false | SOME pattern => process_pattern pattern in match1 orelse match2 end (* Prop-matching with a NAT_ORDER item. *) val nat_order_matcher = {pre_match = nat_order_pre_match, match = nat_order_match} (* Use x < y to match x ~= y. *) val nat_order_noteq_matcher = let fun get_insts pat item ctxt (id, inst) = if not (is_neg pat andalso is_eq_term (dest_not pat)) then ([], []) else let val (A, B) = pat |> dest_not |> dest_eq val (cx, cy, n, _) = get_nat_order_info item in if n >= 0 then ([], []) - else (Matcher.rewrite_match_list + else (Auto2_Setup.Matcher.rewrite_match_list ctxt [(false, (A, cx)), (false, (B, cy))] (id, inst), - Matcher.rewrite_match_list + Auto2_Setup.Matcher.rewrite_match_list ctxt [(false, (B, cx)), (false, (A, cy))] (id, inst)) end fun pre_match pat item ctxt = if not (is_neg pat andalso is_eq_term (dest_not pat)) then false else let val (A, B) = pat |> dest_not |> dest_eq val (cx, cy, n, _) = get_nat_order_info item in - n < 0 andalso ((Matcher.pre_match ctxt (A, cx) andalso - Matcher.pre_match ctxt (B, cy)) orelse - (Matcher.pre_match ctxt (A, cy) andalso - Matcher.pre_match ctxt (A, cx))) + n < 0 andalso ((Auto2_Setup.Matcher.pre_match ctxt (A, cx) andalso + Auto2_Setup.Matcher.pre_match ctxt (B, cy)) orelse + (Auto2_Setup.Matcher.pre_match ctxt (A, cy) andalso + Auto2_Setup.Matcher.pre_match ctxt (A, cx))) end fun match pat item ctxt (id, inst) = if Util.has_vars pat then [] else let val (instAB, instBA) = get_insts pat item ctxt (id, inst) fun process_inst (inst', ths) = let (* eq_x: (A/B) = x, eq_y: (B/A) = y *) val (eq_x, eq_y) = the_pair ths val (_, _, n, prop) = item |> get_nat_order_info |> rewrite_info_x (meta_sym eq_x) |> rewrite_info_y (meta_sym eq_y) in (inst', [prop, nat_less_th 0 (~n)] MRS @{thm nat_ineq_impl_not_eq}) end fun switch_ineq (inst', th) = (inst', th RS @{thm HOL.not_sym}) in map process_inst instAB @ (map (switch_ineq o process_inst) instBA) end in {pre_match = pre_match, match = match} end (* Given pattern pat, find substitutions of pat that give rise to a contradiction. *) fun find_contradiction pat ctxt (id, inst) = let fun process_pattern (pat_l, pat_r, order_ty) = map (fn (id', eq_th) => (order_ty, (id', eq_th, Thm.reflexive (Thm.cterm_of ctxt pat_r)))) - (RewriteTable.equiv_info_t ctxt id (pat_l, pat_r)) + (Auto2_Setup.RewriteTable.equiv_info_t ctxt id (pat_l, pat_r)) val insts = process_pattern (analyze_pattern pat) val insts2 = case analyze_pattern2 pat of NONE => [] | SOME pattern => process_pattern pattern fun process_inst (order_ty, (id', eq_th1, eq_th2)) = let val assum = pat |> Util.subst_term_norm inst |> mk_Trueprop |> Thm.cterm_of ctxt val norm_th = assum |> Thm.assume |> apply_to_thm' (Conv.arg1_conv (rewr_side eq_th1)) |> apply_to_thm' (Conv.arg_conv (rewr_side eq_th2)) |> to_normal_th order_ty val (lhs, rhs) = Util.dest_binop_args (prop_of' norm_th) in if is_plus_const lhs andalso dest_arg1 lhs aconv rhs then let val d = dest_numc (dest_arg lhs) in if d = 0 then [] else [((id', inst), ([nat_less_th 0 d, norm_th] MRS @{thm single_resolve}) |> Thm.implies_intr assum - |> apply_to_thm UtilLogic.rewrite_from_contra_form)] + |> apply_to_thm Auto2_Setup.UtilLogic.rewrite_from_contra_form)] end else [] end in maps process_inst (insts @ insts2) end (* For patterns in the form m < n, where m and n are constants. *) fun find_contradiction_trivial pat ctxt (id, inst) = let val (A, B) = Util.dest_binop_args pat in if not (is_numc A andalso is_numc B) then [] else if not (is_less pat andalso dest_numc A >= dest_numc B) andalso not (is_less_eq pat andalso dest_numc A > dest_numc B) then [] else let val assum = pat |> mk_Trueprop |> Thm.cterm_of ctxt in [((id, inst), (UtilArith.contra_by_arith ctxt [Thm.assume assum]) |> Thm.implies_intr assum - |> apply_to_thm UtilLogic.rewrite_from_contra_form)] + |> apply_to_thm Auto2_Setup.UtilLogic.rewrite_from_contra_form)] end end (* For patterns in the form x < 0. *) fun find_contradiction_trivial2 pat ctxt (id, inst) = let val (_, B) = Util.dest_binop_args pat in if is_less pat andalso B aconv nat0 then let val assum = pat |> mk_Trueprop |> Thm.cterm_of ctxt in [((id, inst), - ([@{thm Nat.not_less0}, Thm.assume assum] MRS UtilBase.contra_triv_th) + ([@{thm Nat.not_less0}, Thm.assume assum] MRS Auto2_UtilBase.contra_triv_th) |> Thm.implies_intr assum - |> apply_to_thm UtilLogic.rewrite_from_contra_form)] + |> apply_to_thm Auto2_Setup.UtilLogic.rewrite_from_contra_form)] end else [] end (* Using term x to justify propositions of the form x <= x + n, where n >= 0. This follows the same general outline as nat_order_match. *) fun nat_order_single_match pat _ ctxt (id, inst) = if not (is_order_pat ctxt pat) orelse Util.has_vars pat then [] else let val is_neg = is_neg pat val neg_pat = get_neg_order pat val (A, B) = Util.dest_binop_args neg_pat val res = if is_less neg_pat andalso B aconv nat0 then find_contradiction_trivial2 neg_pat ctxt (id, inst) else if is_numc A andalso is_numc B then find_contradiction_trivial neg_pat ctxt (id, inst) else find_contradiction neg_pat ctxt (id, inst) in if is_neg then res else res |> map (apsnd (apply_to_thm' UtilArith.neg_ineq_cv)) end val nat_order_single_matcher = {pre_match = (fn pat => fn _ => fn ctxt => is_order_pat ctxt pat), match = nat_order_single_match} (* Shadow the second item if it is looser than the first (same x and y, but larger diff. *) val shadow_nat_order_prfstep = - ProofStep.gen_prfstep + Auto2_Setup.ProofStep.gen_prfstep "shadow_nat_order" [WithItem (TY_NAT_ORDER, @{term_pat "(?x, ?y, ?m)"}), WithItem (TY_NAT_ORDER, @{term_pat "(?x, ?y, ?n)"}), Filter (fn _ => fn (_, inst) => dest_numc (lookup_inst inst "m") <= dest_numc (lookup_inst inst "n")), ShadowSecond] (* Shadow the given item if it is trivial. There are two cases: when two sides are equal, and when the left side is zero (both with nonnegative diff). *) val shadow_nat_order_single = - ProofStep.gen_prfstep + Auto2_Setup.ProofStep.gen_prfstep "shadow_nat_order_single" [WithItem (TY_NAT_ORDER, @{term_pat "(?x, ?x, ?n)"}), Filter (fn _ => fn (_, inst) => dest_numc (lookup_inst inst "n") >= 0), ShadowFirst] val shadow_nat_order_single_zero = - ProofStep.gen_prfstep + Auto2_Setup.ProofStep.gen_prfstep "shadow_nat_order_single_zero" [WithItem (TY_NAT_ORDER, @{term_pat "(0, ?x, ?n)"}), Filter (fn _ => fn (_, inst) => dest_numc (lookup_inst inst "n") >= 0), ShadowFirst] fun string_of_nat_order ctxt (x, y, diff, th) = if x aconv nat0 andalso diff < 0 then (* 0 + n <= y *) th |> apply_to_thm' (Conv.arg1_conv nat_fold_conv0_left) |> Thm.prop_of |> Syntax.string_of_term ctxt else if y aconv nat0 andalso diff >= 0 then (* x <= 0 + n *) th |> apply_to_thm' (Conv.arg_conv nat_fold_conv0_left) |> Thm.prop_of |> Syntax.string_of_term ctxt else if diff = 0 then (* x <= y + 0 *) th |> apply_to_thm' (Conv.arg_conv nat_fold_conv0_right) |> Thm.prop_of |> Syntax.string_of_term ctxt else th |> Thm.prop_of |> Syntax.string_of_term ctxt fun output_nat_order ctxt (ts, th) = let val (x, y, diff_t) = the_triple ts val diff = dest_numc diff_t in "ORDER " ^ string_of_nat_order ctxt (x, y, diff, th) end fun shadow_nat_order _ _ (ts1, cts2) = let val (x1, y1, m) = the_triple ts1 val (cx2, cy2, n) = the_triple cts2 val m = dest_numc m val n = dest_numc (Thm.term_of n) in if m < n then false else x1 aconv Thm.term_of cx2 andalso y1 aconv Thm.term_of cy2 end val add_nat_order_proofsteps = - ItemIO.add_item_type ( + Auto2_Setup.ItemIO.add_item_type ( TY_NAT_ORDER, SOME (take 2), SOME output_nat_order, SOME shadow_nat_order - ) #> ItemIO.add_typed_matcher ( + ) #> Auto2_Setup.ItemIO.add_typed_matcher ( TY_NAT_ORDER, nat_order_typed_matcher - ) #> fold ItemIO.add_prop_matcher [ + ) #> fold Auto2_Setup.ItemIO.add_prop_matcher [ (TY_NAT_ORDER, nat_order_matcher), (TY_NAT_ORDER, nat_order_noteq_matcher), (TY_NULL, nat_order_single_matcher) - ] #> Normalizer.add_normalizer ( + ] #> Auto2_Setup.Normalizer.add_normalizer ( ("nat_order", nat_order_normalizer) ) #> fold add_prfstep [ nat_eq_diff_prfstep, transitive, transitive_resolve, single_resolve, single_resolve_zero, shadow_nat_order_prfstep, shadow_nat_order_single, shadow_nat_order_single_zero ] end val _ = Theory.setup (Nat_Order.add_nat_order_proofsteps) diff --git a/thys/Auto2_HOL/HOL/order_test.ML b/thys/Auto2_HOL/HOL/order_test.ML --- a/thys/Auto2_HOL/HOL/order_test.ML +++ b/thys/Auto2_HOL/HOL/order_test.ML @@ -1,223 +1,223 @@ (* File: order_test.ML Author: Bohua Zhan Unit test for order.ML. *) local val ctxt = @{context} val thy = @{theory} (* Set up rewrite table and incremental context. *) val x = Free ("x", natT) val x' = Free ("x'", natT) val y = Free ("y", natT) val y' = Free ("y'", natT) val z = Free ("z", natT) val w = Free ("w", natT) val px = Var (("x", 0), natT) val py = Var (("y", 0), natT) val ctxt' = ctxt |> fold Variable.declare_term [x, x', y, y', z, w, px, py] - |> RewriteTable.add_rewrite ([], assume_eq thy (x, x')) |> snd + |> Auto2_Setup.RewriteTable.add_rewrite ([], assume_eq thy (x, x')) |> snd in val test_fold_double_plus = let val err_str = "test_fold_double_plus" val test_data = [ ("x", "x"), ("x + 1", "x + 1"), ("(x + 1) + 1", "x + 2"), ("((x + 1) + 1) + 1", "x + 3") ] in map (Util.test_conv ctxt' Nat_Order.fold_double_plus err_str) test_data end (* Convert proposition to nat_order item, by applying one of the proofsteps in to_nat_order_prfsteps. *) fun convert_prop_to_nat_order prop = let val th = prop |> mk_Trueprop |> Util.assume_thm ctxt' in Nat_Order.th_to_normed_ritems th end val test_parse_prop = let fun string_of_spec (x, y, diff) = "(" ^ (Util.string_of_terms ctxt [x, y, mk_nat diff]) ^ ")" fun ritem_to_spec ritem = case ritem of Fact ("NAT_ORDER", tname, _) => let val (x, y, diff_t) = the_triple tname in (x, y, UtilArith.dest_numc diff_t) end | _ => raise Fail "ritem_to_spec" fun eq_specs ((x1, y1, diff1), (x2, y2, diff2)) = x1 aconv x2 andalso y1 aconv y2 andalso diff1 = diff2 fun test (prop_str, specs_str) = let val prop = Syntax.read_term ctxt' prop_str fun read_spec (x_str, y_str, diff) = (Syntax.read_term ctxt' x_str, Syntax.read_term ctxt' y_str, diff) val specs = map read_spec specs_str val specs' = convert_prop_to_nat_order prop |> map ritem_to_spec in if eq_set eq_specs (specs, specs') then () else let val _ = tracing ( "Expected:" ^ (Util.string_of_list string_of_spec specs) ^ "\n" ^ "Actual:" ^ (Util.string_of_list string_of_spec specs')) in raise Fail "test_parse_prop" end end val test_data = [ ("x <= y", [("x", "y", 0)]), ("x + 1 <= y", [("x", "y", ~1)]), ("x - 1 <= y", [("x", "y", 1), ("x-1", "y", 0)]), ("x <= y + 1", [("x", "y", 1)]), ("x <= y - 1", [("x", "y", 0), ("x", "y-1", 0)]), (* special case *) ("x < y", [("x", "y", ~1)]), ("x + 1 < y", [("x", "y", ~2)]), ("x - 1 < y", [("x", "y", 0), ("x-1", "y", ~1)]), ("x < y + 1", [("x", "y", 0)]), ("x < y - 1", [("x", "y", ~2), ("x", "y-1", ~1)]), ("x <= 3", [("x", "0::nat", 3)]), ("x < 3", [("x", "0::nat", 2)]), ("x <= 0", [("x", "0::nat", 0)]), ("x < 0", [("x", "0::nat", ~1)]), ("x >= 3", [("0::nat", "x", ~3)]), ("x > 3", [("0::nat", "x", ~4)]), ("x >= 0", [("0::nat", "x", 0)]), ("x > 0", [("0::nat", "x", ~1)]), ("x + 0 <= y", [("x", "y", 0)]), ("x + 1 <= 3", [("x", "0::nat", 2)]), ("x - 1 <= 3", [("x", "0::nat", 4), ("x-1", "0::nat", 3)]), ("(x + 1) + 1 <= y + 1", [("x", "y", ~1)]), ("x + 1 <= (y + 1) + 1", [("x", "y", 1)]), ("(x + 1) + 1 <= (y + 1) + 1", [("x", "y", 0)]) ] val _ = map test test_data in () end val test_nat_order_match = let fun test match_fn (pat_str, prop_str, res_strs) = let val pat = Proof_Context.read_term_pattern ctxt' pat_str val prop = Syntax.read_term ctxt' prop_str val ritems = convert_prop_to_nat_order prop val ritem = the_single ritems handle List.Empty => raise Fail "ambiguous input" - val item = BoxItem.mk_box_item ctxt (0, [], 0, ritem) + val item = Auto2_Setup.BoxItem.mk_box_item ctxt (0, [], 0, ritem) val res = map (mk_Trueprop o Syntax.read_term ctxt') res_strs val res' = match_fn pat item ctxt' ([], fo_init) fun check_inst ((_, inst), th) = (Util.subst_term_norm inst pat) aconv (prop_of' th) val check_res = eq_list (op aconv) (res, map (Thm.prop_of o snd) res') in if forall check_inst res' andalso check_res then () else let val _ = trace_tlist ctxt' "pat, prop:" [pat, prop] val _ = trace_tlist ctxt' "Expected:" res val _ = trace_tlist ctxt' "Actual:" (map (Thm.prop_of o snd) res') in raise Fail "test_nat_order_match" end end val test_data = [ ("x <= y", "x <= y", ["x <= y"]), ("x < y", "x <= y", []), ("x < y", "x < y", ["x < y"]), ("x <= y", "x < y", ["x <= y"]), ("x <= y", "x' <= y", ["x <= y"]), ("x' <= y", "x <= y", ["x' <= y"]), ("x + 1 <= y", "x <= y", []), ("x - 1 <= y", "x <= y", ["x - 1 <= y"]), ("x <= y + 1", "x <= y", ["x <= y + 1"]), ("min ?a ?b <= z", "min x y <= z", ["min x y <= z"]), ("min ?a ?b + 1 <= z", "min x y <= z", []), ("min ?a ?b - 1 <= z", "min x y <= z", ["min x y - 1 <= z"]), ("min ?a ?b <= z + 1", "min x y <= z", ["min x y <= z + 1"]), ("?x <= y", "x < y", ["x <= y"]), ("?x <= ?y", "x < y", ["x <= y"]), ("?x < ?y + 1", "x < y", ["x < y + 1"]), ("?x < ?y - 1", "x + 1 < y", ["x < y - 1"]), ("0 < ?x", "x > 3", ["0 < x"]), ("0 < ?x", "x >= 0", []), ("0 <= ?x", "x >= 0", ["0 <= x"]), ("?x <= 0", "x < 1", ["x <= 0"]), ("?x <= 0", "x < 2", []), ("~ x < y", "x >= y", ["~ x < y"]), ("~ x < 3", "x > 3", ["~ x < 3"]), ("x <= 3", "x < 3", ["x <= 3"]), ("x <= 3", "x < 4", ["x <= 3"]), ("~ ?x < ?y", "x >= y", ["~ x < y"]), ("~ x <= y", "x > y", ["~ x <= y"]), ("~ min (?a::nat) ?b <= min ?c ?d", "min x y < min z w", ["~ min z w <= min x y"]) ] val _ = map (test Nat_Order.nat_order_match) test_data val test_noteq_data = [ ("x ~= y", "x < y", ["x ~= y"]), ("x ~= y", "y < x", ["x ~= y"]), ("x ~= y", "x <= y", []), ("x ~= y", "y <= x", []), ("x ~= 0", "x > 0", ["x ~= 0"]), ("x ~= 0", "x > 3", ["x ~= 0"]), ("?a ~= y", "x < y", []) ] val _ = map (test (#match Nat_Order.nat_order_noteq_matcher)) test_noteq_data in () end val test_nat_order_single_match = let fun test ineq_str = let val ineq = Syntax.read_term ctxt' ineq_str val res = Nat_Order.nat_order_single_match - ineq BoxItem.null_item ctxt' ([], fo_init) + ineq Auto2_Setup.BoxItem.null_item ctxt' ([], fo_init) in if length res = 1 andalso res |> the_single |> snd |> prop_of' aconv ineq then () else let val _ = trace_t ctxt' "ineq:" ineq val _ = trace_tlist ctxt' "Output:" (map (Thm.prop_of o snd) res) in raise Fail "test_nat_order_single_match" end end val test_data = [ "x >= x", "x + 1 > x", "x >= x'", "x + 1 > x'", "x' + 1 > x", "x >= x - 1", "x >= x' - 1", "x' >= x - 1", "~ x > x", "~ x < x", "(3::nat) < 4", "(3::nat) <= 3" ] val _ = map test test_data in () end end diff --git a/thys/Auto2_HOL/HOL/rewrite_test.ML b/thys/Auto2_HOL/HOL/rewrite_test.ML --- a/thys/Auto2_HOL/HOL/rewrite_test.ML +++ b/thys/Auto2_HOL/HOL/rewrite_test.ML @@ -1,406 +1,406 @@ (* File: rewrite_test.ML Author: Bohua Zhan Unit test for rewrite.ML. *) local val init = @{context} fun add_prim_id (prim_id, id) ctxt = let val (prim_id', ctxt') = BoxID.add_prim_id id ctxt val _ = assert (prim_id = prim_id') "add_prim_id" in ctxt' end (* Conversion between (id, th) used in rewrite.ML and (id, t) used in test cases. *) fun to_term_info (id, th) = (id, Util.rhs_of th) (* Comparison of list of box_ids. *) val eq_id_list = eq_set (op =) (* Comparison of list of (box_id, term) pairs. *) val eq_info_list = eq_set (eq_pair (op =) (op aconv)) fun print_term_info ctxt (id, t) = "(" ^ (BoxID.string_of_box_id id) ^ ", " ^ (Syntax.string_of_term ctxt t) ^ ")" fun print_term_infos ctxt lst = commas (map (print_term_info ctxt) lst) val string_of_box_id_list = Util.string_of_list BoxID.string_of_box_id fun done str _ = let val _ = tracing ("Finished " ^ str) in () end (* info is expected, in box_id * term form, info' is returned value, in box_id * thm form (to be printed). *) fun assert_eq_info (info, info') txt ctxt = if eq_info_list (info, map to_term_info info') then ctxt else let val _ = tracing ("Expected: " ^ print_term_infos ctxt info ^ "\n" ^ "Actual: " ^ print_infos ctxt info') in raise Fail txt end fun assert_rew_info str info_str ctxt = let val t = Syntax.read_term ctxt str val ct = Thm.cterm_of ctxt t val info = map (apsnd (Syntax.read_term ctxt)) info_str in - assert_eq_info (info, RewriteTable.get_rewrite_info ctxt ct) + assert_eq_info (info, Auto2_Setup.RewriteTable.get_rewrite_info ctxt ct) "assert_rew_info" ctxt end fun assert_srew_info str info_str ctxt = let val t = Syntax.read_term ctxt str val ct = Thm.cterm_of ctxt t val info = map (apsnd (Syntax.read_term ctxt)) info_str in - assert_eq_info (info, RewriteTable.get_subterm_rewrite_info ctxt ct) + assert_eq_info (info, Auto2_Setup.RewriteTable.get_subterm_rewrite_info ctxt ct) "assert_srew_info" ctxt end fun assert_head_rep (id, str) info_str ctxt = let val t = Syntax.read_term ctxt str val ct = Thm.cterm_of ctxt t val info = map (apsnd (Syntax.read_term ctxt)) info_str val th = Thm.reflexive ct in - assert_eq_info (info, RewriteTable.get_head_rep_with_id_th ctxt (id, th)) + assert_eq_info (info, Auto2_Setup.RewriteTable.get_head_rep_with_id_th ctxt (id, th)) "assert_head_rep" ctxt end fun assert_equiv_gen exp_equiv id (str1, str2) ctxt = let val (t1, t2) = (Syntax.read_term ctxt str1, Syntax.read_term ctxt str2) in - if exp_equiv = RewriteTable.is_equiv_t id ctxt (t1, t2) then ctxt + if exp_equiv = Auto2_Setup.RewriteTable.is_equiv_t id ctxt (t1, t2) then ctxt else let val _ = tracing ("id: " ^ (BoxID.string_of_box_id id)) val _ = trace_tlist ctxt "Input:" [t1, t2] in raise Fail (if exp_equiv then "assert_equiv" else "assert_not_equiv") end end val assert_equiv = assert_equiv_gen true val assert_not_equiv = assert_equiv_gen false fun assert_simpl_info str info_str ctxt = let val t = Syntax.read_term ctxt str val ct = Thm.cterm_of ctxt t val info = map (apsnd (Syntax.read_term ctxt)) info_str in - assert_eq_info (info, RewriteTable.simplify_info ctxt ct) + assert_eq_info (info, Auto2_Setup.RewriteTable.simplify_info ctxt ct) "assert_simpl_info" ctxt end fun assert_simpl id (str, str') ctxt = let val (t, t') = (Syntax.read_term ctxt str, Syntax.read_term ctxt str') val ct = Thm.cterm_of ctxt t - val res = Util.rhs_of (RewriteTable.simplify id ctxt ct) + val res = Util.rhs_of (Auto2_Setup.RewriteTable.simplify id ctxt ct) in if res aconv t' then ctxt else let val _ = trace_t ctxt "Input:" t val _ = trace_t ctxt "Expected:" t' val _ = trace_t ctxt "Actual:" res in raise Fail "assert_simpl" end end (* Collect all head equivs under box id. The return value of get_head_equiv is in the form [(head, [((id, th), groups), ...]), ...], and we want to collect the set of (id, th). *) fun assert_head_equivs (id, str) info_str ctxt = let val cu = Thm.cterm_of ctxt (Syntax.read_term ctxt str) val info = map (apsnd (Syntax.read_term ctxt)) info_str - val res = (RewriteTable.get_head_equiv ctxt cu) + val res = (Auto2_Setup.RewriteTable.get_head_equiv ctxt cu) |> BoxID.merge_box_with_info ctxt id in assert_eq_info (info, res) "assert_head_equivs" ctxt end fun index_reps ctxt = let fun index_reps_t t = - fold RewriteTable.update_subsimp - (RewriteTable.get_subterm_rewrite_info ctxt t) + fold Auto2_Setup.RewriteTable.update_subsimp + (Auto2_Setup.RewriteTable.get_subterm_rewrite_info ctxt t) in - fold index_reps_t (RewriteTable.get_all_terms ctxt) ctxt + fold index_reps_t (Auto2_Setup.RewriteTable.get_all_terms ctxt) ctxt end fun declare_term str ctxt = Proof_Context.augment (Syntax.read_term ctxt str) ctxt (* First part: test internal functions. *) (* Modification functions using terms. *) fun add_equiv id (str1, str2) ctxt = let val (t1, t2) = (Syntax.read_term ctxt str1, Syntax.read_term ctxt str2) val thy = Proof_Context.theory_of ctxt val eq_th = Util.assume_meta_eq thy (t1, t2) in - ctxt |> RewriteTable.add_equiv (id, eq_th) + ctxt |> Auto2_Setup.RewriteTable.add_equiv (id, eq_th) end fun add_term str ctxt = - ctxt |> RewriteTable.add_term ([], Thm.cterm_of ctxt (Syntax.read_term ctxt str)) |> snd + ctxt |> Auto2_Setup.RewriteTable.add_term ([], Thm.cterm_of ctxt (Syntax.read_term ctxt str)) |> snd fun assert_eq_edges (str1, str2) ids ctxt = let val (t1, t2) = (Syntax.read_term ctxt str1, Syntax.read_term ctxt str2) - val ids' = (RewriteTable.equiv_neighs ctxt t1) + val ids' = (Auto2_Setup.RewriteTable.equiv_neighs ctxt t1) |> map to_term_info |> filter (fn (_, t2') => t2 aconv t2') |> map fst in if eq_id_list (ids, ids') then ctxt else let val _ = trace_tlist ctxt "Input:" [t1, t2] val _ = tracing ("Expected: " ^ (string_of_box_id_list ids) ^ "\n" ^ "Actual: " ^ (string_of_box_id_list ids')) in raise Fail "assert_eq_edges" end end fun update_simp id (str1, str2) ctxt = let val (t1, t2) = (Syntax.read_term ctxt str1, Syntax.read_term ctxt str2) val thy = Proof_Context.theory_of ctxt val eq_th = Util.assume_meta_eq thy (t1, t2) in - ctxt |> RewriteTable.update_simp (id, eq_th) + ctxt |> Auto2_Setup.RewriteTable.update_simp (id, eq_th) end fun assert_rewrite id (str, str') ctxt = let val (t, t') = (Syntax.read_term ctxt str, Syntax.read_term ctxt str') - val res = Util.rhs_of (RewriteTable.get_rewrite id ctxt (Thm.cterm_of ctxt t)) + val res = Util.rhs_of (Auto2_Setup.RewriteTable.get_rewrite id ctxt (Thm.cterm_of ctxt t)) in if res aconv t' then ctxt else let val _ = trace_t ctxt "Input:" t val _ = trace_t ctxt "Expected:" t' val _ = trace_t ctxt "Actual:" res in raise Fail "assert_rewrite" end end fun add_rewrite id (str1, str2) ctxt = let val (t1, t2) = (Syntax.read_term ctxt str1, Syntax.read_term ctxt str2) val thy = Proof_Context.theory_of ctxt val th = assume_eq thy (t1, t2) in - ctxt |> RewriteTable.add_rewrite (id, th) |> snd + ctxt |> Auto2_Setup.RewriteTable.add_rewrite (id, th) |> snd end in val test_equiv_basic = init |> fold add_prim_id [(1, []), (2, [1]), (3, [1]), (4, [2])] |> declare_term "[a, b, c, d]" |> fold add_term ["a", "b", "c", "d"] |> add_equiv [2] ("a", "b") |> assert_eq_edges ("a", "b") [[2]] |> add_equiv [3] ("b", "c") |> assert_eq_edges ("b", "c") [[3]] |> add_equiv [] ("a", "b") |> assert_eq_edges ("a", "b") [[]] |> add_equiv [4] ("a", "d") |> assert_eq_edges ("a", "d") [[4]] |> add_equiv [3] ("a", "d") |> assert_eq_edges ("a", "d") [[3], [4]] |> add_equiv [2] ("a", "d") |> assert_eq_edges ("a", "d") [[2], [3]] |> add_equiv [] ("a", "d") |> assert_eq_edges ("a", "d") [[]] |> done "test_equiv_basic" val test_simp_basic = init |> fold add_prim_id [(1, []), (2, [1]), (3, [1]), (4, [2])] |> declare_term "[a, b, c, d, e]" |> fold add_term ["a", "b", "c", "d", "e"] |> update_simp [2] ("e", "d") |> assert_rewrite [2, 3] ("e", "d") |> update_simp [3] ("e", "c") |> assert_rewrite [3] ("e", "c") |> update_simp [3] ("e", "b") |> assert_rewrite [2, 3] ("e", "b") |> update_simp [4] ("e", "a") |> assert_rewrite [2, 3] ("e", "b") |> update_simp [2] ("e", "b") |> update_simp [2, 3] ("e", "a") |> assert_rew_info "e" [([], "e"), ([4], "a"), ([3], "b"), ([2], "b"), ([2, 3], "a")] |> done "test_simp_basic" val test_head_rep = init |> fold add_prim_id [(1, []), (2, [1]), (3, [1])] |> fold declare_term ["[a, b, c, d]", "f a b"] |> add_term "f b d" |> update_simp [2] ("d", "c") |> update_simp [3] ("b", "a") |> assert_srew_info "f b d" [([], "f b d"), ([2], "f b c"), ([3], "f a d"), ([2, 3], "f a c")] |> index_reps |> assert_head_rep ([], "f a c") [([2, 3], "f b d")] |> assert_head_rep ([], "f b c") [([2], "f b d")] |> assert_head_rep ([], "f a d") [([3], "f b d")] |> assert_head_rep ([], "f b d") [([], "f b d")] |> assert_head_rep ([3], "f b c") [([2, 3], "f b d")] |> done "test_head_rep" val test_head_rep3 = init |> fold add_prim_id [(1, []), (2, [1]), (3, [1])] |> fold declare_term ["[a, b, c, d]", "f a b"] |> add_term "f b c" |> add_term "f a d" |> update_simp [2] ("d", "c") |> update_simp [3] ("b", "a") |> index_reps |> assert_head_rep ([], "f a c") [([2], "f a d"), ([3], "f b c")] |> assert_head_rep ([2], "f a c") [([2], "f a d"), ([2, 3], "f b c")] |> assert_head_rep ([3], "f a c") [([3], "f b c"), ([2, 3], "f a d")] |> done "test_head_rep3" val test_add_term = init |> fold add_prim_id [(1, []), (2, [1]), (3, [1])] |> fold declare_term ["[a, b, c, d]", "f a b"] |> add_rewrite [2] ("d", "c") |> add_rewrite [3] ("b", "a") |> index_reps |> add_term "f b c" |> add_term "f a d" |> assert_head_rep ([], "f a c") [([2], "f a d"), ([3], "f b c")] |> assert_eq_edges ("f a d", "f b c") [[2, 3]] |> add_term "f a c" |> add_term "f b d" |> assert_head_rep ([], "f a c") [([], "f a c"), ([2], "f a d"), ([3], "f b c"), ([2, 3], "f b d")] |> assert_equiv [2, 3] ("f a c", "f b d") |> assert_not_equiv [2] ("f a c", "f b d") |> assert_not_equiv [3] ("f a c", "f b d") |> done "test_add_term" val test_add_rewrite = init |> fold add_prim_id [(1, []), (2, [1]), (3, [1])] |> fold declare_term ["[a, b, c, d]", "f a b"] |> add_term "f a d" |> add_rewrite [2] ("d", "c") |> assert_head_rep ([], "f a c") [([2], "f a d")] |> add_term "f b d" |> assert_head_rep ([], "f b c") [([2], "f b d")] |> add_rewrite [3] ("a", "b") |> assert_eq_edges ("f a d", "f b d") [[3]] |> assert_rew_info "f b d" [([], "f b d"), ([2], "f b c"), ([3], "f a d"), ([2, 3], "f a c")] |> add_term "f b c" |> assert_head_rep ([], "f a c") [([2], "f a d"), ([3], "f b c"), ([2, 3], "f b d")] |> done "test_add_rewrite" val test_simplify = init |> fold add_prim_id [(1, []), (2, [1]), (3, [1])] |> fold declare_term ["[a, b, c, d]", "[f a b, e]"] |> add_rewrite [1] ("f b d", "e") |> add_rewrite [2] ("d", "c") |> add_rewrite [3] ("b", "a") |> assert_simpl_info "f a c" [([], "f a c"), ([2, 3], "e")] |> assert_simpl_info "f a d" [([], "f a d"), ([2], "f a c"), ([3], "e")] |> assert_simpl_info "f b c" [([], "f b c"), ([3], "f a c"), ([2], "e")] |> assert_simpl_info "f b d" [([], "f b d"), ([1], "e")] |> assert_simpl [2] ("f a c", "f a c") |> assert_simpl [2, 3] ("f a c", "e") |> assert_simpl [2] ("f a d", "f a c") |> assert_simpl [3] ("f a d", "e") |> assert_simpl [2, 3] ("f a d", "e") |> add_rewrite [] ("f b d", "e") |> assert_simpl_info "f b d" [([], "e")] |> add_term "f a c" |> assert_equiv [2, 3] ("f a c", "f b d") |> done "test_simplify" val test_simplify2 = init |> fold add_prim_id [(1, []), (2, [1]), (3, [1]), (4, [2])] |> declare_term "[a, b, c, d]" |> add_rewrite [3] ("b", "c") |> add_rewrite [4] ("a", "b") |> add_rewrite [2] ("c", "d") |> assert_simpl_info "c" [([], "c"), ([3], "b"), ([3, 4], "a")] |> assert_simpl_info "d" [([], "d"), ([2], "c"), ([2, 3], "b"), ([3, 4], "a")] |> done "test_simplify2" val test_simplify3 = init |> fold add_prim_id [(1, []), (2, [1])] |> declare_term "[a, b, c, d, e, g]" |> add_rewrite [1] ("e", "c") |> add_rewrite [1] ("g", "d") |> add_rewrite [2] ("e", "b") |> add_rewrite [2] ("g", "a") |> add_rewrite [] ("e", "g") |> assert_simpl_info "e" [([], "e"), ([1], "c"), ([2], "a")] |> assert_simpl_info "g" [([], "e"), ([1], "c"), ([2], "a")] |> done "test_simplify3" val test_head_equivs = init |> fold add_prim_id [(1, []), (2, [1]), (3, [1])] |> fold declare_term ["[a, b, c, d]", "[f a b, e]"] |> add_rewrite [1] ("f b d", "e") |> add_rewrite [2] ("d", "c") |> add_rewrite [3] ("b", "a") |> assert_head_equivs ([], "c") [([], "c"), ([2], "d")] |> assert_head_equivs ([], "f b d") [([], "f b d"), ([1], "e")] |> assert_head_equivs ([2], "f b c") [([2], "f b c"), ([2], "e"), ([2], "f b d")] |> assert_head_equivs ([], "f b c") [([], "f b c"), ([2], "e"), ([2], "f b d")] |> assert_head_equivs ([], "f a c") [([], "f a c"), ([2, 3], "e"), ([2, 3], "f b d")] |> add_term "f a c" |> assert_head_equivs ([], "f b d") [([], "f b d"), ([1], "e"), ([2, 3], "f a c")] |> assert_head_equivs ([], "f a c") [([], "f a c"), ([2, 3], "e"), ([2, 3], "f b d")] |> assert_head_equivs ([], "f b c") [([], "f b c"), ([2], "e"), ([2], "f b d"), ([3], "f a c")] |> done "test_head_equivs" val test_head_equivs2 = init |> fold add_prim_id [(1, []), (2, [1]), (3, [1]), (4, [2])] |> declare_term "[a, b, c, d]" |> add_rewrite [3] ("b", "c") |> add_rewrite [4] ("a", "b") |> assert_head_equivs ([], "b") [([], "b"), ([3], "c"), ([4], "a")] |> add_rewrite [2] ("c", "d") |> assert_head_equivs ([], "b") [([], "b"), ([3], "c"), ([4], "a"), ([2, 3], "d")] |> done "test_head_equivs2" val test_head_equivs3 = init |> fold add_prim_id [(1, []), (2, [1]), (3, [1])] |> fold declare_term ["[a, b, c, d]", "[f a b, e]"] |> add_rewrite [] ("b", "a") |> add_rewrite [2] ("f b d", "e") |> add_rewrite [3] ("f a d", "e") (* Both fbd and fad are ok in the result. *) |> assert_head_equivs ([], "f a d") [([], "f b d"), ([2], "e"), ([3], "e"), ([], "f a d")] |> assert_head_equivs ([], "f b d") [([], "f a d"), ([2], "e"), ([3], "e"), ([], "f b d")] |> done "test_head_equivs3" val test_head_equivs4 = init |> fold add_prim_id [(1, []), (2, [1]), (3, [1])] |> fold declare_term ["[a, b, c, d]", "[f a b, e]"] |> add_rewrite [2] ("b", "a") |> add_rewrite [] ("f b c", "e") |> add_rewrite [3] ("f a c", "e") (* Note fbc ~ fac under 3. *) |> assert_head_equivs ([], "f b c") [([], "f b c"), ([], "e"), ([3], "f a c"), ([2], "f a c")] |> assert_head_equivs ([], "f a c") [([], "f a c"), ([2], "e"), ([3], "e"), ([3], "f b c"), ([2], "f b c")] |> done "test_head_equivs4" (* Simplify with nat variables. *) val test_simplify_nat = init |> declare_term "[a::nat, b, c]" |> add_rewrite [] ("a + b", "b + a") |> add_rewrite [] ("a", "c") |> assert_simpl [] ("b + c", "a + b") |> assert_equiv [] ("b + c", "a + b") |> done "test_simplify_nat" end; (* local *) diff --git a/thys/Auto2_HOL/HOL/unfolding.ML b/thys/Auto2_HOL/HOL/unfolding.ML --- a/thys/Auto2_HOL/HOL/unfolding.ML +++ b/thys/Auto2_HOL/HOL/unfolding.ML @@ -1,67 +1,76 @@ (* File: unfolding.ML Author: Bohua Zhan Unfolding of functional definitions. *) signature UNFOLDING = sig val get_unfold_thms_by_name: theory -> string -> thm list val get_unfold_thms: theory -> term -> thm list val unfold: theory -> conv val unfold_cmd: string -> Proof.state -> Proof.state end; -structure Unfolding : UNFOLDING = +signature UNFOLDING_KEYWORD = +sig + val unfold : string * Position.T +end + +functor Unfolding( + structure Auto2_Outer: AUTO2_OUTER; + structure Basic_UtilLogic: BASIC_UTIL_LOGIC; + structure Unfolding_Keyword: UNFOLDING_KEYWORD; + ): UNFOLDING = struct fun get_unfold_thms_by_name thy nm = let val simp_nm = nm ^ ".simps" val def_nm = nm ^ "_def" in Global_Theory.get_thms thy simp_nm handle ERROR _ => Global_Theory.get_thms thy def_nm handle ERROR _ => raise Fail "get_unfold_thms" end fun get_unfold_thms thy t = get_unfold_thms_by_name thy (Util.get_head_name t) (* Unfold the given term. *) fun unfold thy ct = let val ths = get_unfold_thms thy (Thm.term_of ct) in - Conv.first_conv (map rewr_obj_eq ths) ct + Conv.first_conv (map Basic_UtilLogic.rewr_obj_eq ths) ct end fun unfold_cmd s state = let val {context = ctxt, ...} = Proof.goal state val thy = Proof_Context.theory_of ctxt val (_, (As, _)) = ctxt |> Auto2_State.get_subgoal |> Util.strip_meta_horn val cAs = map (Thm.cterm_of ctxt) As val t = Syntax.read_term ctxt s val eq_th = t |> Thm.cterm_of ctxt |> unfold thy - |> to_obj_eq + |> Basic_UtilLogic.to_obj_eq |> fold Thm.implies_intr (rev cAs) val _ = writeln ("Obtained " ^ (eq_th |> Thm.concl_of |> Syntax.string_of_term ctxt)) val after_qed = Auto2_Outer.have_after_qed ctxt eq_th in state |> Proof.map_contexts (Auto2_State.map_head_th after_qed) end val _ = - Outer_Syntax.command @{command_keyword "@unfold"} "unfold a term" + Outer_Syntax.command Unfolding_Keyword.unfold "unfold a term" (Parse.term >> (fn s => Toplevel.proof (fn state => unfold_cmd s state))) end (* structure Unfolding *) diff --git a/thys/Auto2_HOL/HOL/util_arith.ML b/thys/Auto2_HOL/HOL/util_arith.ML --- a/thys/Auto2_HOL/HOL/util_arith.ML +++ b/thys/Auto2_HOL/HOL/util_arith.ML @@ -1,146 +1,146 @@ (* File: util_arith.ML Author: Bohua Zhan Utility functions related to arithmetic. *) signature UTIL_ARITH = sig (* Types. *) val natT: typ val intT: typ val ratT: typ val rat_zero: Rat.rat (* Terms. *) val is_numc: term -> bool val dest_numc: term -> int val dest_numc_rat: term -> Rat.rat val is_order: term -> bool val is_linorder: Proof.context -> term -> bool val is_plus: term -> bool val is_minus: term -> bool val is_times: term -> bool val is_divide: term -> bool val is_zero: term -> bool val is_one: term -> bool (* Theorems. *) val neg_ineq_cv: conv val neg_ineq_back_cv: conv (* Arith tactic. *) val prove_by_arith: Proof.context -> thm list -> term -> thm val contra_by_arith: Proof.context -> thm list -> thm end; structure UtilArith : UTIL_ARITH = struct val natT = HOLogic.natT val intT = @{typ int} val ratT = @{typ rat} val rat_zero = Rat.of_int 0 (* Test if a term represents a numerical constant. In addition to use dest_number from HOLogic, test for inverse, uminus, of_rat, etc. *) fun is_numc t = case t of Const (@{const_name inverse}, _) $ t' => is_numc t' | Const (@{const_name uminus}, _) $ t' => is_numc t' | Const (@{const_name of_rat}, _) $ r => is_numc r | Const (@{const_name Fract}, _) $ n $ d => is_numc n andalso is_numc d | _ => let val _ = HOLogic.dest_number t in true end handle TERM ("dest_number", _) => false (* Deconstruct numerical constant. Discard type. *) fun dest_numc t = HOLogic.dest_number t |> snd handle TERM ("dest_number", _) => raise Fail "dest_numc" (* Rational numbers version of dest_numc. *) fun dest_numc_rat t = case t of Const (@{const_name inverse}, _) $ t' => let val r' = dest_numc_rat t' in if r' = rat_zero then rat_zero else Rat.inv r' end | Const (@{const_name uminus}, _) $ t' => Rat.neg (dest_numc_rat t') | Const (@{const_name of_rat}, _) $ r => dest_numc_rat r | Const (@{const_name Fract}, _) $ n $ d => Rat.make (dest_numc n, dest_numc d) | _ => Rat.of_int (dest_numc t) (* Whether the given term is a < b or a <= b. *) fun is_order t = let val _ = assert (fastype_of t = boolT) "is_order: wrong type" in case t of Const (@{const_name less}, _) $ _ $ _ => true | Const (@{const_name less_eq}, _) $ _ $ _ => true | _ => false end fun is_linorder ctxt t = let val T = fastype_of (dest_arg t) val thy = Proof_Context.theory_of ctxt in is_order t andalso Sign.of_sort thy (T, ["Orderings.linorder"]) end (* Check whether t is in the form a + b. *) fun is_plus t = case t of Const (@{const_name plus}, _) $ _ $ _ => true | _ => false (* Check whether t is in the form a - b. *) fun is_minus t = case t of Const (@{const_name minus}, _) $ _ $ _ => true | _ => false fun is_times t = case t of Const (@{const_name times}, _) $ _ $ _ => true | _ => false fun is_divide t = case t of Const (@{const_name divide}, _) $ _ $ _ => true | _ => false fun is_zero t = case t of Const (@{const_name zero_class.zero}, _) => true | _ => false fun is_one t = case t of Const (@{const_name one_class.one}, _) => true | _ => false (* Convert ~ x < y to y <= x, and ~ x <= y to y < x. *) val neg_ineq_cv = (Conv.try_conv o Conv.first_conv) (map rewr_obj_eq [@{thm Orderings.linorder_not_less}, @{thm Orderings.linorder_not_le}]) (* Convert x < y to ~ y <= x, and x <= y to ~ y < x. *) val neg_ineq_back_cv = (Conv.try_conv o Conv.first_conv) (map (rewr_obj_eq o obj_sym) [@{thm Orderings.linorder_not_less}, @{thm Orderings.linorder_not_le}]) -val prove_by_arith = UtilLogic.prove_by_tac Arith_Data.arith_tac -val contra_by_arith = UtilLogic.contra_by_tac Arith_Data.arith_tac +val prove_by_arith = Auto2_Setup.UtilLogic.prove_by_tac Arith_Data.arith_tac +val contra_by_arith = Auto2_Setup.UtilLogic.contra_by_tac Arith_Data.arith_tac end (* structure UtilArith *) val natT = UtilArith.natT val intT = UtilArith.intT diff --git a/thys/Auto2_HOL/HOL/util_test.ML b/thys/Auto2_HOL/HOL/util_test.ML --- a/thys/Auto2_HOL/HOL/util_test.ML +++ b/thys/Auto2_HOL/HOL/util_test.ML @@ -1,61 +1,61 @@ (* File: util_test.ML Author: Bohua Zhan Unit test for util.ML. *) local val ctxt = @{context} in val test_normalize_meta_all_imp = let val test_data = [ ("!!x. (A ==> B x)", "A ==> (!!x. B x)") ] in map (Util.test_conv ctxt (Util.normalize_meta_all_imp ctxt) "normalize_meta_all_imp") test_data end val test_to_obj_conv = let fun err n = "test_to_obj_conv: " ^ (string_of_int n) fun assert_eq th ct txt = let val ct' = Thm.rhs_of th in if ct' aconvc ct then () else raise Fail txt end val ct1 = @{cprop "A ==> B ==> (!!(n::nat). C n) ==> D"} val ct2 = @{cprop "A --> B --> (!(n::nat). C n) --> D"} val ct3 = @{cprop "!(y::nat) (x::nat). P x y --> (!z. Q x y z)"} - val _ = assert_eq (UtilLogic.to_obj_conv ctxt ct1) ct2 (err 0) - val _ = assert_eq (UtilLogic.to_meta_conv ctxt ct2) + val _ = assert_eq (Auto2_Setup.UtilLogic.to_obj_conv ctxt ct1) ct2 (err 0) + val _ = assert_eq (Auto2_Setup.UtilLogic.to_meta_conv ctxt ct2) @{cprop "A ==> B ==> (!(n::nat). C n) ==> D"} (err 1) - val _ = assert_eq (UtilLogic.to_meta_conv ctxt ct3) + val _ = assert_eq (Auto2_Setup.UtilLogic.to_meta_conv ctxt ct3) @{cprop "!!(y::nat) (x::nat) z. P x y ==> Q x y z"} (err 2) - val _ = assert_eq (UtilLogic.to_obj_conv_on_horn ctxt ct1) + val _ = assert_eq (Auto2_Setup.UtilLogic.to_obj_conv_on_horn ctxt ct1) @{cprop "A ==> B ==> (!(n::nat). C n) ==> D"} (err 4) in () end val test_is_pattern = let fun test b str = let val t = Proof_Context.read_term_pattern ctxt str in if b = Util.is_pattern t then () else raise Fail "test_is_pattern" end val test_positive = ["?f", "!n. ?f n", "!m n. ?f m n", "!n. ?f n < ?f (n + 1)", "!n. ?f (n + 1) < ?f n", "!n. ?g (?f n) & ?g n"] val test_negative = ["?f ?n", "!n. ?f n n", "!n. ?f (?f n)", "!n. (?f n < ?g (n + 1)) & (?f (n + 1) < ?g n)"] in map (test true) test_positive @ map (test false) test_negative end end diff --git a/thys/Auto2_HOL/ROOT b/thys/Auto2_HOL/ROOT --- a/thys/Auto2_HOL/ROOT +++ b/thys/Auto2_HOL/ROOT @@ -1,21 +1,21 @@ chapter AFP session Auto2_HOL (AFP) = HOL + description \ Instantiation of Auto2 for Isabelle/HOL. \ options [timeout = 600] sessions "HOL-Library" directories "HOL" theories [document = false] - (* Core setup *) + (* Core setup tests *) "HOL/Auto2_Test" theories (* Simple examples *) "HOL/Pelletier" "HOL/Primes_Ex" document_files "root.tex" "root.bib" diff --git a/thys/Auto2_HOL/auto2.ML b/thys/Auto2_HOL/auto2.ML --- a/thys/Auto2_HOL/auto2.ML +++ b/thys/Auto2_HOL/auto2.ML @@ -1,607 +1,626 @@ (* Title: auto2.ML Author: Bohua Zhan Main file defining inference algorithm and tactic. *) signature SCORES = sig val item_score: raw_item -> int val items_score: raw_item list -> int val get_score: raw_update -> int end; structure Scores : SCORES = struct fun item_score ritem = case ritem of Handler _ => 0 | Fact (ty_str, tname, _) => if ty_str = TY_VAR then 10 (* cost of introducing a variable. *) else Integer.sum (map size_of_term tname) fun items_score raw_items = Util.max int_ord (map item_score raw_items) fun get_score raw_updt = case raw_updt of AddItems {raw_items, sc, ...} => (case sc of NONE => items_score raw_items | SOME n => n) | AddBoxes {init_assum, sc, ...} => (case sc of NONE => 20 + 5 * size_of_term init_assum | SOME n => n) | ResolveBox _ => ~1 | ShadowItem _ => ~1 end (* structure SCORES. *) (* Flags specifying output options. *) val print_trace = Attrib.setup_config_bool @{binding "print_trace"} (K false) val print_intended = Attrib.setup_config_bool @{binding "print_intended"} (K false) val print_term = Attrib.setup_config_bool @{binding "print_term"} (K false) val print_shadow = Attrib.setup_config_bool @{binding "print_shadow"} (K false) val print_score = Attrib.setup_config_bool @{binding "print_score"} (K false) (* Flag specifying the maximum number of steps. *) val max_steps = Attrib.setup_config_int @{binding "max_steps"} (K 2000) (* Proof status. Manages changes to proof status, and main loop carrying out proof steps and adding new proof steps to the queue. *) signature PROOFSTATUS = sig val check_hyps: box_id -> thm -> status -> unit val scoring: proofstep -> int -> box_item list -> status -> (box_item -> raw_update list) -> update list val process_shadow: update -> status -> status val process_resolve: update -> status -> status val apply_update_instant: update -> status -> status val process_fact_all: box_id -> int -> box_item list -> status -> status val process_add_items: update -> status -> status val process_add_boxes: update -> status -> status val apply_update: update -> status -> status val init_status: Proof.context -> term -> status val solve_root: int * status -> int * status end; -structure ProofStatus : PROOFSTATUS = +functor ProofStatus( + structure Auto2Data: AUTO2_DATA; + structure BoxItem: BOXITEM; + structure ItemIO: ITEM_IO; + structure Normalizer: NORMALIZER; + structure Logic_ProofSteps: LOGIC_PROOFSTEPS; + structure Property: PROPERTY; + structure PropertyData: PROPERTY_DATA; + structure ProofStepData: PROOFSTEP_DATA; + structure RewriteTable: REWRITE_TABLE; + structure Status: STATUS; + structure UtilBase: UTIL_BASE; + structure UtilLogic: UTIL_LOGIC; + structure Update: UPDATE; + structure WellformData: WELLFORM_DATA; + ): PROOFSTATUS = struct fun check_hyps id th (st as {ctxt, ...}) = let val hyps = Thm.hyps_of th val inits = Status.get_init_assums st id val handlers = (Status.get_handlers st) |> filter (fn (id', _) => BoxID.is_eq_ancestor ctxt id' id) |> map (fn (_, (_, t, _)) => t) val extra = subtract (op aconv) (inits @ handlers) hyps in if null extra then () else let val _ = trace_tlist ctxt "extra:" extra in raise Fail "illegal hyp" end end (* Perform preliminary checks before matching the last item of a proofstep, and process the resulting updates. *) fun scoring {name, ...} sc items (st as {ctxt, ...}) func = let val merged_id = BoxItem.merged_id ctxt items in if BoxID.is_box_resolved ctxt merged_id orelse exists (Status.query_shadowed st merged_id) items then [] else let val scs = map #sc items val max_sc = fold (curry Int.max) scs sc val raw_updts = (func (List.last items)) |> map Update.replace_id_of_update fun process_raw_updt raw_updt = let val id = Update.target_of_update raw_updt in (* Perform resolved and shadowed tests again since id may be descendent of merged_id. *) if BoxID.is_box_resolved ctxt id then [] else if exists (Status.query_shadowed st id) items then [] else [{sc = Scores.get_score raw_updt + max_sc, prfstep_name = name, source = items, raw_updt = raw_updt}] end in maps process_raw_updt raw_updts end end fun process_shadow (updt as {raw_updt, ...}) (st as {ctxt, ...}) = case raw_updt of ShadowItem {id = shadow_id, item as {ty_str, ...}} => let val _ = if not (Config.get ctxt print_trace) orelse not (Config.get ctxt print_shadow) then () else if ty_str = TY_TERM andalso not (Config.get ctxt print_term) then () else tracing ("Shadowing " ^ ItemIO.string_of_item ctxt item ^ " at box " ^ BoxID.string_of_box_id shadow_id ^ " (" ^ Update.source_info updt ^ ")") in st |> Status.add_shadowed (shadow_id, item) end | _ => raise Fail "process_shadow: wrong type of update" (* When a box, whether primitive or composite, is resolved, perform the following two actions: 1. Resolve current and all descendent boxes. 2. Add the appropriate fact to each of the immediate parent boxes. *) fun process_resolve (updt as {sc, raw_updt, ...}) (st as {ctxt, ...}) = case raw_updt of ResolveBox {id, th} => let val _ = if Config.get ctxt print_trace then tracing ("Finished box " ^ BoxID.string_of_box_id id ^ " (" ^ Update.source_info updt ^ ")") else () fun update_one i st = let val cur_parent = BoxID.get_parent_at_i ctxt id i val res_th = Status.get_on_resolve st id i th val _ = check_hyps cur_parent res_th st val res_obj_th = apply_to_thm (UtilLogic.to_obj_conv ctxt) res_th val updt = {sc = sc, prfstep_name = "$RESOLVE", source = [], raw_updt = Logic_ProofSteps.logic_thm_update ctxt (cur_parent, res_obj_th)} in st |> apply_update_instant updt end in if BoxID.is_box_resolved ctxt id then st else if id = [BoxID.home_id] then st |> Status.set_resolve_th (Status.get_on_resolve st id 0 th) |> Status.add_resolved id else st |> fold update_one (0 upto (length id) - 1) |> Status.add_resolved id end | _ => raise Fail "process_resolve: wrong type of update" (* Directly apply Shadow and Resolve updates. Put the remaining updates in queue. *) and apply_update_instant (updt as {raw_updt, ...}) (st as {ctxt, ...}) = let val id = Update.target_of_update raw_updt val _ = assert (not (BoxID.has_incr_id id)) in if BoxID.is_box_resolved ctxt id then st else case raw_updt of ResolveBox _ => process_resolve updt st | ShadowItem _ => process_shadow updt st | _ => Status.add_to_queue updt st end fun process_step_single sc items prfstep (st as {ctxt, ...}) = let val {args, func, ...} = prfstep val arg = the_single (filter_out ItemIO.is_side_match args) val items' = filter (ItemIO.pre_match_arg ctxt arg) items val f = case func of OneStep f => f | TwoStep _ => raise Fail "process_step_single: wrong type of func." fun process_one_item item st = let val updts = scoring prfstep sc [item] st (f ctxt) in st |> fold apply_update_instant updts end in st |> fold process_one_item items' end fun process_step_pair sc items items' prfstep (st as {ctxt, ...}) = let val {args, func, ...} = prfstep val (arg1, arg2) = the_pair (filter_out ItemIO.is_side_match args) val f = case func of TwoStep f => f | _ => raise Fail "process_step_pair: wrong type of func." (* Filter the two lists of items using pre_match_arg. *) fun filter_pairs (left, right) = if length left < length right then let val left' = filter (ItemIO.pre_match_arg ctxt arg1) left in if null left' then ([], []) else (left', filter (ItemIO.pre_match_arg ctxt arg2) right) end else let val right' = filter (ItemIO.pre_match_arg ctxt arg2) right in if null right' then ([], []) else (filter (ItemIO.pre_match_arg ctxt arg1) left, right') end fun process_pairs (left, right) st = let val (left', right') = filter_pairs (left, right) (* One step in the iteration, fixing both left and right items. Do not match item with itself. *) fun process_pair left_item func right_item st = if BoxItem.eq_item (left_item, right_item) then st else let val updts = scoring prfstep sc [left_item, right_item] st func in st |> fold apply_update_instant updts end (* Iterate over the right items, fixing left item. *) fun loop_right left_item st = let val func = f ctxt left_item in st |> fold (process_pair left_item func) right' end in st |> fold loop_right left' end in (* We know items is a subset of items'. On the second round, match with the extra elements in items' on the left. *) st |> process_pairs (items, items') |> process_pairs (subtract BoxItem.eq_item items items', items) end fun process_prfstep sc items items' prfstep (st as {ctxt, ...}) = let val {args, ...} = prfstep val items = filter_out ( fn item => BoxID.is_box_resolved ctxt (#id item) orelse Status.query_removed st item) items in if null items then st else if length (filter_out ItemIO.is_side_match args) = 1 then st |> process_step_single sc items prfstep else st |> process_step_pair sc items items' prfstep end (* List of terms to be added to the rewrite table. *) fun get_rewr_terms ctxt item = let val {ty_str, tname, ...} = item val ts = map Thm.term_of tname val rewr_terms = ItemIO.rewr_terms_of_item ctxt (ty_str, ts) val terms = rewr_terms |> maps UtilLogic.get_all_subterms |> distinct (op aconv) |> filter_out Util.has_vars val headt = if ty_str = TY_PROP then - if is_neg (the_single ts) then - [dest_not (the_single ts)] + if UtilLogic.is_neg (the_single ts) then + [UtilLogic.dest_not (the_single ts)] else ts else [] in headt @ terms end (* List of terms to be added as TERM items. *) fun get_rewr_terms2 ctxt item = let val {ty_str, tname, ...} = item val ts = map Thm.term_of tname val rewr_terms = ItemIO.rewr_terms_of_item ctxt (ty_str, ts) in rewr_terms |> maps UtilLogic.get_all_subterms_skip_if |> distinct (op aconv) |> filter_out Util.has_vars - |> filter_out (fn t => fastype_of t = boolT) + |> filter_out (fn t => fastype_of t = UtilBase.boolT) end fun process_add_terms id sc items (st as {ctxt, ...}) = let (* Add terms to the rewrite table. *) val term_infos = items |> maps (get_rewr_terms ctxt) |> distinct (op aconv) |> map (pair id o Thm.cterm_of ctxt) val old_items = (Status.get_items st) |> filter_out (fn {id, ...} => BoxID.is_box_resolved ctxt id) |> subtract BoxItem.eq_item items val ctxt' = ctxt |> Auto2Data.add_terms old_items term_infos (* Add new terms as updates. *) fun exists_item t = - Status.find_ritem_exact st id (Fact (TY_TERM, [t], true_th)) + Status.find_ritem_exact st id (Fact (TY_TERM, [t], UtilBase.true_th)) val terms2 = items |> maps (get_rewr_terms2 ctxt) |> distinct (op aconv) |> filter_out exists_item - |> filter_out (fn t => fastype_of t = boolT) + |> filter_out (fn t => fastype_of t = UtilBase.boolT) (* New terms have score of the source item plus 1. *) val updt = {sc = sc, prfstep_name = "$TERM", source = [], raw_updt = AddItems {id = id, sc = NONE, raw_items = map BoxItem.term_to_fact terms2}} in st |> (if length terms2 > 0 then Status.add_to_queue updt else I) |> Status.map_context (K ctxt') end (* Process the given items. *) fun process_fact_all id sc items (st as {ctxt, ...}) = if null items then st else let val thy = Proof_Context.theory_of ctxt val items' = (Status.get_items st) |> filter_out (fn {id, ...} => BoxID.is_box_resolved ctxt id) val old_items = subtract BoxItem.eq_item items items' (* Incremental context and the list of relevant terms for the increment. *) val ctxt' = Auto2Data.get_incr_type old_items items ctxt val ts = (maps Auto2Data.relevant_terms_single items @ - maps (Property.strip_property_field thy o dest_arg o prop_of' o snd) + maps (Property.strip_property_field thy o dest_arg o UtilLogic.prop_of' o snd) (PropertyData.get_new_property ctxt') @ map fst (WellformData.get_new_wellform_data ctxt')) |> distinct (op aconv) |> RewriteTable.get_reachable_terms true ctxt (* List of items to consider for init and incr rounds. *) val init_items = filter_out (BoxItem.match_ty_str TY_PROPERTY) items fun incr_filt {id = id', tname, ...} = exists (Util.has_subterm ts) (map Thm.term_of tname) andalso BoxID.is_box_unresolved ctxt (BoxID.merge_boxes ctxt (id, id')) val incr_items = if null ts then [] else old_items |> filter incr_filt |> cons BoxItem.null_item (* List of proofsteps to consider for init and incr rounds. *) val thy = Proof_Context.theory_of ctxt val prfsteps = ProofStepData.get_prfsteps thy val all_items = init_items @ incr_items in st |> Status.map_context (K ctxt') |> fold (process_prfstep sc all_items items') prfsteps |> Status.map_context Auto2Data.get_single_type |> Status.clear_incr end fun process_add_items (updt as {sc, prfstep_name, raw_updt, ...}) (st as {ctxt, ...}) = case raw_updt of AddItems {id, raw_items = ritems, ...} => let val (ctxt', subst) = BoxItem.obtain_variant_frees (ctxt, ritems) val ritems' = ritems |> map (BoxItem.instantiate subst) |> (if prfstep_name = "$RESOLVE" orelse prfstep_name = "$INIT_BOX" then maps (Normalizer.normalize_keep ctxt') else maps (Normalizer.normalize ctxt')) |> distinct BoxItem.eq_ritem val (dup_ritems, new_ritems) = if sc = 0 then ([], ritems') else filter_split (Status.find_ritem_exact st id) ritems' val _ = if null dup_ritems orelse not (Config.get ctxt print_trace) orelse not (Config.get ctxt print_intended) then () else if prfstep_name = "$TERM" andalso not (Config.get ctxt print_term) then () else tracing ("Intend to add " ^ Update.update_info ctxt' id dup_ritems ^ " (" ^ Update.source_info updt ^ ")") val (handlers, ritems') = filter_split BoxItem.is_handler_raw new_ritems val uid_incr = length ritems' val uid_next = Status.get_num_items st val uid_string = if uid_incr = 1 then string_of_int uid_next else string_of_int uid_next ^ "-" ^ string_of_int (uid_next + uid_incr - 1) val sc_string = if Config.get ctxt print_score then (string_of_int sc) ^ ", " else "" val _ = if null new_ritems orelse not (Config.get ctxt print_trace) then () else if prfstep_name = "$TERM" andalso not (Config.get ctxt print_term) then () else tracing ("Add " ^ Update.update_info ctxt' id ritems' ^ " (" ^ uid_string ^ ", " ^ sc_string ^ Update.source_info updt ^ ")") (* Produce the actual items. *) val items = map (fn i => BoxItem.mk_box_item ctxt' (uid_next + i, id, sc, nth ritems' i)) (0 upto (length ritems' - 1)) val handlers_info = map (pair id o BoxItem.dest_handler_raw) handlers in st |> Status.map_context (K ctxt') |> fold Status.add_handler handlers_info |> process_add_terms id sc items |> fold Status.add_item (map BoxItem.item_with_incr items) |> process_fact_all id sc (map BoxItem.item_with_incr items) end | _ => raise Fail "process_add_items: wrong type of update" fun process_add_boxes (updt as {sc, raw_updt, ...}) (st as {ctxt, ...}) = case raw_updt of AddBoxes {id, init_assum, ...} => let val ritem = init_assum |> Thm.cterm_of ctxt |> Thm.assume |> Update.thm_to_ritem val _ = assert (fastype_of init_assum = propT) "process_add_boxes: assumption is not of type prop." (* Find neg_form and check if already present. *) val neg_form_opt = if id = [] then NONE else - init_assum |> dest_Trueprop |> get_neg + init_assum |> UtilLogic.dest_Trueprop |> UtilLogic.get_neg |> Status.find_fact st id val prev_prim_box = Status.find_prim_box st id init_assum in (* Do nothing if there is already a box with the same assumptions and conclusions. *) if is_some prev_prim_box then st else if is_some neg_form_opt then let val th = the neg_form_opt val updt = {sc = sc, prfstep_name = "$RESOLVE", source = [], raw_updt = Logic_ProofSteps.logic_thm_update ctxt (id, th)} in st |> apply_update_instant updt end (* Otherwise, proceed to create the box. *) else let val _ = if Config.get ctxt print_trace then tracing ("Add box under " ^ BoxID.string_of_box_id id ^ " (" ^ Update.source_info updt ^ ")") else () val (prim_id, st') = st |> Status.map_context (K ctxt) |> Status.add_prim_box id init_assum val new_id = [prim_id] val _ = if Config.get ctxt print_trace then tracing (BoxID.string_of_box_id new_id ^ ": " ^ Syntax.string_of_term ctxt init_assum) else () val ritems' = case ritem of Handler _ => [ritem] | Fact (ty_str, _, th) => - if ty_str = TY_PROP andalso is_neg (prop_of' th) then + if ty_str = TY_PROP andalso UtilLogic.is_neg (UtilLogic.prop_of' th) then map Update.thm_to_ritem ( Logic_ProofSteps.split_not_imp_th th) else [ritem] val item_updt = { sc = sc, prfstep_name = "$INIT_BOX", source = [], raw_updt = AddItems {id = new_id, sc = NONE, raw_items = ritems'}} in st' |> process_add_items item_updt end end | _ => raise Fail "process_add_boxes: wrong type of update" fun apply_update (updt as {raw_updt, ...}) (st as {ctxt, ...}) = let val id = Update.target_of_update raw_updt val _ = assert (not (BoxID.has_incr_id id)) in if BoxID.is_box_resolved ctxt id then st else case raw_updt of AddItems _ => process_add_items updt st | AddBoxes _ => process_add_boxes updt st | ResolveBox _ => process_resolve updt st | ShadowItem _ => process_shadow updt st end (* Initialize status, given the subgoal in pure logic form. *) fun init_status ctxt subgoal = let (* Free variables are implicitly quantified over at the front. *) - val raw_updt = AddBoxes {id = [], sc = NONE, init_assum = get_neg' subgoal} + val raw_updt = AddBoxes {id = [], sc = NONE, init_assum = UtilLogic.get_neg' subgoal} val updt = {sc = 0, prfstep_name = "$INIT", source = [], raw_updt = raw_updt} in ctxt |> Status.empty_status |> Status.add_item BoxItem.null_item |> apply_update updt end (* Given a condition status -> bool, step until the condition is satisfied. Return steps remaining as well as updated status. *) fun solve_root (steps, (st as {queue, ctxt, ...})) = if BoxID.is_box_resolved ctxt [BoxID.home_id] then (steps, st) else if steps = 0 then error "Maximum number of steps reached" else let val updt = Updates_Heap.min queue handle List.Empty => error "No more moves" val st' = st |> Status.delmin_from_queue |> apply_update updt in solve_root (steps - 1, st') end end (* structure ProofStatus. *) (* Definition of auto2. *) signature AUTO2 = sig val auto2_tac: Proof.context -> tactic end -structure Auto2 : AUTO2 = +functor Auto2( + structure ProofStatus: PROOFSTATUS; + structure Status: STATUS; + structure UtilLogic: UTIL_LOGIC; + ): AUTO2 = struct fun auto2_tac ctxt state = let val subgoals = state |> Thm.cprop_of |> Drule.strip_imp_prems val steps = Config.get ctxt max_steps in if null subgoals then Seq.empty else let val c_subgoal = hd subgoals val subgoal = Thm.term_of c_subgoal val _ = tracing ( "Subgoal 1 of " ^ (string_of_int (length subgoals)) ^ ":\n" ^ (Syntax.string_of_term ctxt subgoal) ^ "\n") val subgoal_norm = Conv.every_conv [ Util.normalize_meta_horn ctxt, UtilLogic.to_obj_conv ctxt] c_subgoal val st = ProofStatus.init_status ctxt (Util.rhs_of subgoal_norm) val (steps', st') = Util.timer ("Total time: ", fn _ => ProofStatus.solve_root (steps, st)) val _ = writeln ("Finished in " ^ (string_of_int (steps - steps')) ^ " steps.") val th = st' |> Status.get_resolve_th |> Thm.equal_elim (meta_sym subgoal_norm) in Seq.single (Thm.implies_elim state th) end end end (* structure Auto2 *) diff --git a/thys/Auto2_HOL/auto2_data.ML b/thys/Auto2_HOL/auto2_data.ML --- a/thys/Auto2_HOL/auto2_data.ML +++ b/thys/Auto2_HOL/auto2_data.ML @@ -1,88 +1,92 @@ (* File: auto2_data.ML Author: Bohua Zhan Updating of all data maintained at proof time. *) signature AUTO2_DATA = sig val relevant_terms_single: box_item -> term list val add_terms: box_item list -> (box_id * cterm) list -> Proof.context -> Proof.context val get_incr_type: box_item list -> box_item list -> Proof.context -> Proof.context val get_single_type: Proof.context -> Proof.context end; -structure Auto2Data : AUTO2_DATA = +functor Auto2Data( + structure PropertyData: PROPERTY_DATA; + structure RewriteTable: REWRITE_TABLE; + structure WellformData: WELLFORM_DATA; + ): AUTO2_DATA = struct (* Procedure to add a new term. Here old_items is the list of existing items. term_infos is a list of (id, ct) pairs. *) fun add_terms old_items term_infos ctxt = let val ts = map (Thm.term_of o snd) term_infos val (edges, ctxt') = RewriteTable.add_term_list term_infos ctxt val new_ts = map (Thm.term_of o snd) (RewriteTable.get_new_terms (ctxt, ctxt')) val imm_properties = maps (PropertyData.apply_property_update_on_term ctxt' []) ts in ctxt' |> PropertyData.process_update_property imm_properties |> fold PropertyData.process_rewrite_property edges |> fold WellformData.initialize_wellform_data ts |> WellformData.complete_wellform_data_for_terms old_items new_ts end (* Helper function for the two functions below. *) fun relevant_terms_single item = let val {ty_str, tname, ...} = item in if ty_str = "EQ" then map Thm.term_of tname else [] end (* Use the given items to update the current context data, producing the incremental context. Here old_items is the list of existing items. items is the list of new items. Update the rewrite table, property table, wellform table, and the custom tables. *) fun get_incr_type old_items items ctxt = let (* List of relevant terms. *) val relevant_terms = items |> maps relevant_terms_single |> RewriteTable.get_reachable_terms true ctxt fun add_one_info item ctxt = let val {id, ty_str, prop, ...} = item in if ty_str = "EQ" then let val (edges, ctxt') = RewriteTable.add_rewrite (id, prop) ctxt in ctxt' |> fold PropertyData.process_rewrite_property edges end else if ty_str = "PROPERTY" then PropertyData.add_property (id, prop) ctxt else ctxt end val match_items = items @ filter (fn {tname, ...} => exists (Util.has_subterm relevant_terms) (map Thm.term_of tname)) old_items in ctxt |> fold add_one_info items |> WellformData.complete_wellform_data match_items end fun get_single_type ctxt = ctxt |> RewriteTable.clear_incr |> PropertyData.clear_incr |> WellformData.clear_incr end (* Auto2Data *) diff --git a/thys/Auto2_HOL/auto2_outer.ML b/thys/Auto2_HOL/auto2_outer.ML --- a/thys/Auto2_HOL/auto2_outer.ML +++ b/thys/Auto2_HOL/auto2_outer.ML @@ -1,667 +1,690 @@ (* File: auto2_outer.ML Author: Bohua Zhan Proof language for auto2. *) signature AUTO2_OUTER = sig val auto2_solve: Proof.context -> cterm -> thm val init_state: Proof.state -> Proof.state val refine_subgoal_th: thm -> thm -> thm val have_resolve: Proof.context -> term list -> thm -> term -> thm val have_after_qed: Proof.context -> thm -> thm -> thm val have_cmd: bool * string * bool -> Proof.state -> Proof.state val subgoal_cmd: string -> Proof.state -> Proof.state val endgoal_cmd: Proof.state -> Proof.state val end_cmd: Proof.state -> Proof.state val qed_cmd: Proof.state -> Proof.context val obtain_resolve: Proof.context -> term list -> thm -> term -> thm val obtain_after_qed: Proof.context -> thm list -> thm -> thm val obtain_cmd: string list * string list * bool -> Proof.state -> Proof.state val case_resolve: Proof.context -> term list -> thm -> thm val case_after_qed: Proof.context -> thm list -> thm -> thm val case_cmd: string * bool -> Proof.state -> Proof.state val contra_resolve: Proof.context -> term list -> term list -> term -> thm val contra_after_qed: Proof.context -> thm -> thm val contra_cmd: Proof.state -> Proof.state val let_resolve: Proof.context -> term list -> term -> term list -> term -> thm val let_after_qed: Proof.context -> term list -> thm -> thm val let_cmd: string list -> Proof.state -> Proof.state end; -structure Auto2_Outer : AUTO2_OUTER = +signature AUTO2_KEYWORDS = +sig + val case': string * Position.T + val contradiction: string * Position.T + val end': string * Position.T + val endgoal: string * Position.T + val have: string * Position.T + val let': string * Position.T + val obtain: string * Position.T + val proof: string * Position.T + val qed: string * Position.T + val subgoal: string * Position.T + val rule: string parser + val with': string parser +end; + +functor Auto2_Outer( + structure Auto2: AUTO2; + structure Auto2_Keywords: AUTO2_KEYWORDS; + structure UtilBase: UTIL_BASE; + structure UtilLogic: UTIL_LOGIC; + ) : AUTO2_OUTER = struct (* Use auto2 to solve the given statement *) fun auto2_solve ctxt stmt = let val goal = stmt |> Thm.trivial |> Goal.protect 1 in (Auto2.auto2_tac ctxt goal) |> Seq.hd |> Goal.conclude end (* Initiate auto2 state, using goal from the Isar state. *) fun init_state state = let val {goal, context = ctxt, ...} = Proof.goal state val _ = assert (Auto2_State.get_num_frame ctxt = 0) "init_state: state not empty." val subgoals = goal |> Thm.cprop_of |> Drule.strip_imp_prems val init_frame = Auto2_State.simple_frame (hd subgoals, NONE) in state |> Proof.map_contexts (Auto2_State.push_head init_frame) end val _ = - Outer_Syntax.command @{command_keyword "@proof"} "begin auto2 proof" + Outer_Syntax.command Auto2_Keywords.proof "begin auto2 proof" (Scan.succeed (Toplevel.proof init_state)) (* Given th of the form A ==> B, and prop of the form B ==> C, return the theorem A ==> C. In effect, we modified (refined) the subgoal in prop from B to A using th. *) fun refine_subgoal_th th prop = let val assum = hd (Drule.cprems_of th) (* A *) in th |> Util.send_first_to_hyps (* [A] ==> B *) |> Thm.implies_elim prop (* [A] ==> C *) |> Thm.implies_intr assum (* A ==> C *) end (* Given theorem As ==> B, a term C, and a list of variables x, obtain the theorem (!!x. As ==> B ==> C) ==> (!!x. As ==> C). *) fun have_resolve ctxt vars th concl = let val prop = Thm.prop_of th val (As, B) = Logic.strip_horn prop val cAs = map (Thm.cterm_of ctxt) As val c_vars = map (Thm.cterm_of ctxt) vars val stmt = (Util.list_meta_horn (vars, (As @ [B], concl))) |> Thm.cterm_of ctxt val thAs = map Thm.assume cAs val th_bc = stmt |> Thm.assume (* !!x. As ==> B ==> C *) |> fold Thm.forall_elim c_vars |> fold Thm.elim_implies thAs (* B ==> C *) val th_b = th |> fold Thm.elim_implies thAs (* B *) val th_c = Thm.implies_elim th_bc th_b (* C *) in th_c |> fold Thm.implies_intr (rev cAs) |> fold Thm.forall_intr (rev c_vars) |> Thm.implies_intr stmt end (* Given prop in the form (!!x. As ==> C) ==> D, and th in the form As ==> B, return the theorem (!!x. As ==> B ==> C) ==> D. In effect, this adds B as an extra assumption in the current subgoal. *) fun have_after_qed ctxt th prop = let val (vars, (_, concl)) = prop |> Thm.prems_of |> the_single |> Util.strip_meta_horn val have_res = have_resolve ctxt vars th concl in refine_subgoal_th have_res prop end (* Implementation of @have command. *) fun have_cmd (is_rule, t, is_with) state = let val {context = ctxt, ...} = Proof.goal state (* Goal to be proved: !!x. As' ==> C *) val (vars, (assums, concl)) = t |> Syntax.read_term ctxt |> UtilLogic.strip_obj_horn (* Current subgoal *) val (_, (As, _)) = ctxt |> Auto2_State.get_subgoal |> Util.strip_meta_horn (* Actual goal: strip vars, and add As as additional assumptions. *) - val stmt = Logic.list_implies (As @ map mk_Trueprop assums, - mk_Trueprop concl) + val stmt = Logic.list_implies (As @ map UtilLogic.mk_Trueprop assums, + UtilLogic.mk_Trueprop concl) (* Post-processing. We begin with th: As ==> As' ==> C, and want to create the theorem As ==> (!!x. As' ==> C), then use it to insert !!x. As' ==> C as an additional assumption. *) fun post_process th = th |> funpow (length As) Util.send_first_to_hyps (* [As] ==> As' ==> C *) |> fold Thm.forall_intr (rev (map (Thm.cterm_of ctxt) vars)) (* [As] ==> !!x. As' ==> C *) |> apply_to_thm (UtilLogic.to_obj_conv ctxt) (* [As] ==> !x. As' --> C *) |> fold Thm.implies_intr (rev (map (Thm.cterm_of ctxt) As)) fun after_qed ths prop = have_after_qed ctxt (post_process (the_single ths)) prop val add_prem_only = if is_rule then I else Auto2_State.add_prem_only (UtilLogic.list_obj_horn (vars, (assums, concl))) val new_vars = filter_out (Variable.is_fixed ctxt o fst o Term.dest_Free) vars in if not is_with then let (* Run auto2 to obtain As ==> As' ==> C. *) val ctxt' = fold Util.declare_free_term new_vars ctxt val th = auto2_solve ctxt' (Thm.cterm_of ctxt' stmt) in state |> Proof.map_contexts (Auto2_State.map_head_th (after_qed [th])) |> Proof.map_contexts add_prem_only end else let val new_frame = Auto2_State.simple_frame ( Thm.cterm_of ctxt stmt, SOME ([], after_qed)) in (* The order here does not matter much. All three actions will be completed before proof begins. *) state |> Proof.map_contexts add_prem_only |> Proof.map_contexts (fold Util.declare_free_term new_vars) |> Proof.map_contexts (Auto2_State.push_head new_frame) end end val read_rule = - Scan.option (Parse.$$$ "(" --| @{keyword "@rule"} |-- Parse.$$$ ")") + Scan.option (Parse.$$$ "(" --| Auto2_Keywords.rule |-- Parse.$$$ ")") val read_with = - Scan.option @{keyword "@with"} + Scan.option Auto2_Keywords.with' val _ = - Outer_Syntax.command @{command_keyword "@have"} "intermediate goal" + Outer_Syntax.command Auto2_Keywords.have "intermediate goal" (read_rule -- Parse.term -- read_with >> (fn ((rule_opt, t), with_opt) => Toplevel.proof (fn state => have_cmd (is_some rule_opt, t, is_some with_opt) state))) (* Whether the goal is already resolved. *) fun is_goal_resolved th = not (Util.is_head (Logic.protectC) (Thm.concl_of th)) (* Match pattern pat with t. Here pat is the pattern for selecting the i'th subgoal. *) fun match_subgoal_pat thy t (i, pat) = let val inst = Pattern.first_order_match thy (pat, t) fo_init val vars = rev (map Var (Term.add_vars pat [])) val ts = map (Util.subst_term_norm inst) vars in SOME (i, ts) end handle Pattern.MATCH => NONE (* Implementation of the @subgoal command. *) fun subgoal_cmd s state = let val {context = ctxt, ...} = Proof.goal state val {goals, selected, ...} = Auto2_State.get_top_frame ctxt - val _ = assert (selected = NONE) "@subgoal: already selected a subgoal." + val _ = assert (selected = NONE) + ((fst Auto2_Keywords.subgoal) ^ ": already selected a subgoal.") (* Figure out which subgoal to select by matching *) val thy = Proof_Context.theory_of ctxt val t = Syntax.read_term ctxt s val res = get_first (match_subgoal_pat thy t) (tag_list 0 (map fst goals)) in case res of - NONE => error "@subgoal: pattern not found" + NONE => error ((fst Auto2_Keywords.subgoal) ^ ": pattern not found") | SOME (i, ts) => let val prop = snd (nth goals i) val _ = assert (not (is_goal_resolved prop)) - "@subgoal: goal already resolved." + ((fst Auto2_Keywords.subgoal) ^ ": goal already resolved.") (* Use the matched variables ts to name the forall variables. *) val t' = case Thm.prop_of prop of imp $ A $ B => imp $ Util.rename_abs_term ts A $ B | _ => raise Fail "subgoal_cmd" val prop = Thm.renamed_prop t' prop val new_ts = filter_out (Variable.is_fixed ctxt o fst o Term.dest_Free) ts in state |> Proof.map_contexts (Auto2_State.set_selected (SOME i)) |> Proof.map_contexts (Auto2_State.map_head_th (K prop)) |> Proof.map_contexts (fold Util.declare_free_term new_ts) end end val _ = - Outer_Syntax.command @{command_keyword "@subgoal"} "select subgoal" + Outer_Syntax.command Auto2_Keywords.subgoal "select subgoal" (Parse.term >> (fn s => Toplevel.proof (fn state => subgoal_cmd s state))) (* Use auto2 to resolve one of the goals. *) fun auto2_solve_goal ctxt prop = if is_goal_resolved prop then prop else let val (vars, (As, C)) = prop |> Thm.prems_of |> the_single |> Util.strip_meta_horn val body = Util.list_meta_horn ([], (As, C)) val th = auto2_solve ctxt (Thm.cterm_of ctxt body) in th |> fold Thm.forall_intr (rev (map (Thm.cterm_of ctxt) vars)) |> Thm.implies_elim prop |> Goal.conclude end (* Implementation of the @endgoal command. There should be more than one goal in the current frame, and one of the goals is selected. Use auto2 to finish the selected goal. *) fun endgoal_cmd state = let val {context = ctxt, ...} = Proof.goal state val {goals, selected, ...} = Auto2_State.get_top_frame ctxt val _ = assert (length goals > 1 andalso is_some selected) - "@endgoal: called without a selection" + ((fst Auto2_Keywords.endgoal) ^ ": called without a selection") val prop = snd (nth goals (the selected)) val solved_prop = auto2_solve_goal ctxt prop in state |> Proof.map_contexts (Auto2_State.map_head_th (K solved_prop)) |> Proof.map_contexts (Auto2_State.set_selected NONE) end val _ = - Outer_Syntax.command @{command_keyword "@endgoal"} "endgoal of with block" + Outer_Syntax.command Auto2_Keywords.endgoal "endgoal of with block" (Scan.succeed ( Toplevel.proof (fn state => endgoal_cmd state))) (* Implementation of the @end command. If there is exactly one goal in the current frame, that goal should be selected and unresolved. Use auto2 to resolve that goal. If there are multiple goals in the current frame, use auto2 to finish all unresolved goals. *) fun end_cmd state = let val {context = ctxt, ...} = Proof.goal state val {goals, selected, after_qed, ...} = Auto2_State.get_top_frame ctxt val _ = assert ((length goals = 1 andalso selected = SOME 0) orelse (length goals > 1 andalso selected = NONE)) - "@end: cannot call within an selection." + ((fst Auto2_Keywords.end') ^ ": cannot call within an selection.") val solved_props = map (auto2_solve_goal ctxt o snd) goals val _ = assert (forall is_goal_resolved solved_props) - "@end: failed to resolve all goals" + ((fst Auto2_Keywords.end') ^ ": failed to resolve all goals") (* List of new variables, and callback function. *) val (new_vars, f) = the after_qed in state |> Proof.map_contexts Auto2_State.pop_head |> Proof.map_contexts (Auto2_State.map_head_th (f solved_props)) |> Proof.map_contexts (fold Util.declare_free_term new_vars) end val _ = - Outer_Syntax.command @{command_keyword "@end"} "end of with block" + Outer_Syntax.command Auto2_Keywords.end' "end of with block" (Scan.succeed ( Toplevel.proof (fn state => end_cmd state))) (* Implementation of the @qed command. *) fun qed_cmd state = let val {context = ctxt, ...} = Proof.goal state val _ = assert (Auto2_State.get_num_frame ctxt = 1) "Qed should be applied outside any 'with' blocks" val prop = Auto2_State.get_selected ctxt val (vars, (As, C)) = prop |> Thm.prems_of |> the_single |> Util.strip_meta_horn val body = Util.list_meta_horn ([], (As, C)) val th = (auto2_solve ctxt (Thm.cterm_of ctxt body)) |> fold Thm.forall_intr (rev (map (Thm.cterm_of ctxt) vars)) val new_prop = Thm.implies_elim prop th val method = K (Context_Tactic.CONTEXT_TACTIC (PRIMITIVE (K new_prop))) in state |> Proof.map_contexts (Auto2_State.map_head_th (K new_prop)) |> Proof.refine_singleton (Method.Basic (K method)) |> Proof.global_done_proof end val _ = - Outer_Syntax.command @{command_keyword "@qed"} "end of proof" + Outer_Syntax.command Auto2_Keywords.qed "end of proof" (Scan.succeed ( Toplevel.end_proof (K qed_cmd))) (* Given theorem As ==> EX y. P(y), a term C, and a list of variables x, obtain the theorem (!!x y. As ==> P(y) ==> C) ==> (!!x. As ==> C). *) fun obtain_resolve ctxt vars th concl = let val prop = Thm.prop_of th (* B is EX y. P(y) *) val (As, B) = Logic.strip_horn prop val cAs = map (Thm.cterm_of ctxt) As val c_vars = map (Thm.cterm_of ctxt) vars - val (new_vars, body) = UtilLogic.strip_exists (dest_Trueprop B) + val (new_vars, body) = UtilLogic.strip_exists (UtilLogic.dest_Trueprop B) - val As' = As @ [mk_Trueprop body] + val As' = As @ [UtilLogic.mk_Trueprop body] val stmt = (Util.list_meta_horn (vars @ new_vars, (As', concl))) |> Thm.cterm_of ctxt val thAs = map Thm.assume cAs val th_b = th |> fold Thm.elim_implies thAs (* P(y) ==> C *) val th_bc = stmt |> Thm.assume |> fold Thm.forall_elim c_vars |> fold Thm.forall_elim (map (Thm.cterm_of ctxt) new_vars) |> fold Thm.elim_implies thAs val th_c = (fold (UtilLogic.ex_elim ctxt) (rev new_vars) th_bc) |> Thm.elim_implies th_b in th_c |> fold Thm.implies_intr (rev cAs) |> fold Thm.forall_intr (rev c_vars) |> Thm.implies_intr stmt end (* Give prop in the form (!!x. As ==> C) ==> D, and th in the form As ==> EX y. P(y), return the theorem (!!x y. As ==> P(y) ==> C) ==> D. In effect, this creates new variables y with property P(y) in the current subgoal. *) fun obtain_after_qed ctxt ths prop = let val th = the_single ths val (vars, (As, C)) = prop |> Thm.prems_of |> the_single |> Util.strip_meta_horn val body = Util.list_meta_horn ([], (As, C)) val concl = Logic.strip_imp_concl body val obtain_res = obtain_resolve ctxt vars th concl in refine_subgoal_th obtain_res prop end (* Implementation of @obtain command. *) fun obtain_cmd (vars, conds, is_with) state = let val {context = ctxt, ...} = Proof.goal state (* First read list of variables and conditions *) val (vars, conds) = (Syntax.read_terms ctxt (vars @ conds)) |> chop (length vars) (* Elements of vars can be in the form x : A, process it into a variable x and a condition x : A. *) val (vars, conds) = - (map (fn t => if is_mem t then dest_arg1 t else t) vars, - filter is_mem vars @ conds) + (map (fn t => if UtilLogic.is_mem t then dest_arg1 t else t) vars, + filter UtilLogic.is_mem vars @ conds) (* Goal: EX vars. conds *) - val C = conds |> list_conj |> fold mk_exists (rev vars) |> mk_Trueprop + val C = conds |> UtilLogic.list_conj |> fold UtilLogic.mk_exists (rev vars) + |> UtilLogic.mk_Trueprop (* Current subgoal *) val (_, (As, _)) = ctxt |> Auto2_State.get_subgoal |> Util.strip_meta_horn (* Actual goal: As ==> EX vars. conds *) val stmt = Logic.list_implies (As, C) val new_vars = filter_out (Variable.is_fixed ctxt o fst o Term.dest_Free) vars val after_qed = obtain_after_qed ctxt in if not is_with then let (* Run auto2 to obtain As ==> EX vars. conds. *) val th = auto2_solve ctxt (Thm.cterm_of ctxt stmt) in state |> Proof.map_contexts (fold Util.declare_free_term new_vars) |> Proof.map_contexts (Auto2_State.map_head_th (after_qed [th])) end else let val new_frame = Auto2_State.simple_frame ( Thm.cterm_of ctxt stmt, SOME (new_vars, after_qed)) in state |> Proof.map_contexts (Auto2_State.push_head new_frame) end end val obtain_param = Scan.repeat Parse.term -- Scan.option (@{keyword "where"} |-- Scan.repeat Parse.term) -- read_with val _ = - Outer_Syntax.command @{command_keyword "@obtain"} "obtain variables" + Outer_Syntax.command Auto2_Keywords.obtain "obtain variables" (obtain_param >> (fn ((vars, conds), with_opt) => Toplevel.proof ( fn state => obtain_cmd (vars, these conds, is_some with_opt) state))) (* Given theorem As ==> B ==> C, and a list of variables x, obtain the theorem (!!x. As ==> ~B ==> C) ==> (!!x. As ==> C). *) fun case_resolve ctxt vars th = let val prop = Thm.prop_of th val ((As, B), C) = prop |> Logic.strip_horn |> apfst split_last val cAs = map (Thm.cterm_of ctxt) As val c_vars = map (Thm.cterm_of ctxt) vars - val nB = mk_Trueprop (Not $ (dest_Trueprop B)) + val nB = UtilLogic.mk_Trueprop (UtilLogic.Not $ (UtilLogic.dest_Trueprop B)) val stmt = (Util.list_meta_horn (vars, (As @ [nB], C))) |> Thm.cterm_of ctxt val thAs = map Thm.assume cAs val th_nbc = stmt |> Thm.assume (* !!x. As ==> ~B ==> C *) |> fold Thm.forall_elim c_vars |> fold Thm.elim_implies thAs (* ~B ==> C *) val th_bc = th |> fold Thm.elim_implies thAs - val inst = fo_init |> Util.update_env (("P",0), dest_Trueprop B) - |> Util.update_env (("Q",0), dest_Trueprop C) + val inst = fo_init |> Util.update_env (("P",0), UtilLogic.dest_Trueprop B) + |> Util.update_env (("Q",0), UtilLogic.dest_Trueprop C) val th_c = (Util.subst_thm ctxt inst UtilBase.case_split_th) |> fold Thm.elim_implies [th_bc, th_nbc] in th_c |> fold Thm.implies_intr (rev cAs) |> fold Thm.forall_intr (rev c_vars) |> Thm.implies_intr stmt end (* Given prop in the form (!!x. As ==> C) ==> D, and th in the form As ==> B ==> C, return the theorem (!!x. As ==> ~B ==> C) ==> D. In effect, this adds ~B as extra assumption in the current subgoal. *) fun case_after_qed ctxt ths prop = let val th = the_single ths val (vars, _) = prop |> Thm.prems_of |> the_single |> Util.strip_meta_horn val case_res = case_resolve ctxt vars th in refine_subgoal_th case_res prop end (* Implementation of @case command. *) fun case_cmd (t, is_with) state = let val {context = ctxt, ...} = Proof.goal state - val B = t |> Syntax.read_term ctxt |> mk_Trueprop + val B = t |> Syntax.read_term ctxt |> UtilLogic.mk_Trueprop val (_, (As, C)) = ctxt |> Auto2_State.get_subgoal |> Util.strip_meta_horn val stmt = Logic.list_implies (As @ [B], C) |> Thm.cterm_of ctxt val after_qed = case_after_qed ctxt in if not is_with then let (* Run auto2 to obtain A_1 ==> ... ==> A_n ==> B ==> C *) val th = auto2_solve ctxt stmt in state |> Proof.map_contexts (Auto2_State.map_head_th (after_qed [th])) end else let val new_frame = Auto2_State.simple_frame (stmt, SOME ([], after_qed)) in state |> Proof.map_contexts (Auto2_State.push_head new_frame) end end val _ = - Outer_Syntax.command @{command_keyword "@case"} "intermediate case" + Outer_Syntax.command Auto2_Keywords.case' "intermediate case" (Parse.term -- read_with >> (fn (t, with_opt) => Toplevel.proof (fn state => case_cmd (t, is_some with_opt) state))) (* Given a list of variables x, assumptions As, and term C, return the theorem (!!x. As ==> ~C ==> False) ==> (!!x. As ==> C). *) fun contra_resolve ctxt vars As C = let - val stmt = (Util.list_meta_horn (vars, (As @ [get_neg' C], pFalse))) + val stmt = (Util.list_meta_horn (vars, (As @ [UtilLogic.get_neg' C], UtilLogic.pFalse))) |> Thm.cterm_of ctxt val cAs = map (Thm.cterm_of ctxt) As val c_vars = map (Thm.cterm_of ctxt) vars val thAs = map Thm.assume cAs val th_c = stmt |> Thm.assume (* !!x. A_1 ==> ... ==> A_n ==> ~C ==> False *) |> fold Thm.forall_elim c_vars |> fold Thm.elim_implies thAs (* ~C ==> False *) |> apply_to_thm UtilLogic.rewrite_from_contra_form (* C *) in th_c |> fold Thm.implies_intr (rev cAs) |> fold Thm.forall_intr (rev c_vars) |> Thm.implies_intr stmt end (* Given prop in the form (!!x. As ==> C) ==> D, return the theorem (!!x. As ==> ~C ==> False) ==> D. *) fun contra_after_qed ctxt prop = let val (vars, (As, C)) = prop |> Thm.prems_of |> the_single |> Util.strip_meta_horn val contra_res = contra_resolve ctxt vars As C in refine_subgoal_th contra_res prop end (* Implementation of @contradiction command. *) fun contra_cmd state = let val {context = ctxt, ...} = Proof.goal state val after_qed = contra_after_qed ctxt in state |> Proof.map_contexts (Auto2_State.map_head_th after_qed) end val _ = - Outer_Syntax.command @{command_keyword "@contradiction"} + Outer_Syntax.command Auto2_Keywords.contradiction "apply proof by contradiction" (Scan.succeed ( Toplevel.proof (fn state => contra_cmd state))) (* Given a list of variables x, an equation y = t, assumptions As, and a term C, return the theorem (!!x y. As ==> y = t ==> C) ==> (!!x. As ==> C). *) fun let_resolve ctxt vars eq As C = let val thy = Proof_Context.theory_of ctxt (* Call obtain_resolve with th as A_i ==> EX y. y = t. *) - val (lhs, rhs) = dest_eq eq - val pat_a = case UtilBase.ex_vardef_th |> prop_of' |> dest_arg of + val (lhs, rhs) = UtilBase.dest_eq eq + val pat_a = case UtilBase.ex_vardef_th |> UtilLogic.prop_of' |> dest_arg of Abs (_, _, b) => dest_arg b | _ => raise Fail "ex_vardef_th" val (x, _) = Term.dest_Free lhs val inst = Pattern.first_order_match thy (pat_a, rhs) fo_init val ex_th = Util.subst_thm ctxt inst UtilBase.ex_vardef_th val t' = case Thm.prop_of ex_th of A $ (B $ Abs (_, T, body)) => A $ (B $ Abs (x, T, body)) | _ => error "let_resolve" val ex_th = ex_th |> Thm.renamed_prop t' |> fold Thm.implies_intr (rev (map (Thm.cterm_of ctxt) As)) in obtain_resolve ctxt vars ex_th C end (* Given prop in the form (!!x. As ==> C) ==> D, and an equation y = t, return the theorem (!!x y. As ==> y = t ==> C) ==> D. In effect, this creates a new variable y with property y = t in the current subgoal. *) fun let_after_qed ctxt eqs prop = let fun fold_one eq prop = let val (vars, (As, C)) = prop |> Thm.prems_of |> the_single |> Util.strip_meta_horn val let_res = let_resolve ctxt vars eq As C in refine_subgoal_th let_res prop end in fold fold_one eqs prop end (* Implementation of the @let command. *) fun let_cmd eqs state = let val {context = ctxt, ...} = Proof.goal state val eqs = Syntax.read_terms ctxt eqs - val vars = map (fst o dest_eq) eqs + val vars = map (fst o UtilBase.dest_eq) eqs val new_vars = filter_out (Variable.is_fixed ctxt o fst o Term.dest_Free) vars val after_qed = let_after_qed ctxt in state |> Proof.map_context (fold Util.declare_free_term new_vars) |> Proof.map_contexts (Auto2_State.map_head_th (after_qed eqs)) end val _ = - Outer_Syntax.command @{command_keyword "@let"} "define a variable" + Outer_Syntax.command Auto2_Keywords.let' "define a variable" (Scan.repeat Parse.term >> (fn eqs => Toplevel.proof (fn state => let_cmd eqs state))) end (* structure Auto2_Outer *) diff --git a/thys/Auto2_HOL/auto2_setup.ML b/thys/Auto2_HOL/auto2_setup.ML new file mode 100644 --- /dev/null +++ b/thys/Auto2_HOL/auto2_setup.ML @@ -0,0 +1,163 @@ +(* + File: auto2_setup.ML + Author: Kevin Kappelmann + + Setup for auto2 +*) + +signature AUTO2_SETUP = +sig + structure Auto2_Keywords: AUTO2_KEYWORDS + structure UtilBase: UTIL_BASE + structure UtilLogic : UTIL_LOGIC + structure Basic_UtilLogic : BASIC_UTIL_LOGIC + structure WfTerm: WFTERM + structure RewriteTable: REWRITE_TABLE + structure PropertyData: PROPERTY_DATA + structure Matcher: MATCHER + structure BoxItem: BOXITEM + structure ItemIO: ITEM_IO + structure WellformData: WELLFORM_DATA + structure Auto2Data: AUTO2_DATA + structure Update: UPDATE + structure Status: STATUS + structure Normalizer: NORMALIZER + structure ProofStep: PROOFSTEP + structure ProofStepData: PROOFSTEP_DATA + structure Logic_ProofSteps: LOGIC_PROOFSTEPS + structure ProofStatus: PROOFSTATUS + structure Auto2: AUTO2 + structure Auto2_Outer: AUTO2_OUTER +end + +functor Auto2_Setup( + structure Auto2_Keywords: AUTO2_KEYWORDS; + structure UtilBase: UTIL_BASE; + ) : AUTO2_SETUP = +struct + +structure Auto2_Keywords = Auto2_Keywords +structure UtilBase = UtilBase +structure UtilLogic = UtilLogic(UtilBase) +structure Basic_UtilLogic: BASIC_UTIL_LOGIC = UtilLogic +structure Property = Property( + structure Basic_UtilBase = UtilBase; + structure UtilLogic = UtilLogic; +) +structure WfTerm = WfTerm( + structure Basic_UtilBase = UtilBase; + structure UtilLogic = UtilLogic; +) +structure RewriteTable = RewriteTable(UtilLogic) +structure PropertyData = PropertyData( + structure Basic_UtilBase = UtilBase; + structure Property = Property; + structure RewriteTable =RewriteTable; + structure UtilLogic = UtilLogic; +) +structure Matcher = Matcher(RewriteTable) +structure BoxItem = BoxItem(UtilBase) +structure ItemIO = ItemIO( + structure Matcher = Matcher; + structure Property = Property; + structure PropertyData = PropertyData; + structure RewriteTable = RewriteTable; + structure UtilBase = UtilBase; + structure UtilLogic = UtilLogic; +) +structure WellformData = WellformData( + structure Basic_UtilBase = UtilBase; + structure Basic_UtilLogic = Basic_UtilLogic; + structure ItemIO = ItemIO; + structure Property = Property; + structure PropertyData = PropertyData; + structure RewriteTable = RewriteTable; + structure WfTerm = WfTerm; +) +structure Auto2Data = Auto2Data( + structure PropertyData = PropertyData; + structure RewriteTable = RewriteTable; + structure WellformData = WellformData; +) +structure Update = Update( + structure BoxItem = BoxItem; + structure ItemIO = ItemIO; + structure UtilLogic = UtilLogic; +) +structure Status = Status( + structure BoxItem = BoxItem; + structure ItemIO = ItemIO; + structure UtilLogic = UtilLogic + structure PropertyData = PropertyData; + structure RewriteTable = RewriteTable; + structure WellformData = WellformData; +) +val _ = Theory.setup (ItemIO.add_basic_item_io) +structure Normalizer = Normalizer( + structure BoxItem = BoxItem; + structure UtilBase = UtilBase; + structure UtilLogic = UtilLogic; + structure Update = Update; +) +structure ProofStep = ProofStep( + structure ItemIO = ItemIO; + structure Matcher = Matcher; + structure PropertyData = PropertyData; + structure RewriteTable = RewriteTable; + structure UtilBase = UtilBase; + structure UtilLogic = UtilLogic; + structure Update = Update; + structure WellformData = WellformData; +) +structure ProofStepData = ProofStepData( + structure ItemIO = ItemIO; + structure Normalizer = Normalizer; + structure ProofStep = ProofStep; + structure Property = Property; + structure UtilBase = UtilBase; + structure UtilLogic = UtilLogic; +) +structure Logic_ProofSteps = Logic_ProofSteps( + structure BoxItem = BoxItem; + structure ItemIO = ItemIO; + structure Matcher = Matcher; + structure Normalizer = Normalizer; + structure Property = Property; + structure ProofStepData = ProofStepData; + structure RewriteTable = RewriteTable; + structure UtilBase = UtilBase; + structure UtilLogic = UtilLogic; + structure Update = Update; +) +val _ = Theory.setup Logic_ProofSteps.add_logic_proofsteps +val _ = Theory.setup Logic_ProofSteps.add_disj_proofsteps +val _ = Theory.setup Logic_ProofSteps.add_disj_normalizers +structure ProofStatus = ProofStatus( + structure Auto2Data = Auto2Data; + structure BoxItem = BoxItem; + structure ItemIO = ItemIO; + structure Normalizer = Normalizer; + structure Logic_ProofSteps = Logic_ProofSteps; + structure Property = Property; + structure PropertyData = PropertyData; + structure ProofStepData = ProofStepData; + structure RewriteTable = RewriteTable; + structure Status = Status; + structure UtilBase = UtilBase; + structure UtilLogic = UtilLogic; + structure Update = Update; + structure WellformData = WellformData; +) +structure Auto2 = Auto2( + structure ProofStatus = ProofStatus; + structure Status = Status; + structure UtilLogic = UtilLogic; +) +structure Auto2_Outer = Auto2_Outer( + structure Auto2 = Auto2; + structure Auto2_Keywords = Auto2_Keywords; + structure UtilBase = UtilBase; + structure UtilLogic = UtilLogic; +) + +end \ No newline at end of file diff --git a/thys/Auto2_HOL/items.ML b/thys/Auto2_HOL/items.ML --- a/thys/Auto2_HOL/items.ML +++ b/thys/Auto2_HOL/items.ML @@ -1,701 +1,708 @@ (* File: items.ML Author: Bohua Zhan Items and matching on items. *) val TY_NULL = "NULL" val TY_EQ = "EQ" val TY_VAR = "VAR" val TY_PROP = "PROP" val TY_TERM = "TERM" val TY_PROPERTY = "PROPERTY" datatype raw_item = Handler of term list * term * thm | Fact of string * term list * thm type box_item = {uid: int, id: box_id, sc: int, ty_str: string, tname: cterm list, prop: thm} signature BOXITEM = sig (* Facts. *) val var_to_fact: term -> raw_item val term_to_fact: term -> raw_item val is_fact_raw: raw_item -> bool val match_ty_str_raw: string -> raw_item -> bool val match_ty_strs_raw: string list -> raw_item -> bool val get_tname_raw: raw_item -> term list val get_thm_raw: raw_item -> thm (* Handlers. *) val is_handler_raw: raw_item -> bool val dest_handler_raw: raw_item -> term list * term * thm (* Misc. functions. *) val eq_ritem: raw_item * raw_item -> bool val instantiate: (cterm * cterm) list -> raw_item -> raw_item val obtain_variant_frees: Proof.context * raw_item list -> Proof.context * (cterm * cterm) list val null_item: box_item val item_with_id: box_id -> box_item -> box_item val item_with_incr: box_item -> box_item val item_replace_incr: box_item -> box_item val eq_item: box_item * box_item -> bool val match_ty_str: string -> box_item -> bool val match_ty_strs: string list -> box_item -> bool (* Misc. functions. *) val merged_id: Proof.context -> box_item list -> box_id val mk_box_item: Proof.context -> int * box_id * int * raw_item -> box_item end; -structure BoxItem : BOXITEM = +functor BoxItem(UtilBase: BASIC_UTIL_BASE) : BOXITEM = struct -fun var_to_fact t = Fact (TY_VAR, [t], true_th) -fun term_to_fact t = Fact (TY_TERM, [t], true_th) +fun var_to_fact t = Fact (TY_VAR, [t], UtilBase.true_th) +fun term_to_fact t = Fact (TY_TERM, [t], UtilBase.true_th) fun is_fact_raw ritem = case ritem of Fact _ => true | _ => false fun match_ty_str_raw s ritem = case ritem of Fact (ty_str, _, _) => s = "" orelse ty_str = s | _ => false fun match_ty_strs_raw slist ritem = case ritem of Fact (ty_str, _, _) => member (op =) slist ty_str | _ => false fun get_tname_raw ritem = case ritem of Fact (_, ts, _) => ts | _ => raise Fail "get_tname_raw" fun get_thm_raw ritem = case ritem of Fact (_, _, th) => th | _ => raise Fail "get_thm_raw" fun is_handler_raw ritem = case ritem of Handler _ => true | _ => false fun dest_handler_raw ritem = case ritem of Handler (vars, t, ex_th) => (vars, t, ex_th) | _ => raise Fail "dest_handler_raw: wrong type" fun eq_ritem (ritem1, ritem2) = case ritem1 of Fact (ty1, ts1, th1) => (case ritem2 of Fact (ty2, ts2, th2) => ty1 = ty2 andalso eq_list (op aconv) (ts1, ts2) andalso Thm.eq_thm_prop (th1, th2) | _ => false) | Handler (vars1, t1, ex_th1) => (case ritem2 of Fact _ => false | Handler (vars2, t2, ex_th2) => eq_list (op aconv) (vars1, vars2) andalso t1 aconv t2 andalso Thm.eq_thm_prop (ex_th1, ex_th2)) (* Given a context and list of raw items, obtain fresh names of free variables for each internal (schematic) variable declared in the raw items, and declare the new variables in context. Return the substitution from internal schematic variables to the new free variables. *) fun obtain_variant_frees (ctxt, ritems) = let (* Original internal variables. *) val all_vars = ritems |> filter (match_ty_str_raw TY_VAR) |> maps get_tname_raw |> filter is_Free |> map dest_Free |> filter (Util.is_just_internal o fst) (* New names for these variables. *) val all_vars' = all_vars |> map (fn (x, T) => (Name.dest_internal x, T)) |> Variable.variant_frees ctxt [] val subst = map (apply2 (Thm.cterm_of ctxt o Free)) (all_vars ~~ all_vars') in (fold Util.declare_free_term (map Free all_vars') ctxt, subst) end (* Here inst is the return value of obtain_variant_frees. Perform the replacement on the ritems. *) fun instantiate subst ritem = let val subst_fun = Term.subst_atomic (map (apply2 Thm.term_of) subst) in case ritem of Handler (vars, t, ex_th) => Handler (map subst_fun vars, subst_fun t, Util.subst_thm_atomic subst ex_th) | Fact (ty_str, tname, th) => Fact (ty_str, map subst_fun tname, Util.subst_thm_atomic subst th) end val null_item = {uid = 0, id = [], sc = 0, ty_str = TY_NULL, - tname = [], prop = true_th} + tname = [], prop = UtilBase.true_th} fun item_with_id id {uid, sc, ty_str, tname, prop, ...} = {uid = uid, id = id, sc = sc, ty_str = ty_str, tname = tname, prop = prop} fun item_with_incr item = item_with_id (BoxID.add_incr_id (#id item)) item fun item_replace_incr item = item_with_id (BoxID.replace_incr_id (#id item)) item fun eq_item (item1, item2) = (#uid item1 = #uid item2) fun match_ty_str s {ty_str, ...} = (s = "" orelse s = ty_str) fun match_ty_strs slist {ty_str, ...} = member (op =) slist ty_str fun merged_id ctxt items = case items of [] => [] | {id, ...} :: items' => BoxID.merge_boxes ctxt (id, merged_id ctxt items') fun mk_box_item ctxt (uid, id, sc, ritem) = case ritem of Handler _ => raise Fail "mk_box_item: ritem must be Fact" | Fact (ty_str, ts, prop) => {uid = uid, id = id, sc = sc, ty_str = ty_str, tname = map (Thm.cterm_of ctxt) ts, prop = prop} end (* structure BoxItem. *) (* Specifies a method for matching patterns against items. - pre_match is a filter function checking whether it is possible for the pattern to match the item, after possibly instantiating some schematic variables in the pattern (for example, this function should always return true if input pattern is ?A). - match is the actual matching function, returning instantiation, as well as theorem justifying the instantiated pattern. If the matcher is for justifying a proposition, the input term to pre_match and match is of type bool. Othewise, the restrictions depend on type of item to match. *) type item_matcher = { pre_match: term -> box_item -> Proof.context -> bool, match: term -> box_item -> Proof.context -> id_inst -> id_inst_th list } (* Output function for items of a given type. *) type item_output = Proof.context -> term list * thm -> string (* Data structure containing methods involved in the input / output of items of a given type. - prop_matchers: methods for matching the item against a desired proposition. - typed_matchers: methods for matching the item against a pattern for items of the same type. - output_fn: printing function of theorems. Input is tname and the proposition. *) type item_io_info = { prop_matchers: (item_matcher * serial) list, typed_matchers: (item_matcher * serial) list, term_fn: ((term list -> term list) * serial) option, output_fn: (item_output * serial) option, shadow_fn: ((Proof.context -> box_id -> term list * cterm list -> bool) * serial) option } datatype match_arg = PropMatch of term | TypedMatch of string * term | TypedUniv of string | PropertyMatch of term | WellFormMatch of term * term type prfstep_filter = Proof.context -> id_inst -> bool signature ITEM_IO = sig val pat_of_match_arg: match_arg -> term val subst_arg: Type.tyenv * Envir.tenv -> match_arg -> match_arg val assert_valid_arg: match_arg -> unit val check_ty_str: string -> match_arg -> bool val is_ordinary_match: match_arg -> bool val is_side_match: match_arg -> bool val add_item_type: string * (term list -> term list) option * item_output option * (Proof.context -> box_id -> term list * cterm list -> bool) option -> theory -> theory val add_prop_matcher: string * item_matcher -> theory -> theory val add_typed_matcher: string * item_matcher -> theory -> theory val get_io_info: theory -> string -> item_io_info val get_prop_matchers: theory -> string -> item_matcher list val get_typed_matchers: theory -> string -> item_matcher list val prop_matcher: item_matcher val term_prop_matcher: item_matcher val null_eq_matcher: item_matcher val term_typed_matcher: item_matcher val eq_tname_typed_matcher: item_matcher val null_property_matcher: item_matcher val term_property_matcher: item_matcher val pre_match_arg: Proof.context -> match_arg -> box_item -> bool val match_arg: Proof.context -> match_arg -> box_item -> id_inst -> id_inst_th list val no_rewr_terms: term list -> term list val rewr_terms_of_item: Proof.context -> string * term list -> term list val output_prop_fn: item_output val string_of_item_info: Proof.context -> string * term list * thm -> string val add_basic_item_io: theory -> theory val string_of_raw_item: Proof.context -> raw_item -> string val string_of_item: Proof.context -> box_item -> string val trace_ritem: Proof.context -> string -> raw_item -> unit val trace_item: Proof.context -> string -> box_item -> unit val trace_ritems: Proof.context -> string -> raw_item list -> unit val trace_items: Proof.context -> string -> box_item list -> unit end; -structure ItemIO : ITEM_IO = +functor ItemIO( + structure Matcher: MATCHER; + structure Property: PROPERTY; + structure PropertyData: PROPERTY_DATA + structure RewriteTable: REWRITE_TABLE; + structure UtilBase: UTIL_BASE; + structure UtilLogic: UTIL_LOGIC + ) : ITEM_IO = struct fun add_prop_matcher_to_info mtch {prop_matchers, typed_matchers, term_fn, output_fn, shadow_fn} : item_io_info = {prop_matchers = (mtch, serial ()) :: prop_matchers, typed_matchers = typed_matchers, term_fn = term_fn, output_fn = output_fn, shadow_fn = shadow_fn} fun add_typed_matcher_to_info mtch {prop_matchers, typed_matchers, term_fn, output_fn, shadow_fn} : item_io_info = {prop_matchers = prop_matchers, typed_matchers = (mtch, serial ()) :: typed_matchers, term_fn = term_fn, output_fn = output_fn, shadow_fn = shadow_fn} structure Data = Theory_Data ( type T = item_io_info Symtab.table val empty = Symtab.empty val merge = Symtab.join (fn _ => fn ({prop_matchers = pm1, typed_matchers = tm1, term_fn = tf1, output_fn = of1, shadow_fn = sf1}, {prop_matchers = pm2, typed_matchers = tm2, term_fn = tf2, output_fn = of2, shadow_fn = sf2}) => {prop_matchers = Library.merge (eq_snd op =) (pm1, pm2), typed_matchers = Library.merge (eq_snd op =) (tm1, tm2), term_fn = (if eq_option (eq_snd op =) (tf1, tf2) then tf1 else raise Fail "join_infos: term_fn non-equal"), output_fn = (if eq_option (eq_snd op =) (of1, of2) then of1 else raise Fail "join_infos: output_fn non-equal"), shadow_fn = (if eq_option (eq_snd op =) (sf1, sf2) then sf1 else raise Fail "join_infos: shadow_fn non-equal")}) ) fun pat_of_match_arg arg = case arg of PropMatch pat => pat | TypedMatch (_, pat) => pat | TypedUniv _ => Term.dummy | PropertyMatch pat => pat | WellFormMatch (_, req) => req fun subst_arg inst arg = case arg of PropMatch pat => PropMatch (Util.subst_term_norm inst pat) | TypedMatch (ty_str, pat) => TypedMatch (ty_str, Util.subst_term_norm inst pat) | TypedUniv ty_str => TypedUniv ty_str | PropertyMatch pat => PropertyMatch (Util.subst_term_norm inst pat) | WellFormMatch (t, req) => WellFormMatch (Util.subst_term_norm inst t, Util.subst_term_norm inst req) fun assert_valid_arg arg = case arg of PropMatch pat => - assert (fastype_of pat = boolT) + assert (fastype_of pat = UtilBase.boolT) "assert_valid_arg: arg for PropMatch should be bool." | TypedMatch _ => () | TypedUniv _ => () | PropertyMatch pat => - assert (fastype_of pat = boolT) + assert (fastype_of pat = UtilBase.boolT) "assert_valid_arg: arg for PropertyMatch should be bool." | WellFormMatch (_, req) => - assert (fastype_of req = boolT) + assert (fastype_of req = UtilBase.boolT) "assert_valid_arg: arg for WellFormMatch should be bool." fun check_ty_str ty_str arg = case arg of TypedMatch (ty_str', _) => ty_str = ty_str' | TypedUniv ty_str' => ty_str = ty_str' | _ => true fun is_ordinary_match arg = case arg of PropMatch _ => true | TypedMatch _ => true | _ => false fun is_side_match arg = case arg of PropertyMatch _ => true | WellFormMatch _ => true | _ => false fun identify NONE = NONE | identify (SOME x) = (SOME (x, serial ())) fun add_item_type (ty_str, term_fn, output_fn, shadow_fn) = let val item_info : item_io_info = {prop_matchers = [], typed_matchers = [], term_fn = identify term_fn, output_fn = identify output_fn, shadow_fn = identify shadow_fn} in Data.map (Symtab.update_new (ty_str, item_info)) end fun add_prop_matcher (ty_str, it_match) = Data.map ( Symtab.map_entry ty_str (add_prop_matcher_to_info it_match)) fun add_typed_matcher (ty_str, it_match) = Data.map ( Symtab.map_entry ty_str (add_typed_matcher_to_info it_match)) fun get_io_info thy ty_str = the (Symtab.lookup (Data.get thy) ty_str) handle Option.Option => raise Fail ("get_io_info: not found " ^ ty_str) fun get_prop_matchers thy ty_str = #prop_matchers (get_io_info thy ty_str) |> map fst fun get_typed_matchers thy ty_str = #typed_matchers (get_io_info thy ty_str) |> map fst (* Prop-matching with a PROP item. *) val prop_matcher = let fun pre_match pat {tname, ...} ctxt = let val ct = the_single tname val t = Thm.term_of ct in - if is_neg pat then - is_neg t andalso - Matcher.pre_match_head ctxt (get_neg pat, UtilLogic.get_cneg ct) + if UtilLogic.is_neg pat then + UtilLogic.is_neg t andalso + Matcher.pre_match_head ctxt (UtilLogic.get_neg pat, UtilLogic.get_cneg ct) else Term.is_Var pat orelse - (not (is_neg t) andalso Matcher.pre_match_head ctxt (pat, ct)) + (not (UtilLogic.is_neg t) andalso Matcher.pre_match_head ctxt (pat, ct)) end fun match pat {tname, prop, ...} ctxt (id, inst) = let val ct = the_single tname val t = Thm.term_of ct in - if is_neg pat andalso is_neg t then + if UtilLogic.is_neg pat andalso UtilLogic.is_neg t then let val insts' = Matcher.rewrite_match_head - ctxt (get_neg pat, UtilLogic.get_cneg ct) (id, inst) + ctxt (UtilLogic.get_neg pat, UtilLogic.get_cneg ct) (id, inst) fun process_inst (inst, eq_th) = let (* This version certainly will not cancel ~~ on two sides. *) val make_neg_eq' = Thm.combination (Thm.reflexive UtilBase.cNot) in (inst, Thm.equal_elim ( - make_trueprop_eq (make_neg_eq' (meta_sym eq_th))) prop) + UtilLogic.make_trueprop_eq (make_neg_eq' (meta_sym eq_th))) prop) end in map process_inst insts' end - else if not (is_neg pat) andalso - (Term.is_Var pat orelse not (is_neg t)) then + else if not (UtilLogic.is_neg pat) andalso + (Term.is_Var pat orelse not (UtilLogic.is_neg t)) then let val insts' = Matcher.rewrite_match_head ctxt (pat, ct) (id, inst) fun process_inst (inst, eq_th) = (inst, Thm.equal_elim ( - make_trueprop_eq (meta_sym eq_th)) prop) + UtilLogic.make_trueprop_eq (meta_sym eq_th)) prop) in map process_inst insts' end else [] end in {pre_match = pre_match, match = match} end (* Prop-matching with a TERM item (used to justify equalities). *) val term_prop_matcher = let fun pre_match pat {tname, ...} ctxt = if not (Util.has_vars pat) then false - else if is_eq_term pat then - Matcher.pre_match ctxt (fst (dest_eq pat), the_single tname) + else if UtilBase.is_eq_term pat then + Matcher.pre_match ctxt (fst (UtilBase.dest_eq pat), the_single tname) else false fun match pat {tname, ...} ctxt (id, inst) = - if not (is_eq_term pat) then [] else + if not (UtilBase.is_eq_term pat) then [] else if not (Util.has_vars pat) then [] else let - val (lhs, rhs) = dest_eq pat + val (lhs, rhs) = UtilBase.dest_eq pat val cu = the_single tname val pairs = if Term.is_Var lhs then [(false, (lhs, cu)), (true, (rhs, cu))] else [(true, (lhs, cu)), (false, (rhs, cu))] val insts' = Matcher.rewrite_match_list ctxt pairs (id, inst) fun process_inst (inst, ths) = let (* th1: lhs(env) == u, th2: rhs(env) == u. *) val (th1, th2) = the_pair ths in - (inst, to_obj_eq (Util.transitive_list [th1, meta_sym th2])) + (inst, UtilLogic.to_obj_eq (Util.transitive_list [th1, meta_sym th2])) end in map process_inst insts' end in {pre_match = pre_match, match = match} end val null_eq_matcher = let - fun pre_match pat _ _ = is_eq_term pat + fun pre_match pat _ _ = UtilBase.is_eq_term pat fun match pat _ ctxt (id, inst) = - if not (is_eq_term pat) then [] else + if not (UtilBase.is_eq_term pat) then [] else if Util.has_vars pat then [] else let - val (lhs, rhs) = dest_eq pat + val (lhs, rhs) = UtilBase.dest_eq pat val infos = RewriteTable.equiv_info_t ctxt id (lhs, rhs) fun process_info (id', th) = - ((id', inst), to_obj_eq th) + ((id', inst), UtilLogic.to_obj_eq th) in map process_info infos end in {pre_match = pre_match, match = match} end (* Typed matching with a TERM item. *) val term_typed_matcher = let fun pre_match pat {tname, ...} ctxt = Matcher.pre_match_head ctxt (pat, the_single tname) (* Return value is (inst, eq), where eq is pat(inst) == tname. *) fun match pat {tname, ...} ctxt (id, inst) = Matcher.rewrite_match_head ctxt (pat, the_single tname) (id, inst) in {pre_match = pre_match, match = match} end (* Typed matching for items representing an equality ?A = ?B, where the tname is the pair (?A, ?B). Pattern is expected to be of the form ?A = ?B. *) val eq_tname_typed_matcher = let fun pre_match pat {tname, ...} ctxt = let val thy = Proof_Context.theory_of ctxt val (lhs, rhs) = the_pair (map Thm.term_of tname) - val _ = Pattern.first_order_match thy (pat, mk_eq (lhs, rhs)) fo_init + val _ = Pattern.first_order_match thy (pat, UtilBase.mk_eq (lhs, rhs)) fo_init in true end handle Pattern.MATCH => false fun match pat {tname, ...} ctxt (id, inst) = let val thy = Proof_Context.theory_of ctxt val (lhs, rhs) = the_pair (map Thm.term_of tname) val inst' = - Pattern.first_order_match thy (pat, mk_eq (lhs, rhs)) inst + Pattern.first_order_match thy (pat, UtilBase.mk_eq (lhs, rhs)) inst in - [((id, inst'), true_th)] + [((id, inst'), UtilBase.true_th)] end handle Pattern.MATCH => [] in {pre_match = pre_match, match = match} end (* Obtain a proposition from the property table. *) val null_property_matcher = let fun pre_match pat _ _ = Property.is_property pat fun match pat _ ctxt (id, inst) = if Util.has_vars pat then [] else if not (Property.is_property pat) then [] else map (fn (id', th) => ((id', inst), th)) (PropertyData.get_property_t ctxt (id, pat)) in {pre_match = pre_match, match = match} end (* Obtain a proposition from the property table, matching the argument of the property with the given term. *) val term_property_matcher = let fun pre_match pat {tname, ...} ctxt = Property.is_property pat andalso Matcher.pre_match_head ctxt ( Property.get_property_arg pat, the_single tname) fun match pat {tname, ...} ctxt (id, inst) = if not (Util.has_vars pat) then [] else if not (Property.is_property pat) then [] else let val arg = Property.get_property_arg pat in let val insts' = Matcher.rewrite_match_head ctxt (arg, the_single tname) (id, inst) fun process_inst ((id', inst'), _) = let val t = Util.subst_term_norm inst' pat in map (fn (id'', th) => ((id'', inst'), th)) (PropertyData.get_property_t ctxt (id', t)) end in maps process_inst insts' end end in {pre_match = pre_match, match = match} end (* Generic pre-matching function. Returns whether there is a possible match among any of the registered matchers. *) fun pre_match_arg ctxt arg (item as {ty_str, ...}) = if not (check_ty_str ty_str arg) then false else let val _ = assert_valid_arg arg val thy = Proof_Context.theory_of ctxt val {prop_matchers, typed_matchers, ...} = get_io_info thy ty_str in case arg of PropMatch pat => Term.is_Var pat orelse - (is_neg pat andalso Term.is_Var (get_neg pat)) orelse + (UtilLogic.is_neg pat andalso Term.is_Var (UtilLogic.get_neg pat)) orelse not (Util.is_pattern pat) orelse exists (fn f => f pat item ctxt) (map (#pre_match o fst) prop_matchers) | TypedMatch (_, pat) => not (Util.is_pattern pat) orelse exists (fn f => f pat item ctxt) (map (#pre_match o fst) typed_matchers) | TypedUniv _ => true | PropertyMatch _ => raise Fail "pre_match_arg: should not be called on PropertyMatch." | WellFormMatch _ => raise Fail "pre_match_arg: should not be called on WellFormMatch." end (* Generic matching function. Returns list of all matches (iterating over all registered matchers for the given item type. Note box_id for item is taken into account here. *) fun match_arg ctxt arg (item as {id, ty_str, ...}) (id', inst) = if not (check_ty_str ty_str arg) then [] else let val _ = assert_valid_arg arg val thy = Proof_Context.theory_of ctxt val {prop_matchers, typed_matchers, ...} = get_io_info thy ty_str val pat = pat_of_match_arg arg val id'' = BoxID.merge_boxes ctxt (id, id') in case arg of PropMatch _ => maps (fn f => f pat item ctxt (id'', inst)) (map (#match o fst) prop_matchers) - | TypedUniv _ => [((id'', inst), true_th)] + | TypedUniv _ => [((id'', inst), UtilBase.true_th)] | TypedMatch _ => maps (fn f => f pat item ctxt (id'', inst)) (map (#match o fst) typed_matchers) | PropertyMatch _ => raise Fail "match_arg: should not be called on PropertyMatch." | WellFormMatch _ => raise Fail "match_arg: should not be called on WellFormMatch." end val no_rewr_terms = K [] fun arg_rewr_terms ts = maps Util.dest_args ts fun prop_rewr_terms ts = let val t = the_single ts in - if is_neg t then t |> dest_arg |> Util.dest_args + if UtilLogic.is_neg t then t |> dest_arg |> Util.dest_args else t |> Util.dest_args end fun rewr_terms_of_item ctxt (ty_str, tname) = let val thy = Proof_Context.theory_of ctxt val {term_fn, ...} = get_io_info thy ty_str in case term_fn of NONE => tname | SOME (f, _) => f tname end fun output_prop_fn ctxt (_, th) = Thm.prop_of th |> Syntax.string_of_term ctxt fun string_of_item_info ctxt (ty_str, ts, th) = let val thy = Proof_Context.theory_of ctxt val {output_fn, ...} = get_io_info thy ty_str in case output_fn of NONE => ty_str ^ " " ^ (Util.string_of_terms ctxt ts) | SOME (f, _) => f ctxt (ts, th) end fun string_of_raw_item ctxt ritem = case ritem of Handler (_, t, _) => "Handler " ^ (Syntax.string_of_term ctxt t) | Fact info => string_of_item_info ctxt info fun string_of_item ctxt {ty_str, tname, prop, ...} = string_of_item_info ctxt (ty_str, map Thm.term_of tname, prop) fun trace_ritem ctxt s ritem = tracing (s ^ " " ^ (string_of_raw_item ctxt ritem)) fun trace_ritems ctxt s ritems = tracing (s ^ "\n" ^ (cat_lines (map (string_of_raw_item ctxt) ritems))) fun trace_item ctxt s item = tracing (s ^ " " ^ (string_of_item ctxt item)) fun trace_items ctxt s items = tracing (s ^ "\n" ^ (cat_lines (map (string_of_item ctxt) items))) (* We assume ts1 is the new item at the given id, while cts2 is for an existing item, at some previous id. *) fun shadow_prop_fn ctxt id (ts1, cts2) = let val (t1, ct2) = (the_single ts1, the_single cts2) - val (lhs, crhs) = if is_neg t1 andalso is_neg (Thm.term_of ct2) then - (get_neg t1, UtilLogic.get_cneg ct2) + val (lhs, crhs) = if UtilLogic.is_neg t1 andalso UtilLogic.is_neg (Thm.term_of ct2) then + (UtilLogic.get_neg t1, UtilLogic.get_cneg ct2) else (t1, ct2) in (Matcher.rewrite_match_head ctxt (lhs, crhs) (id, fo_init)) |> filter (fn ((id', _), _) => id' = id) |> not o null end fun shadow_term_fn ctxt id (ts1, cts2) = let val (lhs, crhs) = (the_single ts1, the_single cts2) in (Matcher.rewrite_match_head ctxt (lhs, crhs) (id, fo_init)) |> filter (fn ((id', _), _) => id' = id) |> not o null end val add_basic_item_io = fold add_item_type [ (TY_NULL, NONE, NONE, NONE), (TY_PROP, SOME prop_rewr_terms, SOME output_prop_fn, SOME shadow_prop_fn), (TY_TERM, SOME no_rewr_terms, NONE, SOME shadow_term_fn), (TY_EQ, NONE, SOME output_prop_fn, NONE), (TY_VAR, NONE, NONE, NONE), (TY_PROPERTY, SOME arg_rewr_terms, NONE, NONE) ] #> fold add_prop_matcher [ (TY_PROP, prop_matcher), (TY_TERM, term_prop_matcher), (TY_NULL, null_eq_matcher), (TY_NULL, null_property_matcher), (TY_TERM, term_property_matcher) ] #> fold add_typed_matcher [ (TY_PROP, prop_matcher), (TY_TERM, term_typed_matcher), (TY_VAR, term_typed_matcher), (TY_EQ, eq_tname_typed_matcher) ] end (* structure ItemIO. *) diff --git a/thys/Auto2_HOL/logic_steps.ML b/thys/Auto2_HOL/logic_steps.ML --- a/thys/Auto2_HOL/logic_steps.ML +++ b/thys/Auto2_HOL/logic_steps.ML @@ -1,1081 +1,1087 @@ (* File: logic_steps.ML Author: Bohua Zhan Core (logic) proofsteps. *) signature LOGIC_PROOFSTEPS = sig (* General logic *) val shadow_prop_item: proofstep val shadow_term_item: proofstep val exists_elim_prfstep: proofstep val add_logic_proofsteps: theory -> theory val mk_all_disj: term list * term list -> term val strip_all_disj: term -> term list * term list val norm_all_disj: Proof.context -> conv val replace_disj_vars: Proof.context -> term list * term list -> term list * term list val disj_prop_match: Proof.context -> id_inst -> term * (term list * term list) * ((indexname * typ) list * cterm list) -> id_inst_th list val norm_conj: conv (* DISJ items. *) val TY_DISJ: string val disj_to_ritems: bool -> term -> thm -> raw_item list val disj_to_update: bool -> term -> box_id * int option * thm -> raw_update val dest_tname_of_disj: cterm list -> term * cterm list val is_match_prem_only: box_item -> bool val analyze_disj_th: Proof.context -> thm -> term * thm val disj_rewr_terms: term list -> term list val output_disj_fn: item_output val disj_prop_matcher: item_matcher val reduce_disj_True: conv val match_update_prfstep: proofstep val match_one_sch_prfstep: proofstep val disj_match_iff_prfstep: proofstep val disj_create_case_prfstep: proofstep val disj_shadow_prfstep: proofstep val add_disj_proofsteps: theory -> theory (* Normalizers *) val split_not_imp_th: thm -> thm list val split_conj_gen_th: Proof.context -> thm -> thm list - val eq_normalizer: Normalizer.normalizer - val property_normalizer: Normalizer.normalizer - val disj_normalizer: Normalizer.normalizer + val eq_normalizer: Normalizer_Type.normalizer + val property_normalizer: Normalizer_Type.normalizer + val disj_normalizer: Normalizer_Type.normalizer val logic_thm_update: Proof.context -> box_id * thm -> raw_update val add_disj_normalizers: theory -> theory end; -structure Logic_ProofSteps : LOGIC_PROOFSTEPS = +functor Logic_ProofSteps( + structure BoxItem: BOXITEM; + structure ItemIO: ITEM_IO; + structure Matcher: MATCHER; + structure Normalizer: NORMALIZER; + structure Property: PROPERTY; + structure ProofStepData: PROOFSTEP_DATA; + structure RewriteTable: REWRITE_TABLE; + structure UtilBase: UTIL_BASE; + structure UtilLogic: UTIL_LOGIC; + structure Update: UPDATE; + ): LOGIC_PROOFSTEPS = struct -fun boolVar s = Var ((s, 0), boolT) +fun boolVar s = Var ((s, 0), UtilBase.boolT) (* Shadowing based on equivalence. For both PROP and TERM items, shadowing is based on subterm equivalence, skipping any Not (~) at head. *) fun shadow_item_fn ctxt item1 item2 = if #sc item1 = 0 andalso #sc item2 = 0 then [] else let val id = BoxItem.merged_id ctxt [item1, item2] val _ = assert (forall (BoxItem.match_ty_strs [TY_TERM, TY_PROP]) [item1, item2]) "shadow_item_fn" val (tname1, tname2) = (the_single (#tname item1), the_single (#tname item2)) handle List.Empty => raise Fail "shadow_item_fn" val (t1, t2) = (Thm.term_of tname1, Thm.term_of tname2) val (lhs, rhs) = - if fastype_of t1 = boolT andalso is_neg t1 andalso - fastype_of t2 = boolT andalso is_neg t2 then + if fastype_of t1 = UtilBase.boolT andalso UtilLogic.is_neg t1 andalso + fastype_of t2 = UtilBase.boolT andalso UtilLogic.is_neg t2 then (UtilLogic.get_cneg tname1, UtilLogic.get_cneg tname2) else (tname1, tname2) val equiv_ids = (RewriteTable.subequiv_info ctxt id (lhs, rhs)) |> map fst |> filter BoxID.has_incr_id |> Util.max_partial (BoxID.is_eq_ancestor ctxt) val item_to_shadow = if #uid item1 > #uid item2 then item1 else item2 fun process_id id' = ShadowItem {id = id', item = item_to_shadow} in map process_id equiv_ids end val shadow_prop_item = {name = "shadow_prop", args = [TypedUniv TY_PROP, TypedUniv TY_PROP], func = TwoStep shadow_item_fn} val shadow_term_item = {name = "shadow_term", args = [TypedUniv TY_TERM, TypedUniv TY_TERM], func = TwoStep shadow_item_fn} fun eq_abs_fn ctxt item1 item2 = let val {id = id1, tname = tname1, ...} = item1 val {id = id2, tname = tname2, ...} = item2 val id = BoxID.merge_boxes ctxt (id1, id2) val (ct1, ct2) = apply2 the_single (tname1, tname2) val (t1, t2) = apply2 Thm.term_of (ct1, ct2) in if not (Util.is_abs t1) orelse not (Util.is_abs t2) then [] else if RewriteTable.is_equiv id ctxt (ct1, ct2) then [] else if Term_Ord.term_ord (t2, t1) = LESS then [] else let fun process_equiv (id', eq_th) = let val (lhs, rhs) = (Util.lhs_of eq_th, Util.rhs_of eq_th) in AddItems {id = id', sc = SOME 1, raw_items = [Fact (TY_EQ, [lhs, rhs], eq_th)]} end in (Matcher.rewrite_match ctxt (t1, ct2) (id, fo_init)) |> map (fn ((id, _), eq_th) => (id, eq_th)) |> filter (BoxID.has_incr_id o fst) |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) |> map process_equiv end end val eq_abs_prfstep = {name = "eq_abs", args = [TypedUniv TY_TERM, TypedUniv TY_TERM], func = TwoStep eq_abs_fn} (* Given an assumption of the form EX x. A, we produce an assumption A with x in A replaced by a free variable. To avoid name collisions, when the update is produced x is replaced by an "internal" free variable, with suffix '_'. When the update is applied, that internal free variable is replaced by a fresh variable as determined by the context. We produce at most two variables at a time. *) fun exists_elim_fn ctxt {id, prop, ...} = if not (BoxID.has_incr_id id) then [] else let - val t = prop_of' prop + val t = UtilLogic.prop_of' prop in - if is_ex t orelse is_bex t orelse - (is_neg t andalso is_obj_all (dest_not t)) orelse - (is_neg t andalso is_ball (dest_not t)) then + if UtilLogic.is_ex t orelse UtilLogic.is_bex t orelse + (UtilLogic.is_neg t andalso UtilLogic.is_obj_all (UtilLogic.dest_not t)) orelse + (UtilLogic.is_neg t andalso UtilLogic.is_ball (UtilLogic.dest_not t)) then let val (ritems, th') = - prop |> apply_to_thm' (UtilLogic.normalize_exists ctxt) + prop |> UtilLogic.apply_to_thm' (UtilLogic.normalize_exists ctxt) |> Update.apply_exists_ritems ctxt in [AddItems {id = id, sc = NONE, raw_items = ritems @ [Update.thm_to_ritem th']}] end else [] end val exists_elim_prfstep = {name = "exists_elim", args = [TypedUniv TY_PROP], func = OneStep exists_elim_fn} val add_logic_proofsteps = - fold add_prfstep [ + fold ProofStepData.add_prfstep [ shadow_prop_item, shadow_term_item, exists_elim_prfstep, eq_abs_prfstep ] (* Given (x_1 ... x_n, A_1 ... A_n), create the corresponding forall-disj term. *) fun mk_all_disj (vars, terms) = case vars of - [] => list_disj terms - | var :: vars' => mk_obj_all var (mk_all_disj (vars', terms)) + [] => UtilLogic.list_disj terms + | var :: vars' => UtilLogic.mk_obj_all var (mk_all_disj (vars', terms)) (* Normalize t into a disjunction of terms. *) fun strip_disj t = - if is_disj t then + if UtilLogic.is_disj t then maps strip_disj [dest_arg1 t, dest_arg t] - else if is_imp t then - maps strip_disj [get_neg (dest_arg1 t), dest_arg t] - else if is_neg t then + else if UtilLogic.is_imp t then + maps strip_disj [UtilLogic.get_neg (dest_arg1 t), dest_arg t] + else if UtilLogic.is_neg t then let - val t' = dest_not t + val t' = UtilLogic.dest_not t in - if is_neg t' then strip_disj (dest_not t') - else if is_conj t' then - maps strip_disj [get_neg (dest_arg1 t'), get_neg (dest_arg t')] + if UtilLogic.is_neg t' then strip_disj (UtilLogic.dest_not t') + else if UtilLogic.is_conj t' then + maps strip_disj [UtilLogic.get_neg (dest_arg1 t'), UtilLogic.get_neg (dest_arg t')] else [t] end else [t] (* Normalize a term into the form !x_1 ... x_n. A_1 | ... | A_n *) fun strip_all_disj t = - if is_obj_all t then + if UtilLogic.is_obj_all t then case t of _ $ (u as Abs _) => let val (v, body) = Term.dest_abs_global u val var = Free v val (vars, disjs) = strip_all_disj body in if exists (Util.occurs_free var) disjs then (var :: vars, disjs) else (vars, disjs) end | f $ arg => strip_all_disj (f $ UtilLogic.force_abs_form arg) | _ => raise Fail "strip_all_disj" - else if is_ball t then + else if UtilLogic.is_ball t then case t of _ $ S $ (u as Abs _) => let val (v, body) = Term.dest_abs_global u val var = Free v - val mem = mk_mem (var, S) + val mem = UtilLogic.mk_mem (var, S) val (vars, disjs) = strip_all_disj body in - (var :: vars, get_neg mem :: disjs) + (var :: vars, UtilLogic.get_neg mem :: disjs) end | f $ S $ arg => strip_all_disj (f $ S $ UtilLogic.force_abs_form arg) | _ => raise Fail "strip_all_disj" - else if is_neg t andalso is_ex (dest_not t) then - case dest_not t of + else if UtilLogic.is_neg t andalso UtilLogic.is_ex (UtilLogic.dest_not t) then + case UtilLogic.dest_not t of _ $ (u as Abs _) => let val (v, body) = Term.dest_abs_global u val var = Free v - val (vars, disjs) = strip_all_disj (get_neg body) + val (vars, disjs) = strip_all_disj (UtilLogic.get_neg body) in if exists (Util.occurs_free var) disjs then (var :: vars, disjs) else (vars, disjs) end - | f $ arg => strip_all_disj (get_neg (f $ UtilLogic.force_abs_form arg)) + | f $ arg => strip_all_disj (UtilLogic.get_neg (f $ UtilLogic.force_abs_form arg)) | _ => raise Fail "strip_all_disj" - else if is_neg t andalso is_bex (dest_not t) then - case dest_not t of + else if UtilLogic.is_neg t andalso UtilLogic.is_bex (UtilLogic.dest_not t) then + case UtilLogic.dest_not t of _ $ S $ (u as Abs _) => let val (v, body) = Term.dest_abs_global u val var = Free v - val mem = mk_mem (var, S) - val (vars, disjs) = strip_all_disj (get_neg body) + val mem = UtilLogic.mk_mem (var, S) + val (vars, disjs) = strip_all_disj (UtilLogic.get_neg body) in - (var :: vars, get_neg mem :: disjs) + (var :: vars, UtilLogic.get_neg mem :: disjs) end - | f $ S $ arg => strip_all_disj (get_neg (f $ S $ UtilLogic.force_abs_form arg)) + | f $ S $ arg => strip_all_disj (UtilLogic.get_neg (f $ S $ UtilLogic.force_abs_form arg)) | _ => raise Fail "strip_all_disj" - else if is_disj t then + else if UtilLogic.is_disj t then let val (v1, ts1) = strip_all_disj (dest_arg1 t) val (v2, ts2) = strip_all_disj (dest_arg t) in (v1 @ v2, ts1 @ ts2) end - else if is_imp t then + else if UtilLogic.is_imp t then let - val (v1, ts1) = strip_all_disj (get_neg (dest_arg1 t)) + val (v1, ts1) = strip_all_disj (UtilLogic.get_neg (dest_arg1 t)) val (v2, ts2) = strip_all_disj (dest_arg t) in (v1 @ v2, ts1 @ ts2) end - else if is_neg t then + else if UtilLogic.is_neg t then let - val t' = dest_not t + val t' = UtilLogic.dest_not t in - if is_neg t' then strip_all_disj (dest_not t') - else if is_conj t' then + if UtilLogic.is_neg t' then strip_all_disj (UtilLogic.dest_not t') + else if UtilLogic.is_conj t' then let - val (v1, ts1) = strip_all_disj (get_neg (dest_arg1 t')) - val (v2, ts2) = strip_all_disj (get_neg (dest_arg t')) + val (v1, ts1) = strip_all_disj (UtilLogic.get_neg (dest_arg1 t')) + val (v2, ts2) = strip_all_disj (UtilLogic.get_neg (dest_arg t')) in (v1 @ v2, ts1 @ ts2) end else ([], [t]) end else ([], [t]) (* Normalize (A_1 | A_2 | ... | A_m) | (B_1 | B_2 | ... | B_n) *) fun norm_disj_clauses ct = let val (arg1, _) = Util.dest_binop_args (Thm.term_of ct) in - if is_disj arg1 then - Conv.every_conv [rewr_obj_eq UtilBase.disj_assoc_th, + if UtilLogic.is_disj arg1 then + Conv.every_conv [UtilLogic.rewr_obj_eq UtilBase.disj_assoc_th, Conv.arg_conv norm_disj_clauses] ct else Conv.all_conv ct end (* Normalize ct. *) fun norm_disj ct = let val t = Thm.term_of ct - val _ = assert (fastype_of t = boolT) "norm_disj: wrong type" + val _ = assert (fastype_of t = UtilBase.boolT) "norm_disj: wrong type" in - if is_disj t then + if UtilLogic.is_disj t then Conv.every_conv [Conv.binop_conv norm_disj, norm_disj_clauses] ct - else if is_imp t then - Conv.every_conv [rewr_obj_eq UtilBase.imp_conv_disj_th, norm_disj] ct - else if is_neg t andalso is_neg (dest_not t) then - Conv.every_conv [rewr_obj_eq UtilBase.nn_cancel_th, norm_disj] ct - else if is_neg t andalso is_conj (dest_not t) then - Conv.every_conv [rewr_obj_eq UtilBase.de_Morgan_conj_th, norm_disj] ct + else if UtilLogic.is_imp t then + Conv.every_conv [UtilLogic.rewr_obj_eq UtilBase.imp_conv_disj_th, norm_disj] ct + else if UtilLogic.is_neg t andalso UtilLogic.is_neg (UtilLogic.dest_not t) then + Conv.every_conv [UtilLogic.rewr_obj_eq UtilBase.nn_cancel_th, norm_disj] ct + else if UtilLogic.is_neg t andalso UtilLogic.is_conj (UtilLogic.dest_not t) then + Conv.every_conv [UtilLogic.rewr_obj_eq UtilBase.de_Morgan_conj_th, norm_disj] ct else Conv.all_conv ct end (* Normalize to forall at the top-level *) fun norm_all_disj ctxt ct = let val t = Thm.term_of ct in - if is_obj_all t then + if UtilLogic.is_obj_all t then if not (Util.is_subterm (Bound 0) (dest_arg t)) then - Conv.every_conv [rewr_obj_eq UtilBase.all_trivial_th, + Conv.every_conv [UtilLogic.rewr_obj_eq UtilBase.all_trivial_th, norm_all_disj ctxt] ct else Conv.every_conv [Conv.binder_conv (norm_all_disj o snd) ctxt] ct - else if is_ball t then - Conv.every_conv [rewr_obj_eq UtilBase.Ball_def_th, norm_all_disj ctxt] ct - else if is_neg t andalso is_ex (dest_not t) then - Conv.every_conv [rewr_obj_eq UtilBase.not_ex_th, norm_all_disj ctxt] ct - else if is_neg t andalso is_bex (dest_not t) then + else if UtilLogic.is_ball t then + Conv.every_conv [UtilLogic.rewr_obj_eq UtilBase.Ball_def_th, norm_all_disj ctxt] ct + else if UtilLogic.is_neg t andalso UtilLogic.is_ex (UtilLogic.dest_not t) then + Conv.every_conv [UtilLogic.rewr_obj_eq UtilBase.not_ex_th, norm_all_disj ctxt] ct + else if UtilLogic.is_neg t andalso UtilLogic.is_bex (UtilLogic.dest_not t) then Conv.every_conv [ - Conv.arg_conv (rewr_obj_eq UtilBase.Bex_def_th), norm_all_disj ctxt] ct - else if is_disj t then + Conv.arg_conv (UtilLogic.rewr_obj_eq UtilBase.Bex_def_th), norm_all_disj ctxt] ct + else if UtilLogic.is_disj t then let val eq1 = Conv.binop_conv (norm_all_disj ctxt) ct val rhs = Util.rhs_of eq1 in - if is_obj_all (dest_arg1 rhs) then + if UtilLogic.is_obj_all (dest_arg1 rhs) then Conv.every_conv [Conv.rewr_conv eq1, - rewr_obj_eq UtilBase.disj_commute_th, - rewr_obj_eq UtilBase.swap_all_disj_th, + UtilLogic.rewr_obj_eq UtilBase.disj_commute_th, + UtilLogic.rewr_obj_eq UtilBase.swap_all_disj_th, norm_all_disj ctxt] ct - else if is_obj_all (dest_arg rhs) then + else if UtilLogic.is_obj_all (dest_arg rhs) then Conv.every_conv [Conv.rewr_conv eq1, - rewr_obj_eq UtilBase.swap_all_disj_th, + UtilLogic.rewr_obj_eq UtilBase.swap_all_disj_th, norm_all_disj ctxt] ct else Conv.every_conv [Conv.rewr_conv eq1, norm_disj_clauses] ct end - else if is_imp t then - Conv.every_conv [rewr_obj_eq UtilBase.imp_conv_disj_th, norm_all_disj ctxt] ct - else if is_neg t andalso is_neg (dest_not t) then - Conv.every_conv [rewr_obj_eq UtilBase.nn_cancel_th, norm_all_disj ctxt] ct - else if is_neg t andalso is_conj (dest_not t) then - Conv.every_conv [rewr_obj_eq UtilBase.de_Morgan_conj_th, norm_all_disj ctxt] ct + else if UtilLogic.is_imp t then + Conv.every_conv [UtilLogic.rewr_obj_eq UtilBase.imp_conv_disj_th, norm_all_disj ctxt] ct + else if UtilLogic.is_neg t andalso UtilLogic.is_neg (UtilLogic.dest_not t) then + Conv.every_conv [UtilLogic.rewr_obj_eq UtilBase.nn_cancel_th, norm_all_disj ctxt] ct + else if UtilLogic.is_neg t andalso UtilLogic.is_conj (UtilLogic.dest_not t) then + Conv.every_conv [UtilLogic.rewr_obj_eq UtilBase.de_Morgan_conj_th, norm_all_disj ctxt] ct else Conv.all_conv ct end fun mk_disj_eq eq_ths = case eq_ths of [] => raise Fail "mk_disj_eq" | [eq] => eq | eq :: eqs' => Drule.binop_cong_rule UtilBase.cDisj eq (mk_disj_eq eqs') (* Sort A | (B_1 | B_2 | ... | B_n), assuming the right side is sorted. *) fun sort_disj_clause_aux ct = - if not (is_disj (Thm.term_of ct)) then Conv.all_conv ct else + if not (UtilLogic.is_disj (Thm.term_of ct)) then Conv.all_conv ct else let val (arg1, arg2) = Util.dest_binop_args (Thm.term_of ct) in - if is_disj arg2 then + if UtilLogic.is_disj arg2 then if Term_Ord.term_ord (dest_arg1 arg2, arg1) = LESS then - Conv.every_conv [rewr_obj_eq (obj_sym UtilBase.disj_assoc_th), - Conv.arg1_conv (rewr_obj_eq UtilBase.disj_commute_th), - rewr_obj_eq UtilBase.disj_assoc_th, + Conv.every_conv [UtilLogic.rewr_obj_eq (UtilLogic.obj_sym UtilBase.disj_assoc_th), + Conv.arg1_conv (UtilLogic.rewr_obj_eq UtilBase.disj_commute_th), + UtilLogic.rewr_obj_eq UtilBase.disj_assoc_th, Conv.arg_conv sort_disj_clause_aux] ct else Conv.all_conv ct else if Term_Ord.term_ord (arg2, arg1) = LESS then - rewr_obj_eq (obj_sym UtilBase.disj_commute_th) ct + UtilLogic.rewr_obj_eq (UtilLogic.obj_sym UtilBase.disj_commute_th) ct else Conv.all_conv ct end (* Sort A_1 | ... | A_n. *) fun sort_disj_clause ct = - if not (is_disj (Thm.term_of ct)) then Conv.all_conv ct + if not (UtilLogic.is_disj (Thm.term_of ct)) then Conv.all_conv ct else Conv.every_conv [Conv.arg_conv sort_disj_clause, sort_disj_clause_aux] ct (* Apply cv on the body of !x y z. ... .*) fun forall_body_conv cv ctxt ct = - if is_obj_all (Thm.term_of ct) then + if UtilLogic.is_obj_all (Thm.term_of ct) then Conv.binder_conv (fn (_, ctxt) => forall_body_conv cv ctxt) ctxt ct else cv ct fun norm_all_disj_sorted ctxt ct = Conv.every_conv [norm_all_disj ctxt, forall_body_conv sort_disj_clause ctxt] ct fun abstract_eq ctxt var eq = let val (x, T) = Term.dest_Free var - val all_const = Const (UtilBase.All_name, (T --> boolT) --> boolT) + val all_const = Const (UtilBase.All_name, (T --> UtilBase.boolT) --> UtilBase.boolT) in Drule.arg_cong_rule (Thm.cterm_of ctxt all_const) (Thm.abstract_rule x (Thm.cterm_of ctxt var) eq) end fun replace_disj_vars ctxt (vars, disjs) = let val vars' = vars |> map Term.dest_Free |> Variable.variant_frees ctxt [] |> map Free val subst = vars ~~ vars' in (vars', map (subst_atomic subst) disjs) end (* Matching for all-disj propositions *) fun disj_prop_match ctxt (id, (tyinst, inst)) (t, (var_t, ts), (var_u, cus)) = let val thy = Proof_Context.theory_of ctxt val us = map Thm.term_of cus in if length var_t <> length var_u orelse length ts <> length us then [] else let (* First match the types (return [] if no match). *) val tys_t = map fastype_of var_t val tys_u = map snd var_u val tyinst' = fold (Sign.typ_match thy) (tys_t ~~ tys_u) tyinst val var_t' = map (Envir.subst_term_types tyinst') var_t val ts' = ts |> map (Envir.subst_term_types tyinst') val var_ct' = map (Thm.cterm_of ctxt) var_t' val cus' = cus |> map (Thm.instantiate_cterm (TVars.empty, Vars.make (var_u ~~ var_ct'))) (* Match the type-instantiated pattern with term. *) val insts = Matcher.rewrite_match_subset ctxt var_t' (ts', cus') (id, (tyinst', inst)) fun process_inst ((id', instsp'), ths) = let (* Equality between normalized t and u *) val eq_th = ths |> mk_disj_eq |> fold (abstract_eq ctxt) (rev var_t') |> apply_to_lhs (norm_all_disj_sorted ctxt) |> apply_to_rhs (norm_all_disj_sorted ctxt) (* Equality between un-normalized t and u *) val t' = Util.subst_term_norm instsp' t val norm1 = norm_all_disj_sorted ctxt (Thm.cterm_of ctxt t') val cu = Thm.cterm_of ctxt (mk_all_disj (var_t', map Thm.term_of cus')) val norm2 = norm_all_disj_sorted ctxt cu val eq_th' = Util.transitive_list [norm1, eq_th, meta_sym norm2] in ((id', instsp'), eq_th') end in map process_inst insts end handle Type.TYPE_MATCH => [] end fun norm_conj_de_Morgan ct = let val t = Thm.term_of ct in - if is_neg t andalso is_disj (dest_not t) then - Conv.every_conv [rewr_obj_eq UtilBase.de_Morgan_disj_th, + if UtilLogic.is_neg t andalso UtilLogic.is_disj (UtilLogic.dest_not t) then + Conv.every_conv [UtilLogic.rewr_obj_eq UtilBase.de_Morgan_disj_th, Conv.arg_conv norm_conj_de_Morgan] ct else Conv.all_conv ct end fun norm_conj_not_imp ct = let val t = Thm.term_of ct in - if is_neg t andalso is_imp (dest_not t) then - Conv.every_conv [rewr_obj_eq UtilBase.not_imp_th, + if UtilLogic.is_neg t andalso UtilLogic.is_imp (UtilLogic.dest_not t) then + Conv.every_conv [UtilLogic.rewr_obj_eq UtilBase.not_imp_th, Conv.arg_conv norm_conj_not_imp] ct else Conv.all_conv ct end fun norm_nn_cancel ct = let val t = Thm.term_of ct in - if is_neg t andalso is_neg (dest_not t) then - Conv.every_conv [rewr_obj_eq UtilBase.nn_cancel_th, + if UtilLogic.is_neg t andalso UtilLogic.is_neg (UtilLogic.dest_not t) then + Conv.every_conv [UtilLogic.rewr_obj_eq UtilBase.nn_cancel_th, norm_nn_cancel] ct else Conv.all_conv ct end fun norm_conj ct = let val t = Thm.term_of ct - val _ = assert (fastype_of t = boolT) "norm_conj: wrong type" + val _ = assert (fastype_of t = UtilBase.boolT) "norm_conj: wrong type" in - if is_neg t andalso is_neg (dest_not t) then + if UtilLogic.is_neg t andalso UtilLogic.is_neg (UtilLogic.dest_not t) then norm_nn_cancel ct - else if is_neg t andalso is_imp (dest_not t) then + else if UtilLogic.is_neg t andalso UtilLogic.is_imp (UtilLogic.dest_not t) then norm_conj_not_imp ct - else if is_neg t andalso is_disj (dest_not t) then + else if UtilLogic.is_neg t andalso UtilLogic.is_disj (UtilLogic.dest_not t) then norm_conj_de_Morgan ct else Conv.all_conv ct end (* DISJ items. *) val TY_DISJ = "DISJ" (* Given a theorem in the form of a disjunction, possibly containing schematic variables, return the corresponding DISJ item. *) fun disj_to_ritems prem_only disj_head th = let - val subs = strip_disj (prop_of' th) + val subs = strip_disj (UtilLogic.prop_of' th) in if length subs = 1 then if Util.has_vars (Thm.prop_of th) then let fun th_to_ritem th = Fact (TY_DISJ, - UtilLogic.term_of_bool prem_only :: disj :: strip_disj (prop_of' th), + UtilLogic.term_of_bool prem_only :: UtilLogic.disj :: strip_disj (UtilLogic.prop_of' th), th) in - map th_to_ritem (th |> apply_to_thm' norm_conj + map th_to_ritem (th |> UtilLogic.apply_to_thm' norm_conj |> UtilLogic.split_conj_th) end else - [Fact (TY_PROP, [prop_of' th], th)] + [Fact (TY_PROP, [UtilLogic.prop_of' th], th)] else let val tname = UtilLogic.term_of_bool prem_only :: disj_head :: subs in [Fact (TY_DISJ, tname, th)] end end fun disj_to_update prem_only disj_head (id, sc, th) = - if Thm.prop_of th aconv pFalse then + if Thm.prop_of th aconv UtilLogic.pFalse then ResolveBox {id = id, th = th} else AddItems {id = id, sc = sc, raw_items = disj_to_ritems prem_only disj_head th} fun get_disj_head t = - if is_disj t then disj - else if is_imp t then get_disj_head (dest_arg t) - else if is_neg t andalso is_neg (dest_not t) then - get_disj_head (dest_not (dest_not t)) - else if is_neg t andalso is_conj (dest_not t) then conj - else if is_obj_all t orelse is_ball t then + if UtilLogic.is_disj t then UtilLogic.disj + else if UtilLogic.is_imp t then get_disj_head (dest_arg t) + else if UtilLogic.is_neg t andalso UtilLogic.is_neg (UtilLogic.dest_not t) then + get_disj_head (UtilLogic.dest_not (UtilLogic.dest_not t)) + else if UtilLogic.is_neg t andalso UtilLogic.is_conj (UtilLogic.dest_not t) then UtilLogic.conj + else if UtilLogic.is_obj_all t orelse UtilLogic.is_ball t then case dest_arg t of u as Abs _ => get_disj_head (snd (Term.dest_abs_global u)) - | _ => imp - else if is_neg t andalso is_ex (dest_not t) then conj - else if is_neg t andalso is_bex (dest_not t) then conj - else imp + | _ => UtilLogic.imp + else if UtilLogic.is_neg t andalso UtilLogic.is_ex (UtilLogic.dest_not t) then UtilLogic.conj + else if UtilLogic.is_neg t andalso UtilLogic.is_bex (UtilLogic.dest_not t) then UtilLogic.conj + else UtilLogic.imp (* Given a theorem th, return equivalent theorem in disjunctive form, with possible schematic variables. Also return whether th is "active", that is, whether it is originally a conjunctive goal or disjunctive fact, as opposed to implications. *) fun analyze_disj_th ctxt th = let - val head = get_disj_head (prop_of' th) - val th' = th |> apply_to_thm' (norm_all_disj ctxt) - |> apply_to_thm (UtilLogic.to_meta_conv ctxt) + val head = get_disj_head (UtilLogic.prop_of' th) + val th' = th |> UtilLogic.apply_to_thm' (norm_all_disj ctxt) + |> Util.apply_to_thm (UtilLogic.to_meta_conv ctxt) |> Util.forall_elim_sch in (head, th') end (* Deconstruct the tname of a DISJ item. *) fun dest_tname_of_disj tname = case tname of _ :: disj_head :: rest => (Thm.term_of disj_head, rest) | _ => raise Fail "dest_tname_of_disj: too few terms in tname." (* Determine whether the item is for matching premises only (from the first entry in tname. *) fun is_match_prem_only {tname, ...} = UtilLogic.bool_of_term (Thm.term_of (hd tname)) -fun is_active_head disj_head = (not (disj_head aconv imp)) +fun is_active_head disj_head = (not (disj_head aconv UtilLogic.imp)) fun disj_rewr_terms ts = if UtilLogic.bool_of_term (hd ts) then [] else drop 2 ts fun output_disj_fn ctxt (ts, _) = let val (match_prem, disj_head, subs) = (hd ts, hd (tl ts), tl (tl ts)) val prefix = if UtilLogic.bool_of_term match_prem then "(match_prem) " else "" in - if disj_head aconv disj then - prefix ^ ((foldr1 mk_disj subs) |> Syntax.string_of_term ctxt) - else if disj_head aconv conj then - prefix ^ ((foldr1 mk_conj (map get_neg subs)) - |> get_neg |> Syntax.string_of_term ctxt) - else if disj_head aconv imp then - prefix ^ ((foldr1 mk_imp (subs |> split_last |> apfst (map get_neg) + if disj_head aconv UtilLogic.disj then + prefix ^ ((foldr1 UtilLogic.mk_disj subs) |> Syntax.string_of_term ctxt) + else if disj_head aconv UtilLogic.conj then + prefix ^ ((foldr1 UtilLogic.mk_conj (map UtilLogic.get_neg subs)) + |> UtilLogic.get_neg |> Syntax.string_of_term ctxt) + else if disj_head aconv UtilLogic.imp then + prefix ^ ((foldr1 UtilLogic.mk_imp (subs |> split_last |> apfst (map UtilLogic.get_neg) |> apsnd single |> (op @))) |> Syntax.string_of_term ctxt) else raise Fail "output_disj_fn: unexpected disj_head." end val disj_prop_matcher = let fun pre_match pat {tname, ...} ctxt = let val (var_t, ts) = (strip_all_disj pat) |> replace_disj_vars ctxt val (_, cus) = dest_tname_of_disj tname val us = map Thm.term_of cus val var_u = fold Term.add_vars us [] in length ts = length us andalso length var_t = length var_u end fun match pat item ctxt (id, instsp) = let val {tname, prop = th, ...} = item val (_, cus) = dest_tname_of_disj tname val us = map Thm.term_of cus val var_u = fold Term.add_vars us [] val (var_t, ts) = (strip_all_disj pat) |> replace_disj_vars ctxt fun process_perm perm = map (pair (map Var perm)) (disj_prop_match ctxt (id, instsp) (pat, (var_t, ts), (perm, cus))) fun process_inst (var_u, ((id', instsp'), eq_th)) = let - val eq_th' = make_trueprop_eq (meta_sym eq_th) + val eq_th' = UtilLogic.make_trueprop_eq (meta_sym eq_th) val forall_th = th |> fold Thm.forall_intr (rev (map (Thm.cterm_of ctxt) var_u)) - |> apply_to_thm (UtilLogic.to_obj_conv ctxt) + |> Util.apply_to_thm (UtilLogic.to_obj_conv ctxt) in ((id', instsp'), Thm.equal_elim eq_th' forall_th) end in if length var_t <> length var_u orelse length ts <> length us then [] else var_u |> Util.all_permutes |> maps process_perm |> map process_inst end in {pre_match = pre_match, match = match} end (* Given ct in the form p_1 | ... | p_n, apply cv to each of p_i. *) fun ac_disj_conv cv ct = - if is_disj (Thm.term_of ct) then + if UtilLogic.is_disj (Thm.term_of ct) then Conv.every_conv [Conv.arg1_conv cv, Conv.arg_conv (ac_disj_conv cv)] ct else cv ct (* Assume ct is a disjunction, associating to the right. *) fun reduce_disj_True ct = - if is_disj (Thm.term_of ct) then - ((rewr_obj_eq UtilBase.disj_True1_th) + if UtilLogic.is_disj (Thm.term_of ct) then + ((UtilLogic.rewr_obj_eq UtilBase.disj_True1_th) else_conv ((Conv.arg_conv reduce_disj_True) - then_conv (rewr_obj_eq UtilBase.disj_True2_th))) ct + then_conv (UtilLogic.rewr_obj_eq UtilBase.disj_True2_th))) ct else Conv.all_conv ct (* Handles also the case where pat is in not-conj or imp form. *) fun match_prop ctxt (id, item2) pat = let val disj_pats = strip_disj pat (* th is pat'(inst), where pat' is one of the disjunctive terms of pat. *) fun process_inst ((id, inst), th) = let (* Construct the theorem pat'(inst) == True. *) - val to_eqT_cv = (th RS UtilBase.eq_True_th) |> rewr_obj_eq |> Conv.try_conv + val to_eqT_cv = (th RS UtilBase.eq_True_th) |> UtilLogic.rewr_obj_eq |> Conv.try_conv (* Rewrite pat(inst) using the above, then rewrite to True. *) val pat_eqT = pat |> Thm.cterm_of ctxt |> norm_disj |> Util.subst_thm ctxt inst |> apply_to_rhs (ac_disj_conv to_eqT_cv) |> apply_to_rhs reduce_disj_True - |> to_obj_eq + |> UtilLogic.to_obj_eq val patT = pat_eqT RS UtilBase.eq_True_inv_th in ((id, inst), patT) end val insts1 = (ItemIO.match_arg ctxt (PropMatch pat) item2 (id, fo_init)) |> filter (BoxID.has_incr_id o fst o fst) val insts2 = if length disj_pats > 1 then map process_inst (maps (match_prop ctxt (id, item2)) disj_pats) else [] in insts1 @ insts2 end (* Given theorem ~P, cancel any disjunct that is aconv to P. It is possible to leave one disjunct P un-cancelled. *) fun disj_cancel_cv ctxt th ct = - if is_disj (Thm.term_of ct) then + if UtilLogic.is_disj (Thm.term_of ct) then Conv.every_conv [Conv.arg_conv (disj_cancel_cv ctxt th), - Conv.try_conv (rewr_obj_eq (th RS UtilBase.or_cancel1_th)), - Conv.try_conv (rewr_obj_eq (th RS UtilBase.or_cancel2_th))] ct + Conv.try_conv (UtilLogic.rewr_obj_eq (th RS UtilBase.or_cancel1_th)), + Conv.try_conv (UtilLogic.rewr_obj_eq (th RS UtilBase.or_cancel2_th))] ct else Conv.all_conv ct (* Given theorem P and a disjuncion theorem, return new disjunction theorem with ~P cancelled. If all disjuncts can be cancelled, return False. *) fun disj_cancel_prop ctxt th prop = let - val th' = if is_neg (prop_of' th) then th + val th' = if UtilLogic.is_neg (UtilLogic.prop_of' th) then th else th RS UtilBase.nn_create_th - val prop' = prop |> apply_to_thm' (disj_cancel_cv ctxt th') - val P = th' |> prop_of' |> dest_not + val prop' = prop |> UtilLogic.apply_to_thm' (disj_cancel_cv ctxt th') + val P = th' |> UtilLogic.prop_of' |> UtilLogic.dest_not in - if prop_of' prop' aconv P then + if UtilLogic.prop_of' prop' aconv P then [th', prop'] MRS UtilBase.contra_triv_th else prop' end (* Reduce a disjunction p_1 | ... | t | ... | p_n by matching ~t with the second item. If the disjunction contains schematic variables, t must have either zero or the largest number of schematic variables. *) fun match_update_fn ctxt item1 item2 = if is_match_prem_only item1 andalso Util.has_vars (Thm.prop_of (#prop item1)) then [] else let val {id, prop, tname, ...} = item1 val thy = Proof_Context.theory_of ctxt val (disj_head, csubs) = dest_tname_of_disj tname val subs = map Thm.term_of csubs fun count_var t = length (Term.add_vars t []) val max_nvar = fold (curry Int.max) (map count_var subs) 0 fun is_priority_term t = - if is_neg t then - exists (is_ex orf is_bex) (strip_conj (dest_not t)) + if UtilLogic.is_neg t then + exists (UtilLogic.is_ex orf UtilLogic.is_bex) (UtilLogic.strip_conj (UtilLogic.dest_not t)) else - exists (is_obj_all orf is_ball) (strip_disj t) + exists (UtilLogic.is_obj_all orf UtilLogic.is_ball) (strip_disj t) - val has_priority_term = exists is_priority_term (map get_neg subs) + val has_priority_term = exists is_priority_term (map UtilLogic.get_neg subs) val (NO_MATCH, SLOW_MATCH, YES_MATCH) = (0, 1, 2) (* Test whether to perform matching on pattern. *) fun test_do_match t = let val nvar = count_var t - val neg_t = get_neg t + val neg_t = UtilLogic.get_neg t in if not (Util.is_pattern t) then NO_MATCH else if length subs > 1 andalso Property.is_property_prem thy neg_t then NO_MATCH else if has_priority_term andalso not (is_priority_term neg_t) then SLOW_MATCH - else if is_mem neg_t andalso Term.is_Var (dest_arg1 neg_t) andalso + else if UtilLogic.is_mem neg_t andalso Term.is_Var (dest_arg1 neg_t) andalso null (Term.add_frees (dest_arg neg_t) []) then SLOW_MATCH else if nvar = 0 orelse nvar = max_nvar then YES_MATCH else NO_MATCH end (* Match the negation of subs[i] with th2. For each match, instantiate in prop all schematic variables in t, so that t becomes ~th2. Then remove t from prop in the instantiated version. *) fun get_matches i = let val t = nth subs i val do_match = test_do_match t fun process_inst ((id', inst), th) = let val prop' = prop |> Util.subst_thm_thy thy inst |> disj_cancel_prop ctxt th val sc = if do_match = SLOW_MATCH then 200 else 10 in disj_to_update false disj_head (id', SOME sc, prop') :: (if count_var t > 0 then [] else [ShadowItem {id = id', item = item1}]) end in if do_match = NO_MATCH then [] - else t |> get_neg |> match_prop ctxt (id, item2) + else t |> UtilLogic.get_neg |> match_prop ctxt (id, item2) |> maps process_inst end fun get_matches_no_var () = let fun process_inst (id', ths) = let val prop' = prop |> fold (disj_cancel_prop ctxt) ths in disj_to_update false disj_head (id', SOME 1, prop') :: (if is_match_prem_only item1 then [] else [ShadowItem {id = id', item = item1}]) end fun get_match_at_id id' insts = insts |> filter (fn ((id, _), _) => BoxID.is_eq_ancestor ctxt id id') |> map snd |> take 1 fun get_matches_at_id all_insts id' = (id', maps (get_match_at_id id') all_insts) fun merge_matches all_insts = let val ids = distinct (op =) (maps (map (fst o fst)) all_insts) in map (get_matches_at_id all_insts) ids end val _ = assert (length subs >= 2) val ts = [hd subs, List.last subs] in - ts |> map get_neg |> map (match_prop ctxt (id, item2)) + ts |> map UtilLogic.get_neg |> map (match_prop ctxt (id, item2)) |> merge_matches |> maps process_inst end in if max_nvar = 0 then get_matches_no_var () else maps get_matches (0 upto (length subs - 1)) end val match_update_prfstep = {name = "disj_match_update", args = [TypedUniv TY_DISJ, PropMatch (boolVar "A")], func = TwoStep match_update_fn} (* For DISJ items with a single term, of form f p1 ... pn, match t against each of p_i. *) fun match_one_sch_fn ctxt item1 item2 = if is_match_prem_only item1 then [] else let val {id, tname, prop = th1, ...} = item1 val thy = Proof_Context.theory_of ctxt val subs = (dest_tname_of_disj tname) |> snd |> map Thm.term_of in if length subs > 1 then [] else let val t = the_single subs val args = Util.dest_args t fun count_var t = length (Term.add_vars t []) val nvar = count_var t fun get_matches i = if count_var (nth args i) < nvar then [] else let val arg = nth args i val targ = TypedMatch (TY_TERM, arg) val insts = (ItemIO.match_arg ctxt targ item2 (id, fo_init)) |> filter (BoxID.has_incr_id o fst o fst) fun inst_to_updt ((id', inst), _) = let val th1' = Util.subst_thm_thy thy inst th1 - val prop' = prop_of' th1' + val prop' = UtilLogic.prop_of' th1' in - if is_eq_term prop' andalso - RewriteTable.is_equiv_t id' ctxt (dest_eq prop') + if UtilBase.is_eq_term prop' andalso + RewriteTable.is_equiv_t id' ctxt (UtilBase.dest_eq prop') then [] else [Update.thm_update (id', th1')] end in maps inst_to_updt insts end in maps get_matches (0 upto (length args - 1)) end end val match_one_sch_prfstep = {name = "disj_match_one_sch", args = [TypedUniv TY_DISJ, TypedUniv TY_TERM], func = TwoStep match_one_sch_fn} fun disj_match_iff_fn ctxt {id, tname, prop, ...} = if not (BoxID.has_incr_id id) then [] else let val (_, csubs) = dest_tname_of_disj tname val subs = map Thm.term_of csubs in if length subs > 1 then [] - else if not (is_eq_term (the_single subs) andalso - fastype_of (dest_arg (the_single subs)) = boolT) then [] + else if not (UtilBase.is_eq_term (the_single subs) andalso + fastype_of (dest_arg (the_single subs)) = UtilBase.boolT) then [] else let - val cv = (UtilLogic.to_obj_conv ctxt) then_conv (Trueprop_conv norm_disj) - val forward = prop |> UtilLogic.equiv_forward_th |> apply_to_thm cv - val backward = prop |> UtilLogic.equiv_backward_th |> apply_to_thm cv + val cv = (UtilLogic.to_obj_conv ctxt) then_conv (UtilLogic.Trueprop_conv norm_disj) + val forward = prop |> UtilLogic.equiv_forward_th |> Util.apply_to_thm cv + val backward = prop |> UtilLogic.equiv_backward_th |> Util.apply_to_thm cv in - [disj_to_update false imp (id, NONE, forward), - disj_to_update false imp (id, NONE, backward)] + [disj_to_update false UtilLogic.imp (id, NONE, forward), + disj_to_update false UtilLogic.imp (id, NONE, backward)] end end val disj_match_iff_prfstep = {name = "disj_match_iff", args = [TypedUniv TY_DISJ], func = OneStep disj_match_iff_fn} (* For active case, create box checking the next case. *) fun disj_create_case_fn _ {id, tname, ...} = if not (BoxID.has_incr_id id) then [] else if exists Util.has_vars (map Thm.term_of tname) then [] else let val (disj_head, csubs) = dest_tname_of_disj tname in if not (is_active_head disj_head) then [] else if length csubs = 1 then [] else let val subs = map Thm.term_of csubs in [AddBoxes {id = id, - sc = NONE, init_assum = mk_Trueprop (hd subs)}] + sc = NONE, init_assum = UtilLogic.mk_Trueprop (hd subs)}] end end val disj_create_case_prfstep = {name = "disj_create_case", args = [TypedUniv TY_DISJ], func = OneStep disj_create_case_fn} (* item1 dominates item2 if the disjunctive terms in item1 is a subset of that for item2. *) fun disj_shadow_fn ctxt (item1 as {tname = tname1, ...}) (item2 as {tname = tname2, ...}) = let val id = BoxItem.merged_id ctxt [item1, item2] val (disj_head1, subs1) = dest_tname_of_disj tname1 val (disj_head2, subs2) = dest_tname_of_disj tname2 in if not (BoxID.has_incr_id id) then [] else if not (is_active_head disj_head1) andalso is_active_head disj_head2 then [] else if is_match_prem_only item1 andalso not (is_match_prem_only item2) then [] else if subset (op aconvc) (subs1, subs2) then [ShadowItem {id = id, item = item2}] else [] end val disj_shadow_prfstep = {name = "disj_shadow", args = [TypedUniv TY_DISJ, TypedUniv TY_DISJ], func = TwoStep disj_shadow_fn} val add_disj_proofsteps = fold ItemIO.add_item_type [ (TY_DISJ, SOME disj_rewr_terms, SOME output_disj_fn, NONE) ] #> fold ItemIO.add_prop_matcher [ (TY_DISJ, disj_prop_matcher) - ] #> fold add_prfstep [ + ] #> fold ProofStepData.add_prfstep [ match_update_prfstep, match_one_sch_prfstep, disj_match_iff_prfstep, disj_create_case_prfstep, disj_shadow_prfstep ] (* Normalizers *) fun split_not_imp_th th = - th |> apply_to_thm' norm_conj_not_imp + th |> UtilLogic.apply_to_thm' norm_conj_not_imp |> UtilLogic.split_conj_th (* Generalized form of splitting A & B. Also deal with cases ~(A | B) and ~(A --> B). *) fun split_conj_gen_th _ th = - th |> apply_to_thm' norm_conj + th |> UtilLogic.apply_to_thm' norm_conj |> UtilLogic.split_conj_th fun eq_normalizer _ ritem = case ritem of Handler _ => [ritem] | Fact (ty_str, _, th) => if ty_str <> TY_PROP then [ritem] - else if is_eq_term (prop_of' th) then + else if UtilBase.is_eq_term (UtilLogic.prop_of' th) then let - val (lhs, rhs) = dest_eq (prop_of' th) + val (lhs, rhs) = UtilBase.dest_eq (UtilLogic.prop_of' th) in - if fastype_of lhs = boolT then + if fastype_of lhs = UtilBase.boolT then map Update.thm_to_ritem (UtilLogic.split_conj_th (th RS UtilBase.iffD_th)) else [Fact (TY_EQ, [lhs, rhs], th)] end else [ritem] fun property_normalizer _ ritem = case ritem of Handler _ => [ritem] | Fact (ty_str, _, th) => if ty_str <> TY_PROP then [ritem] - else if Property.is_property (prop_of' th) then - [Fact (TY_PROPERTY, [prop_of' th], th)] + else if Property.is_property (UtilLogic.prop_of' th) then + [Fact (TY_PROPERTY, [UtilLogic.prop_of' th], th)] else [ritem] fun disj_normalizer ctxt ritem = case ritem of Handler _ => [ritem] | Fact (ty_str, _, th) => if ty_str <> TY_PROP then [ritem] else let - val t = prop_of' th + val t = UtilLogic.prop_of' th in - if is_neg t andalso is_conj (dest_not t) orelse - is_disj t orelse is_imp t orelse - is_obj_all t orelse is_ball t orelse - is_neg t andalso is_ex (dest_not t) orelse - is_neg t andalso is_bex (dest_not t) + if UtilLogic.is_neg t andalso UtilLogic.is_conj (UtilLogic.dest_not t) orelse + UtilLogic.is_disj t orelse UtilLogic.is_imp t orelse + UtilLogic.is_obj_all t orelse UtilLogic.is_ball t orelse + UtilLogic.is_neg t andalso UtilLogic.is_ex (UtilLogic.dest_not t) orelse + UtilLogic.is_neg t andalso UtilLogic.is_bex (UtilLogic.dest_not t) then let val (disj_head, disj_th) = analyze_disj_th ctxt th val prem_only = Auto2_State.lookup_prem_only ctxt t val disj_th = if prem_only then disj_th else snd (Normalizer.meta_use_vardefs disj_th) in disj_to_ritems prem_only disj_head disj_th end else [ritem] end fun logic_thm_update ctxt (id, th) = let - val t = prop_of' th + val t = UtilLogic.prop_of' th in - if is_obj_all t orelse is_ball t orelse - is_neg t andalso is_ex (dest_not t) orelse - is_neg t andalso is_bex (dest_not t) orelse - is_disj t orelse is_imp t orelse - is_neg t andalso is_conj (dest_not t) + if UtilLogic.is_obj_all t orelse UtilLogic.is_ball t orelse + UtilLogic.is_neg t andalso UtilLogic.is_ex (UtilLogic.dest_not t) orelse + UtilLogic.is_neg t andalso UtilLogic.is_bex (UtilLogic.dest_not t) orelse + UtilLogic.is_disj t orelse UtilLogic.is_imp t orelse + UtilLogic.is_neg t andalso UtilLogic.is_conj (UtilLogic.dest_not t) then let val (disj_head, disj_th) = analyze_disj_th ctxt th val raw_items = disj_to_ritems true disj_head disj_th in AddItems {id = id, sc = NONE, raw_items = raw_items} end else Update.thm_update (id, th) end val add_disj_normalizers = Normalizer.add_th_normalizer ( "split_conj_gen", split_conj_gen_th ) #> fold Normalizer.add_normalizer [ ("eq", eq_normalizer), ("property", property_normalizer), ("disj", disj_normalizer) ] end (* structure Logic_ProofSteps. *) - -val _ = Theory.setup Logic_ProofSteps.add_logic_proofsteps -val _ = Theory.setup Logic_ProofSteps.add_disj_proofsteps -val _ = Theory.setup Logic_ProofSteps.add_disj_normalizers -val TY_DISJ = Logic_ProofSteps.TY_DISJ diff --git a/thys/Auto2_HOL/matcher.ML b/thys/Auto2_HOL/matcher.ML --- a/thys/Auto2_HOL/matcher.ML +++ b/thys/Auto2_HOL/matcher.ML @@ -1,409 +1,409 @@ (* File: matcher.ML Author: Bohua Zhan Matching up to equivalence (E-matching) using a rewrite table. *) signature MATCHER = sig (* Internal *) val check_type_term: theory -> term * term -> id_inst -> (id_inst * term) option val check_type: theory -> typ * typ -> id_inst -> id_inst option val update_inst: term list -> indexname -> cterm -> id_inst -> id_inst_th list (* THe actual matching function. These are defined together. *) val match_list: Proof.context -> term list -> term list * cterm list -> id_inst -> id_inst_ths list val match_comb: Proof.context -> term list -> term * cterm -> id_inst -> id_inst_th list val match_head: Proof.context -> term list -> term * cterm -> id_inst -> id_inst_th list val match_all_head: Proof.context -> term list -> term * cterm -> id_inst -> id_inst_th list val match: Proof.context -> term list -> term * cterm -> id_inst -> id_inst_th list (* Defined in terms of match. *) val rewrite_match: Proof.context -> term * cterm -> id_inst -> id_inst_th list val rewrite_match_head: Proof.context -> term * cterm -> id_inst -> id_inst_th list val rewrite_match_list: Proof.context -> (bool * (term * cterm)) list -> id_inst -> id_inst_ths list val rewrite_match_subset: Proof.context -> term list -> term list * cterm list -> id_inst -> id_inst_ths list (* Prematching. *) val pre_match_type: Proof.context -> typ * typ -> bool val pre_match_comb: Proof.context -> term * cterm -> bool val pre_match_head': Proof.context -> term * cterm -> bool val pre_match_all_head: Proof.context -> term * cterm -> bool val pre_match: Proof.context -> term * cterm -> bool val pre_match_head: Proof.context -> term * cterm -> bool end; -structure Matcher : MATCHER = +functor Matcher(RewriteTable: REWRITE_TABLE) : MATCHER = struct fun compare_inst (((id1, inst1), _), ((id2, inst2), _)) = eq_list (op =) (id1, id2) andalso Util.eq_env (inst1, inst2) (* Match type at the top level for t and u. If there is no match, return NONE. Otherwise, return the updated instsp as well as t instantiated with the new type. *) fun check_type_term thy (t, u) (id, (tyinst, inst)) = let val (T, U) = (fastype_of t, fastype_of u) in if T = U then SOME ((id, (tyinst, inst)), t) else let val tyinst' = Sign.typ_match thy (T, U) tyinst val t' = Envir.subst_term_types tyinst' t in SOME ((id, (tyinst', inst)), t') end end handle Type.TYPE_MATCH => NONE (* Match two types. *) fun check_type thy (T, U) (id, (tyinst, inst)) = let val tyinst' = if T = U then tyinst else Sign.typ_match thy (T, U) tyinst in SOME (id, (tyinst', inst)) end handle Type.TYPE_MATCH => NONE (* Starting here, bd_vars is the list of free variables substituted for bound variables, when matching goes inside abstractions. *) fun is_open bd_vars u = case bd_vars of [] => false | _ => length (inter (op aconv) bd_vars (map Free (Term.add_frees u []))) > 0 (* Assign schematic variable with indexname ixn to u. Type of the schematic variable is determined by the type of u. *) fun update_inst bd_vars ixn cu (id, inst) = let val u = Thm.term_of cu in if is_open bd_vars u then [] else [((id, Util.update_env (ixn, u) inst), Thm.reflexive cu)] end (* Matching an order list of patterns against terms. *) fun match_list ctxt bd_vars (ts, cus) instsp = if null ts andalso null cus then [(instsp, [])] else if null ts orelse null cus then [] else let (* Two choices, one of which should always work (encounter no illegal higher-order patterns. *) fun hd_first () = let val insts_t = match ctxt bd_vars (hd ts, hd cus) instsp fun process_inst_t (instsp', th) = let val insts_ts' = match_list ctxt bd_vars (tl ts, tl cus) instsp' in map (apsnd (cons th)) insts_ts' end in maps process_inst_t insts_t end fun tl_first () = let val insts_ts' = match_list ctxt bd_vars (tl ts, tl cus) instsp fun process_inst_ts' (instsp', ths) = let val insts_t = match ctxt bd_vars (hd ts, hd cus) instsp' in map (apsnd (fn th => th :: ths)) insts_t end in maps process_inst_ts' insts_ts' end in hd_first () handle Fail "invalid pattern" => tl_first () end (* Match a non-AC function. *) and match_comb ctxt bd_vars (t, cu) (instsp as (_, (_, inst))) = let val (tf, targs) = Term.strip_comb t val (cuf, cuargs) = Drule.strip_comb cu val uf = Thm.term_of cuf in if Term.aconv_untyped (tf, uf) then let val instsps' = match_list ctxt bd_vars (targs, cuargs) instsp fun process_inst (instsp', ths) = (instsp', Util.comb_equiv (cuf, ths)) in map process_inst instsps' end else if is_Var tf then let val (ixn, _) = Term.dest_Var tf in case Vartab.lookup inst ixn of NONE => if subset (op aconv) (targs, bd_vars) andalso not (has_duplicates (op aconv) targs) then let val cu' = cu |> Thm.term_of |> fold Util.lambda_abstract (rev targs) |> Thm.cterm_of ctxt in map (fn (instsp', _) => (instsp', Thm.reflexive cu)) (update_inst bd_vars ixn cu' instsp) end else raise Fail "invalid pattern" | SOME (_, tf') => let val t' = Term.list_comb (tf', targs) |> Envir.beta_norm in match ctxt bd_vars (t', cu) instsp end end else [] end (* Match t and u at head. *) and match_head ctxt bd_vars (t, cu) (instsp as (_, (_, inst))) = let val u = Thm.term_of cu in if fastype_of t <> fastype_of u then [] else case (t, u) of (Var ((x, i), _), _) => (case Vartab.lookup inst (x, i) of NONE => if x = "NUMC" andalso not (Consts.is_const_ctxt ctxt u) then [] else if x = "FREE" andalso not (Term.is_Free u) then [] else update_inst bd_vars (x, i) cu instsp | SOME (_, u') => match_head ctxt bd_vars (u', cu) instsp) | (Free (a, _), Free (b, _)) => if a = b then [(instsp, Thm.reflexive cu)] else [] | (Const (a, _), Const (b, _)) => if a = b then [(instsp, Thm.reflexive cu)] else [] | (_ $ _, _) => match_comb ctxt bd_vars (t, cu) instsp | _ => [] end (* With fixed t, match with all equivalences of u. *) and match_all_head ctxt bd_vars (t, cu) (id, env) = let val u = Thm.term_of cu val u_equivs = if is_open bd_vars u then [(id, Thm.reflexive cu)] else RewriteTable.get_head_equiv_with_t ctxt (id, cu) t fun process_equiv (id', eq_u) = let val cu' = Thm.rhs_of eq_u val insts_u' = match_head ctxt bd_vars (t, cu') (id', env) fun process_inst ((id', env'), eq_th) = (* eq_th: t(env') == u', eq_u: u == u'. *) ((id', env'), Util.transitive_list [eq_th, meta_sym eq_u]) in map process_inst insts_u' end in maps process_equiv u_equivs end (* Match t and u, possibly by rewriting u at head. *) and match ctxt bd_vars (t, cu) (instsp as (_, (_, inst))) = case check_type_term (Proof_Context.theory_of ctxt) (t, Thm.term_of cu) instsp of NONE => [] | SOME (instsp', t') => case (t', Thm.term_of cu) of (Var ((x, i), _), _) => (case Vartab.lookup inst (x, i) of NONE => if member (op =) ["NUMC", "FREE"] x then match_all_head ctxt bd_vars (t', cu) instsp' else update_inst bd_vars (x, i) cu instsp' | SOME (_, u') => match ctxt bd_vars (u', cu) instsp') | (Abs (_, T, t'), Abs (x, U, _)) => ( case check_type (Proof_Context.theory_of ctxt) (T, U) instsp' of NONE => [] | SOME (instsp'' as (_, (tyinst', _))) => let val (cv, cu') = Thm.dest_abs_global cu val v = Thm.term_of cv val t'' = Envir.subst_term_types tyinst' t' val t_free = Term.subst_bound (v, t'') val insts' = match ctxt (v :: bd_vars) (t_free, cu') instsp'' fun process_inst (inst', th') = (inst', Thm.abstract_rule x cv th') in map process_inst insts' end) | _ => (* Free, Const, and comb case *) (match_all_head ctxt bd_vars (t', cu) instsp') |> distinct compare_inst (* Function for export *) fun rewrite_match_gen at_head ctxt (t, cu) (id, env) = if at_head then case check_type_term (Proof_Context.theory_of ctxt) (t, Thm.term_of cu) (id, env) of NONE => [] | SOME ((id, env), t) => match_head ctxt [] (t, cu) (id, env) else match ctxt [] (t, cu) (id, env) val rewrite_match = rewrite_match_gen false val rewrite_match_head = rewrite_match_gen true (* pairs is a list of (at_head, (t, u)). Match the pairs in sequence, and return a list of ((id, inst), ths), where ths is the list of equalities t(env) == u. *) fun rewrite_match_list ctxt pairs instsp = case pairs of [] => [(instsp, [])] | (at_head, (t, cu)) :: pairs' => let val insts_t = rewrite_match_gen at_head ctxt (t, cu) instsp fun process_inst_t (instsp', th) = let val insts_ts' = rewrite_match_list ctxt pairs' instsp' in map (apsnd (cons th)) insts_ts' end in maps process_inst_t insts_t end (* Given two lists of terms (ts, us), match ts with a subset of us. Return a list of ((id, inst), ths), where ths is the list of equalities t_i(env) == u_j. *) fun rewrite_match_subset ctxt bd_vars (ts, cus) instsp = case ts of [] => [(instsp, [])] | t :: ts' => let fun match_i i = map (pair i) (match ctxt bd_vars (t, nth cus i) instsp) fun process_match_i (i, (instsp', th)) = let val insts_ts' = rewrite_match_subset ctxt bd_vars (ts', nth_drop i cus) instsp' in map (apsnd (cons th)) insts_ts' end in (0 upto (length cus - 1)) |> maps match_i |> maps process_match_i end handle Fail "invalid pattern" => if length ts' > 0 andalso Term_Ord.term_ord (t, hd ts') = LESS then rewrite_match_subset ctxt bd_vars (ts' @ [t], cus) instsp else raise Fail "rewrite_match_subset: invalid pattern" (* Fast function for determining whether there can be a match between t and u. *) fun pre_match_type ctxt (T, U) = let val thy = Proof_Context.theory_of ctxt val _ = Sign.typ_match thy (T, U) Vartab.empty in true end handle Type.TYPE_MATCH => false fun pre_match_comb ctxt (t, cu) = let val (tf, targs) = Term.strip_comb t val (cuf, cuargs) = Drule.strip_comb cu val uf = Thm.term_of cuf in is_Var tf orelse (Term.aconv_untyped (tf, uf) andalso length targs = length cuargs andalso forall (pre_match ctxt) (targs ~~ cuargs)) end and pre_match_head' ctxt (t, cu) = let val u = Thm.term_of cu in case (t, u) of (Var ((x, _), T), _) => if Term.is_open u then false else if not (pre_match_type ctxt (T, fastype_of u)) then false else if x = "NUMC" then Consts.is_const_ctxt ctxt u else if x = "FREE" then Term.is_Free u else true | (Free (a, T), Free (b, U)) => (a = b andalso pre_match_type ctxt (T, U)) | (Const (a, T), Const (b, U)) => (a = b andalso pre_match_type ctxt (T, U)) | (_ $ _, _) => pre_match_comb ctxt (t, cu) | _ => false end and pre_match_all_head ctxt (t, cu) = let val thy = Proof_Context.theory_of ctxt val u = Thm.term_of cu val tyinst = Sign.typ_match thy (fastype_of t, fastype_of u) Vartab.empty val t' = Envir.subst_term_types tyinst t in if Term.is_open u then pre_match_head' ctxt (t', cu) else let val u_equivs = (RewriteTable.get_head_equiv_with_t ctxt ([], cu) t') |> map snd |> map Thm.rhs_of in exists (fn cu' => pre_match_head' ctxt (t', cu')) u_equivs end end handle Type.TYPE_MATCH => false and pre_match ctxt (t, cu) = let val u = Thm.term_of cu in case (t, u) of (Var ((x, _), T), _) => if Term.is_open u then false else if not (pre_match_type ctxt (T, fastype_of u)) then false else if member (op =) ["NUMC", "FREE"] x then pre_match_all_head ctxt (t, cu) else true | (Abs (_, T, t'), Abs (_, U, _)) => if not (pre_match_type ctxt (T, U)) then false else let val (cv, cu') = Thm.dest_abs_global cu val t'' = subst_bound (Thm.term_of cv, t') in pre_match ctxt (t'', cu') end | (Bound i, Bound j) => (i = j) | (_, _) => pre_match_all_head ctxt (t, cu) end fun pre_match_head ctxt (t, cu) = let val thy = Proof_Context.theory_of ctxt val u = Thm.term_of cu val tyinst = Sign.typ_match thy (fastype_of t, fastype_of u) Vartab.empty val t' = Envir.subst_term_types tyinst t in pre_match_head' ctxt (t', cu) end handle Type.TYPE_MATCH => false end (* structure Matcher. *) diff --git a/thys/Auto2_HOL/normalize.ML b/thys/Auto2_HOL/normalize.ML --- a/thys/Auto2_HOL/normalize.ML +++ b/thys/Auto2_HOL/normalize.ML @@ -1,293 +1,309 @@ (* File: normalize.ML Author: Bohua Zhan Normalization procedure for facts obtained during a proof. *) +signature NORMALIZER_TYPE = +sig + type normalizer = Proof.context -> raw_item -> raw_item list +end + +structure Normalizer_Type = +struct +type normalizer = Proof.context -> raw_item -> raw_item list; +end + signature NORMALIZER = sig - type normalizer = Proof.context -> raw_item -> raw_item list + include NORMALIZER_TYPE val add_normalizer: string * normalizer -> theory -> theory val add_th_normalizer: string * (Proof.context -> thm -> thm list) -> theory -> theory val add_rewr_normalizer: string * thm -> theory -> theory val get_normalizers: theory -> (string * normalizer) list val normalize: Proof.context -> raw_item -> raw_item list val normalize_keep: Proof.context -> raw_item -> raw_item list (* Normalization of definition of variable *) val add_inj_struct_data: thm -> theory -> theory val is_def_eq: theory -> term -> bool val swap_eq_to_front: conv val meta_use_vardef: thm -> (term * term) list * thm val meta_use_vardefs: thm -> (term * term) list * thm val def_subst: (term * term) list -> term -> term end; -structure Normalizer : NORMALIZER = +functor Normalizer( + structure BoxItem: BOXITEM; + structure UtilBase: UTIL_BASE; + structure UtilLogic: UTIL_LOGIC; + structure Update: UPDATE; + ): NORMALIZER = struct (* Keeps list of normalizers. *) -type normalizer = Proof.context -> raw_item -> raw_item list; +open Normalizer_Type structure Data = Theory_Data ( type T = (normalizer * serial) Symtab.table; val empty = Symtab.empty; val merge = Symtab.merge (eq_snd op =) ) fun add_normalizer (norm_name, norm_fun) = Data.map (Symtab.update_new (norm_name, (norm_fun, serial ()))) (* Apply norm_fun: thm -> thm list to any PROP item. *) fun th_normalizer norm_fun ctxt ritem = case ritem of Handler _ => [ritem] | Fact (ty_str, _, th) => if ty_str = TY_PROP then map Update.thm_to_ritem (norm_fun ctxt th) else [ritem] fun add_th_normalizer (norm_name, norm_fun) = add_normalizer (norm_name, th_normalizer norm_fun) (* eq_th is a meta equality. *) fun rewr_normalizer eq_th ctxt ritem = let val cv = (Conv.top_conv (K (Conv.try_conv (Conv.rewr_conv eq_th))) ctxt) then_conv (Thm.beta_conversion true) in case ritem of Handler _ => [ritem] | Fact (ty_str, _, th) => if ty_str = TY_PROP then [Update.thm_to_ritem (apply_to_thm cv th)] else if ty_str = TY_EQ then if Util.is_meta_eq (Thm.prop_of th) then (* Equality between lambda terms *) [ritem] else let (* Apply to right side *) - val th' = apply_to_thm' (Conv.arg_conv cv) th - val (lhs, rhs) = dest_eq (prop_of' th') + val th' = UtilLogic.apply_to_thm' (Conv.arg_conv cv) th + val (lhs, rhs) = UtilBase.dest_eq (UtilLogic.prop_of' th') in [Fact (TY_EQ, [lhs, rhs], th')] end else [ritem] end fun add_rewr_normalizer (norm_name, eq_th) = add_normalizer (norm_name, rewr_normalizer eq_th) fun get_normalizers thy = Symtab.dest (Data.get thy) |> map (apsnd #1) fun normalize ctxt ritem = let val norms = get_normalizers (Proof_Context.theory_of ctxt) fun apply_norm (_, norm_fun) ritems = maps (norm_fun ctxt) ritems val norm_once = fold apply_norm norms [ritem] in case norm_once of [ritem'] => if BoxItem.eq_ritem (ritem, ritem') then [ritem'] else normalize ctxt ritem' | _ => maps (normalize ctxt) norm_once end (* Perform normalization, but keep the original ritem. *) fun normalize_keep ctxt ritem = let val norm_ritems = normalize ctxt ritem in if length norm_ritems = 1 then norm_ritems else ritem :: norm_ritems end (* Normalization of variable definition *) structure InjStructData = Theory_Data ( type T = thm Symtab.table; val empty = Symtab.empty; val merge = Symtab.merge Thm.eq_thm ) fun add_inj_struct_data th thy = let - val (lhs, _) = th |> prop_of' |> dest_eq |> fst |> dest_eq + val (lhs, _) = th |> UtilLogic.prop_of' |> UtilBase.dest_eq |> fst |> UtilBase.dest_eq val (f, _) = Term.strip_comb lhs in case f of Const (nm, _) => InjStructData.map (Symtab.update_new (nm, th)) thy | _ => raise Fail "add_inj_struct_data" end (* Check whether t is of the form ?A = t, where t does not contain ?A. *) fun inj_args thy t = if Term.is_Var t then [t] else let val (f, args) = Term.strip_comb t in if Term.is_Const f andalso Symtab.defined (InjStructData.get thy) (Util.get_head_name f) then maps (inj_args thy) args else [t] end fun is_def_eq thy t = - if not (is_eq_term t) then false + if not (UtilBase.is_eq_term t) then false else let - val (lhs, rhs) = dest_eq t + val (lhs, rhs) = UtilBase.dest_eq t val (l_args, r_args) = apply2 (inj_args thy) (lhs, rhs) in forall Term.is_Var l_args andalso not (forall Term.is_Var r_args) andalso forall (fn t => not (Util.is_subterm t rhs)) l_args end fun is_def_eq' thy t = - is_Trueprop t andalso is_def_eq thy (dest_Trueprop t) + UtilLogic.is_Trueprop t andalso is_def_eq thy (UtilLogic.dest_Trueprop t) fun is_neg_def_eq thy t = - is_neg t andalso is_def_eq thy (dest_not t) + UtilLogic.is_neg t andalso is_def_eq thy (UtilLogic.dest_not t) (* Given t of the form A_1 ==> ... ==> A_n ==> C, swap all A_i of the form ?A = t to the front. *) fun swap_eq_to_front ct = let val t = Thm.term_of ct val thy = Thm.theory_of_cterm ct in if Util.is_implies t then if is_def_eq' thy (dest_arg1 t) then Conv.all_conv ct else Conv.every_conv [Conv.arg_conv swap_eq_to_front, Conv.rewr_conv Drule.swap_prems_eq] ct else Conv.no_conv ct end (* Given th where the first premise is in the form ?A = t, or f ?A_1 ... ?A_n = t, where f is injective, replace ?A or each ?A_i in the rest of th, and remove the first premise. *) fun meta_use_vardef th = if not (Util.is_implies (Thm.prop_of th)) then ([], th) else let val cprem = th |> Thm.cprop_of |> Thm.dest_arg1 |> Thm.dest_arg val prem = Thm.term_of cprem val thy = Thm.theory_of_thm th in - if is_conj prem then - th |> apply_to_thm (Conv.rewr_conv (meta_sym @{thm atomize_conjL})) + if UtilLogic.is_conj prem then + th |> apply_to_thm (Conv.rewr_conv (meta_sym UtilBase.atomize_conjL_th)) |> meta_use_vardef else if is_def_eq thy prem then let - val (c_lhs, c_rhs) = cdest_eq cprem - val (lhs, rhs) = dest_eq prem + val (c_lhs, c_rhs) = UtilBase.cdest_eq cprem + val (lhs, rhs) = UtilBase.dest_eq prem in if Term.is_Var lhs then let val (pairs, th') = th |> Util.subst_thm_atomic [(c_lhs, c_rhs)] |> apply_to_thm (Conv.arg1_conv UtilBase.to_meta_eq_cv) |> Thm.elim_implies (Thm.reflexive c_rhs) |> meta_use_vardef in ((lhs, rhs) :: pairs, th') end else let val nm = Util.get_head_name lhs val data = InjStructData.get thy val inj_th = the (Symtab.lookup data nm) handle Option.Option => raise Fail "meta_use_vardef" in th |> apply_to_thm ( - Conv.arg1_conv (Conv.arg_conv (rewr_obj_eq inj_th))) + Conv.arg1_conv (Conv.arg_conv (UtilLogic.rewr_obj_eq inj_th))) |> meta_use_vardef end end else ([], th) end fun disj_ts t = - if is_disj t then dest_arg1 t :: disj_ts (dest_arg t) + if UtilLogic.is_disj t then dest_arg1 t :: disj_ts (dest_arg t) else [t] fun swap_disj ct = - Conv.every_conv [rewr_obj_eq (obj_sym UtilBase.disj_assoc_th), - Conv.arg1_conv (rewr_obj_eq UtilBase.disj_commute_th), - rewr_obj_eq UtilBase.disj_assoc_th] ct + Conv.every_conv [UtilLogic.rewr_obj_eq (UtilLogic.obj_sym UtilBase.disj_assoc_th), + Conv.arg1_conv (UtilLogic.rewr_obj_eq UtilBase.disj_commute_th), + UtilLogic.rewr_obj_eq UtilBase.disj_assoc_th] ct fun disj_swap_eq_to_front' ct = let val t = Thm.term_of ct val thy = Thm.theory_of_cterm ct in - if is_disj t then + if UtilLogic.is_disj t then if is_neg_def_eq thy (dest_arg1 t) then Conv.all_conv ct - else if is_disj (dest_arg t) then + else if UtilLogic.is_disj (dest_arg t) then Conv.every_conv [Conv.arg_conv disj_swap_eq_to_front', swap_disj] ct else if is_neg_def_eq thy (dest_arg t) then - rewr_obj_eq UtilBase.disj_commute_th ct + UtilLogic.rewr_obj_eq UtilBase.disj_commute_th ct else Conv.no_conv ct else Conv.no_conv ct end fun disj_swap_eq_to_front ct = Conv.every_conv [ - Trueprop_conv disj_swap_eq_to_front', - Trueprop_conv (rewr_obj_eq (obj_sym UtilBase.imp_conv_disj_th)), + UtilLogic.Trueprop_conv disj_swap_eq_to_front', + UtilLogic.Trueprop_conv + (UtilLogic.rewr_obj_eq (UtilLogic.obj_sym UtilBase.imp_conv_disj_th)), Conv.rewr_conv (meta_sym UtilBase.atomize_imp_th) ] ct fun meta_use_vardefs th = let val thy = Thm.theory_of_thm th in if exists (is_def_eq' thy) (Thm.prems_of th) then let val (pairs, th') = th |> apply_to_thm swap_eq_to_front |> meta_use_vardef val (pairs', th'') = meta_use_vardefs th' in (pairs @ pairs', th'') end - else if is_Trueprop (Thm.prop_of th) then + else if UtilLogic.is_Trueprop (Thm.prop_of th) then let - val ts = disj_ts (prop_of' th) + val ts = disj_ts (UtilLogic.prop_of' th) in if length ts > 1 andalso exists (is_neg_def_eq thy) ts then let val (pairs, th') = th |> apply_to_thm disj_swap_eq_to_front |> meta_use_vardef val (pairs', th'') = meta_use_vardefs th' in (pairs @ pairs', th'') end else ([], th) end else ([], th) end fun def_subst pairs t = fold (fn p => Term.subst_atomic [p]) pairs t end (* structure Normalizer. *) diff --git a/thys/Auto2_HOL/proofsteps.ML b/thys/Auto2_HOL/proofsteps.ML --- a/thys/Auto2_HOL/proofsteps.ML +++ b/thys/Auto2_HOL/proofsteps.ML @@ -1,967 +1,976 @@ (* File: proofsteps.ML Author: Bohua Zhan Definition of type proofstep, and facility for adding basic proof steps. *) datatype proofstep_fn = OneStep of Proof.context -> box_item -> raw_update list | TwoStep of Proof.context -> box_item -> box_item -> raw_update list type proofstep = { name: string, args: match_arg list, func: proofstep_fn } datatype prfstep_descriptor = WithFact of term | WithItem of string * term | WithProperty of term | WithWellForm of term * term | WithScore of int | GetFact of term * thm | ShadowFirst | ShadowSecond | CreateCase of term | CreateConcl of term | Filter of prfstep_filter +(* Conditional prfstep_descriptors. *) +type pre_prfstep_descriptor = Proof.context -> prfstep_descriptor + signature PROOFSTEP = sig val eq_prfstep: proofstep * proofstep -> bool val apply_prfstep: Proof.context -> box_item list -> proofstep -> raw_update list val WithGoal: term -> prfstep_descriptor val WithTerm: term -> prfstep_descriptor val WithProp: term -> prfstep_descriptor val string_of_desc: theory -> prfstep_descriptor -> string val string_of_descs: theory -> prfstep_descriptor list -> string (* prfstep_filter *) val all_insts: prfstep_filter val neq_filter: term -> prfstep_filter val order_filter: string -> string -> prfstep_filter val size1_filter: string -> prfstep_filter val not_type_filter: string -> typ -> prfstep_filter (* First level proofstep writing functions. *) val apply_pat_r: Proof.context -> id_inst_ths -> term * thm -> thm val retrieve_args: prfstep_descriptor list -> match_arg list val retrieve_pats_r: prfstep_descriptor list -> (term * thm) list val retrieve_filts: prfstep_descriptor list -> prfstep_filter val retrieve_cases: prfstep_descriptor list -> term list val retrieve_shadows: prfstep_descriptor list -> int list val get_side_ths: Proof.context -> id_inst -> match_arg list -> (box_id * thm list) list val prfstep_custom: string -> prfstep_descriptor list -> (id_inst_ths -> box_item list -> Proof.context -> raw_update list) -> proofstep val gen_prfstep: string -> prfstep_descriptor list -> proofstep val prfstep_pre_conv: string -> prfstep_descriptor list -> (Proof.context -> conv) -> proofstep val prfstep_conv: string -> prfstep_descriptor list -> conv -> proofstep end; -structure ProofStep : PROOFSTEP = +functor ProofStep( + structure ItemIO: ITEM_IO; + structure Matcher: MATCHER; + structure PropertyData: PROPERTY_DATA; + structure RewriteTable: REWRITE_TABLE; + structure UtilBase: UTIL_BASE; + structure UtilLogic: UTIL_LOGIC; + structure Update: UPDATE; + structure WellformData: WELLFORM_DATA; + ): PROOFSTEP = struct fun eq_prfstep (prfstep1, prfstep2) = (#name prfstep1 = #name prfstep2) fun apply_prfstep ctxt items {func, ...} = case func of OneStep f => f ctxt (the_single items) | TwoStep f => f ctxt (hd items) (nth items 1) fun WithGoal t = let - val _ = assert (type_of t = boolT) "WithGoal: pat should have type bool." + val _ = assert (type_of t = UtilBase.boolT) "WithGoal: pat should have type bool." in - WithFact (get_neg t) + WithFact (UtilLogic.get_neg t) end fun WithTerm t = WithItem (TY_TERM, t) fun WithProp t = let - val _ = assert (type_of t = boolT) "WithProp: pat should have type bool." + val _ = assert (type_of t = UtilBase.boolT) "WithProp: pat should have type bool." in WithItem (TY_PROP, t) end fun string_of_desc thy desc = let val print = Syntax.string_of_term_global thy in case desc of WithFact t => - if is_neg t then "WithGoal " ^ (print (get_neg t)) + if UtilLogic.is_neg t then "WithGoal " ^ (print (UtilLogic.get_neg t)) else "WithFact " ^ (print t) | WithItem (ty_str, t) => if ty_str = TY_TERM then "WithTerm " ^ (print t) else "WithItem " ^ ty_str ^ " " ^ (print t) | WithProperty t => "WithProperty " ^ (print t) | WithWellForm (_, req) => "WithWellForm " ^ (print req) | WithScore n => "WithScore " ^ (string_of_int n) | GetFact (t, th) => if t aconv @{term False} then "GetResolve " ^ (Util.name_of_thm th) - else if is_neg t then - "GetGoal (" ^ (print (get_neg t)) ^ ", " ^ (Util.name_of_thm th) ^ ")" + else if UtilLogic.is_neg t then + "GetGoal (" ^ (print (UtilLogic.get_neg t)) ^ ", " ^ (Util.name_of_thm th) ^ ")" else "GetFact (" ^ (print t) ^ ", " ^ (Util.name_of_thm th) ^ ")" | ShadowFirst => "Shadow first" | ShadowSecond => "Shadow second" | CreateCase assum => "CreateCase " ^ (print assum) | CreateConcl concl => "CreateConcl " ^ (print concl) | Filter _ => "Filter (...)" end fun string_of_descs thy descs = let fun is_filter desc = case desc of Filter _ => true | _ => false val (filts, non_filts) = filter_split is_filter descs in (cat_lines (map (string_of_desc thy) non_filts)) ^ (if length filts > 0 then (" + " ^ (string_of_int (length filts)) ^ " filters") else "") end (* prfstep_filter *) val all_insts = fn _ => fn _ => true fun neq_filter cond ctxt (id, inst) = let val (lhs, rhs) = - cond |> dest_not |> dest_eq + cond |> UtilLogic.dest_not |> UtilBase.dest_eq handle Fail "dest_not" => raise Fail "neq_filter: not an inequality." | Fail "dest_eq" => raise Fail "neq_filter: not an inequality." val _ = assert (null (Term.add_frees cond [])) "neq_filter: should not contain free variable." val t1 = Util.subst_term_norm inst lhs val t2 = Util.subst_term_norm inst rhs in if Util.has_vars t1 andalso Util.has_vars t2 then true else if Util.has_vars t1 then (Matcher.rewrite_match ctxt (t1, Thm.cterm_of ctxt t2) (id, fo_init)) |> filter (fn ((id', _), _) => id = id') |> null else if Util.has_vars t2 then (Matcher.rewrite_match ctxt (t2, Thm.cterm_of ctxt t1) (id, fo_init)) |> filter (fn ((id', _), _) => id = id') |> null else not (RewriteTable.is_equiv_t id ctxt (t1, t2)) end fun order_filter s1 s2 _ (_, inst) = not (Term_Ord.term_ord (lookup_inst inst s2, lookup_inst inst s1) = LESS) fun size1_filter s1 ctxt (id, inst) = size_of_term (RewriteTable.simp_val_t id ctxt (lookup_inst inst s1)) = 1 fun not_type_filter s ty _ (_, inst) = not (Term.fastype_of (lookup_inst inst s) = ty) (* First level proofstep writing functions. *) fun apply_pat_r ctxt ((_, inst), ths) (pat_r, th) = let - val _ = assert (fastype_of pat_r = boolT) + val _ = assert (fastype_of pat_r = UtilBase.boolT) "apply_pat_r: pat_r should be of type bool" (* Split into meta equalities (usually produced by term matching, not applied to th, and others (assumptions for th). *) val (eqs, ths') = ths |> filter_split (Util.is_meta_eq o Thm.prop_of) val _ = assert (length ths' = Thm.nprems_of th) "apply_pat_r: wrong number of assumptions." - val inst_new = Util.subst_term_norm inst (mk_Trueprop pat_r) + val inst_new = Util.subst_term_norm inst (UtilLogic.mk_Trueprop pat_r) val th' = th |> Util.subst_thm ctxt inst |> fold Thm.elim_implies ths' val _ = if inst_new aconv (Thm.prop_of th') then () else raise Fail "apply_pat_r: conclusion mismatch" (* Rewrite on subterms, top sweep order. *) fun rewr_top eq_th = Conv.top_sweep_rewrs_conv [eq_th] ctxt in th' |> apply_to_thm (Conv.every_conv (map rewr_top eqs)) end fun retrieve_args descs = maps (fn desc => case desc of WithFact t => [PropMatch t] | WithItem (ty_str, t) => [TypedMatch (ty_str, t)] | WithProperty t => [PropertyMatch t] | WithWellForm t => [WellFormMatch t] | _ => []) descs fun retrieve_pats_r descs = maps (fn desc => case desc of GetFact (pat_r, th) => [(pat_r, th)] | _ => []) descs fun retrieve_filts descs = let val filts = maps (fn Filter filt => [filt] | _ => []) descs in fn ctxt => fn inst => forall (fn f => f ctxt inst) filts end fun retrieve_cases descs = let fun retrieve_case desc = case desc of - CreateCase assum => [mk_Trueprop assum] - | CreateConcl concl => [mk_Trueprop (get_neg concl)] + CreateCase assum => [UtilLogic.mk_Trueprop assum] + | CreateConcl concl => [UtilLogic.mk_Trueprop (UtilLogic.get_neg concl)] | _ => [] in maps retrieve_case descs end fun retrieve_shadows descs = let fun retrieve_shadow desc = case desc of ShadowFirst => [0] | ShadowSecond => [1] | _ => [] in maps retrieve_shadow descs end fun retrieve_score descs = let fun retrieve_score desc = case desc of WithScore n => SOME n | _ => NONE in get_first retrieve_score descs end (* Given list of PropertyMatch and WellFormMatch arguments, attempt to find the corresponding theorems in the rewrite table. Return the list of theorems for each possible (mutually non-comparable) box IDs. *) fun get_side_ths ctxt (id, inst) side_args = if null side_args then [(id, [])] else let val side_args' = map (ItemIO.subst_arg inst) side_args fun process_side_arg side_arg = case side_arg of PropertyMatch prop => PropertyData.get_property_t ctxt (id, prop) | WellFormMatch (t, req) => (WellformData.get_wellform_t ctxt (id, t)) - |> filter (fn (_, th) => prop_of' th aconv req) + |> filter (fn (_, th) => UtilLogic.prop_of' th aconv req) | _ => raise Fail "get_side_ths: wrong kind of arg." val side_ths = map process_side_arg side_args' in if exists null side_ths then [] else side_ths |> BoxID.get_all_merges_info ctxt |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) end (* Creates a proofstep with specified patterns and filters (in descs), and a custom function converting any instantiations into updates. *) fun prfstep_custom name descs updt_fn = let val args = retrieve_args descs val (item_args, side_args) = filter_split ItemIO.is_ordinary_match args val filt = retrieve_filts descs val shadows = retrieve_shadows descs (* Processing an instantiation after matching the one or two main matchers: apply filters, remove trivial True from matchings, find properties, and replace incremental ids. *) fun process_inst ctxt ((id, inst), ths) = (get_side_ths ctxt (id, inst) side_args) |> filter (BoxID.has_incr_id o fst) |> map (fn (id', p_ths) => ((id', inst), p_ths @ ths)) |> filter (filt ctxt o fst) fun shadow_to_update items ((id, _), _) n = ShadowItem {id = id, item = nth items n} in if length item_args = 1 then let val arg = the_single item_args fun prfstep ctxt item = let val inst_ths = (ItemIO.match_arg ctxt arg item ([], fo_init)) |> map (fn (inst, th) => (inst, [th])) |> maps (process_inst ctxt) fun process_inst inst_th = updt_fn inst_th [item] ctxt @ map (shadow_to_update [item] inst_th) shadows in maps process_inst inst_ths end in {name = name, args = args, func = OneStep prfstep} end else if length item_args = 2 then let val (arg1, arg2) = the_pair item_args fun prfstep1 ctxt item1 = let val inst_ths = ItemIO.match_arg ctxt arg1 item1 ([], fo_init) fun process_inst1 item2 ((id, inst), th) = let val arg2' = ItemIO.subst_arg inst arg2 val inst_ths' = (ItemIO.match_arg ctxt arg2' item2 (id, inst)) |> map (fn (inst', th') => (inst', [th, th'])) |> maps (process_inst ctxt) fun process_inst inst_th' = updt_fn inst_th' [item1, item2] ctxt @ map (shadow_to_update [item1, item2] inst_th') shadows in maps process_inst inst_ths' end in fn item2 => maps (process_inst1 item2) inst_ths end in {name = name, args = args, func = TwoStep prfstep1} end else raise Fail "prfstep_custom: must have 1 or 2 patterns." end (* Create a proofstep from a list of proofstep descriptors. See datatype prfstep_descriptor for allowed types of descriptors. *) fun gen_prfstep name descs = let val args = retrieve_args descs val pats_r = retrieve_pats_r descs val cases = retrieve_cases descs val sc = retrieve_score descs val input_descs = filter (fn desc => case desc of GetFact _ => false | CreateCase _ => false | CreateConcl _ => false | _ => true) descs (* Verify that all schematic variables appearing in pats_r / cases appear in pats. *) val pats = map ItemIO.pat_of_match_arg args val vars = map Var (fold Term.add_vars pats []) fun check_pat_r (pat_r, _) = subset (op aconv) (map Var (Term.add_vars pat_r []), vars) fun check_case assum = subset (op aconv) (map Var (Term.add_vars assum []), vars) val _ = assert (forall check_pat_r pats_r andalso forall check_case cases) "gen_prfstep: new schematic variable in pats_r / cases." fun pats_r_to_update ctxt (inst_ths as ((id, _), _)) = if null pats_r then [] else let val ths = map (apply_pat_r ctxt inst_ths) pats_r in if length ths = 1 andalso - Thm.prop_of (the_single ths) aconv pFalse then + Thm.prop_of (the_single ths) aconv UtilLogic.pFalse then [ResolveBox {id = id, th = the_single ths}] else [AddItems {id = id, sc = sc, raw_items = map Update.thm_to_ritem ths}] end fun case_to_update ((id, inst), _) assum = AddBoxes {id = id, sc = sc, init_assum = Util.subst_term_norm inst assum} fun cases_to_update inst_ths = map (case_to_update inst_ths) cases fun updt_fn inst_th _ ctxt = pats_r_to_update ctxt inst_th @ cases_to_update inst_th in prfstep_custom name input_descs updt_fn end fun prfstep_pre_conv name descs pre_cv = let val args = retrieve_args descs val _ = case args of [TypedMatch ("TERM", _)] => () | _ => raise Fail ("prfstep_conv: should have exactly one " ^ "term pattern.") val filt = retrieve_filts descs fun prfstep ctxt item = let val inst_ths = (ItemIO.match_arg ctxt (the_single args) item ([], fo_init)) |> filter (BoxID.has_incr_id o fst o fst) |> filter (filt ctxt o fst) fun inst_to_updt ((id, _), eq1) = (* Here eq1 is meta_eq from pat(inst) to item. *) let val ct = Thm.lhs_of eq1 val err = name ^ ": cv failed." val eq_th = pre_cv ctxt ct handle CTERM _ => raise Fail err in if Thm.is_reflexive eq_th then [] else if RewriteTable.is_equiv id ctxt (Thm.rhs_of eq1, Thm.rhs_of eq_th) then [] else let - val th = to_obj_eq (Util.transitive_list [meta_sym eq1, eq_th]) + val th = UtilLogic.to_obj_eq (Util.transitive_list [meta_sym eq1, eq_th]) in [Update.thm_update (id, th)] end end in maps inst_to_updt inst_ths end in {name = name, args = args, func = OneStep prfstep} end fun prfstep_conv name descs cv = prfstep_pre_conv name descs (K cv) end (* structure ProofStep *) -val WithTerm = ProofStep.WithTerm -val WithGoal = ProofStep.WithGoal -val WithProp = ProofStep.WithProp -val neq_filter = ProofStep.neq_filter -val order_filter = ProofStep.order_filter -val size1_filter = ProofStep.size1_filter -val not_type_filter = ProofStep.not_type_filter - signature PROOFSTEP_DATA = sig val add_prfstep: proofstep -> theory -> theory val del_prfstep_pred: (string -> bool) -> theory -> theory val del_prfstep: string -> theory -> theory val del_prfstep_thm: thm -> theory -> theory val del_prfstep_thm_str: string -> thm -> theory -> theory val del_prfstep_thm_eqforward: thm -> theory -> theory val get_prfsteps: theory -> proofstep list val add_prfstep_custom: (string * prfstep_descriptor list * (id_inst_ths -> box_item list -> Proof.context -> raw_update list)) -> theory -> theory val add_gen_prfstep: string * prfstep_descriptor list -> theory -> theory val add_prfstep_pre_conv: string * prfstep_descriptor list * (Proof.context -> conv) -> theory -> theory val add_prfstep_conv: string * prfstep_descriptor list * conv -> theory -> theory (* Constructing conditional prfstep_descriptors. *) - type pre_prfstep_descriptor = Proof.context -> prfstep_descriptor + type pre_prfstep_descriptor = pre_prfstep_descriptor val with_term: string -> pre_prfstep_descriptor val with_cond: string -> pre_prfstep_descriptor val with_conds: string list -> pre_prfstep_descriptor list val with_filt: prfstep_filter -> pre_prfstep_descriptor val with_filts: prfstep_filter list -> pre_prfstep_descriptor list val with_score: int -> pre_prfstep_descriptor (* Second level proofstep writing functions. *) datatype prfstep_mode = MODE_FORWARD | MODE_FORWARD' | MODE_BACKWARD | MODE_BACKWARD1 | MODE_BACKWARD2 | MODE_RESOLVE val add_prfstep_check_req: string * string -> theory -> theory val add_forward_prfstep_cond: thm -> pre_prfstep_descriptor list -> theory -> theory val add_forward'_prfstep_cond: thm -> pre_prfstep_descriptor list -> theory -> theory val add_backward_prfstep_cond: thm -> pre_prfstep_descriptor list -> theory -> theory val add_backward1_prfstep_cond: thm -> pre_prfstep_descriptor list -> theory -> theory val add_backward2_prfstep_cond: thm -> pre_prfstep_descriptor list -> theory -> theory val add_resolve_prfstep_cond: thm -> pre_prfstep_descriptor list -> theory -> theory val add_forward_prfstep: thm -> theory -> theory val add_forward'_prfstep: thm -> theory -> theory val add_backward_prfstep: thm -> theory -> theory val add_backward1_prfstep: thm -> theory -> theory val add_backward2_prfstep: thm -> theory -> theory val add_resolve_prfstep: thm -> theory -> theory val add_rewrite_rule_cond: thm -> pre_prfstep_descriptor list -> theory -> theory val add_rewrite_rule_back_cond: thm -> pre_prfstep_descriptor list -> theory -> theory val add_rewrite_rule_bidir_cond: thm -> pre_prfstep_descriptor list -> theory -> theory val add_rewrite_rule: thm -> theory -> theory val add_rewrite_rule_back: thm -> theory -> theory val add_rewrite_rule_bidir: thm -> theory -> theory val setup_attrib: (thm -> theory -> theory) -> attribute context_parser end; -structure ProofStepData : PROOFSTEP_DATA = +functor ProofStepData( + structure ItemIO: ITEM_IO; + structure Normalizer: NORMALIZER; + structure ProofStep: PROOFSTEP; + structure Property: PROPERTY; + structure UtilBase: UTIL_BASE; + structure UtilLogic: UTIL_LOGIC; + ): PROOFSTEP_DATA = struct structure Data = Theory_Data ( type T = proofstep list; val empty = []; fun merge (ps1, ps2) = Library.merge ProofStep.eq_prfstep (ps1, ps2) ) (* Add the given proof step. *) fun add_prfstep (prfstep as {args, ...}) = Data.map (fn prfsteps => if Util.is_prefix_str "$" (#name prfstep) then error "Add prfstep: names beginning with $ is reserved." else let val num_args = length (filter_out ItemIO.is_side_match args) in if num_args >= 1 andalso num_args <= 2 then prfsteps @ [prfstep] - else error "add_proofstep: need 1 or 2 patterns." + else error "add_prfstep: need 1 or 2 patterns." end) (* Deleting a proofstep. For string inputs, try adding theory name. For theorem inputs, try all @-suffixes. *) fun del_prfstep_pred pred = Data.map (fn prfsteps => let val names = map #name prfsteps val to_delete = filter pred names fun eq_name (key, {name, ...}) = (key = name) in if null to_delete then error "Delete prfstep: not found" else let val _ = writeln (cat_lines (map (fn name => "Delete " ^ name) to_delete)) in subtract eq_name to_delete prfsteps end end) fun del_prfstep prfstep_name thy = del_prfstep_pred (equal prfstep_name) thy (* Delete all proofsteps for a given theorem. *) fun del_prfstep_thm th = let val th_name = Util.name_of_thm th in del_prfstep_pred (equal th_name orf Util.is_prefix_str (th_name ^ "@")) end (* Delete proofsteps for a given theorem, with the given postfix. *) fun del_prfstep_thm_str str th = del_prfstep_pred (equal (Util.name_of_thm th ^ str)) val del_prfstep_thm_eqforward = del_prfstep_thm_str "@eqforward" fun get_prfsteps thy = Data.get thy fun add_prfstep_custom (name, descs, updt_fn) = add_prfstep (ProofStep.prfstep_custom name descs updt_fn) fun add_gen_prfstep (name, descs) = add_prfstep (ProofStep.gen_prfstep name descs) fun add_prfstep_pre_conv (name, descs, pre_cv) = add_prfstep (ProofStep.prfstep_pre_conv name descs pre_cv) fun add_prfstep_conv (name, descs, cv) = add_prfstep (ProofStep.prfstep_conv name descs cv) (* Constructing conditional prfstep_descriptors. *) -type pre_prfstep_descriptor = Proof.context -> prfstep_descriptor +type pre_prfstep_descriptor = pre_prfstep_descriptor fun with_term str ctxt = let val t = Proof_Context.read_term_pattern ctxt str val _ = assert (null (Term.add_frees t [])) "with_term: should not contain free variable." in - WithTerm t + ProofStep.WithTerm t end fun with_cond str ctxt = - Filter (neq_filter (Proof_Context.read_term_pattern ctxt str)) + Filter (ProofStep.neq_filter (Proof_Context.read_term_pattern ctxt str)) fun with_conds strs = map with_cond strs fun with_filt filt = K (Filter filt) fun with_filts filts = map with_filt filts fun with_score n = K (WithScore n) (* Second level proofstep writing functions. *) fun add_and_print_prfstep prfstep_name descs thy = let val _ = writeln (prfstep_name ^ "\n" ^ (ProofStep.string_of_descs thy descs)) in add_gen_prfstep (prfstep_name, descs) thy end (* Add a proofstep checking a requirement. *) fun add_prfstep_check_req (t_str, req_str) thy = let val ctxt = Proof_Context.init_global thy val t = Proof_Context.read_term_pattern ctxt t_str val vars = map Free (Term.add_frees t []) val c = Util.get_head_name t val ctxt' = fold Util.declare_free_term vars ctxt val req = Proof_Context.read_term_pattern ctxt' req_str fun get_subst var = case var of Free (x, T) => (var, Var ((x, 0), T)) | _ => raise Fail "add_prfstep_check_req" val subst = map get_subst vars val t' = Term.subst_atomic subst t val req' = Term.subst_atomic subst req in add_and_print_prfstep - (c ^ "_case") [WithTerm t', CreateConcl req'] thy + (c ^ "_case") [ProofStep.WithTerm t', CreateConcl req'] thy end datatype prfstep_mode = MODE_FORWARD | MODE_FORWARD' | MODE_BACKWARD | MODE_BACKWARD1 | MODE_BACKWARD2 | MODE_RESOLVE (* Maximum number of term matches for the given mode. *) fun max_term_matches mode = case mode of MODE_FORWARD => 2 | MODE_FORWARD' => 1 | MODE_BACKWARD => 1 | MODE_RESOLVE => 1 | _ => 0 (* Obtain the first several premises of th that are either properties or wellformed-ness data. ts is the list of term matches. *) fun get_side_prems thy mode ts th = let val (prems, concl) = UtilLogic.strip_horn' th val _ = assert (length ts <= max_term_matches mode) "get_side_prems: too many term matches." (* Helper function. Consider the case where the first n premises are side conditions. Find the additional terms to match against for each mode. *) fun additional_matches n = let val prems' = drop n prems in case mode of MODE_FORWARD => take (2 - length ts) prems' | MODE_FORWARD' => if null ts andalso length prems' >= 2 then [hd prems', List.last prems'] else [List.last prems'] - | MODE_BACKWARD => [get_neg concl] - | MODE_BACKWARD1 => [get_neg concl, List.last prems'] - | MODE_BACKWARD2 => [get_neg concl, hd prems'] + | MODE_BACKWARD => [UtilLogic.get_neg concl] + | MODE_BACKWARD1 => [UtilLogic.get_neg concl, List.last prems'] + | MODE_BACKWARD2 => [UtilLogic.get_neg concl, hd prems'] | MODE_RESOLVE => if null ts andalso length prems' > 0 then - [get_neg concl, List.last prems'] - else [get_neg concl] + [UtilLogic.get_neg concl, List.last prems'] + else [UtilLogic.get_neg concl] end (* Determine whether t is a valid side premises, relative to the matches ts'. If yes, return the corresponding side matching. Otherwise return NONE. *) fun to_side_prems ts' t = case WellForm.is_subterm_wellform_data thy t ts' of SOME (t, req) => SOME (WithWellForm (t, req)) | NONE => if Property.is_property_prem thy t then SOME (WithProperty t) else NONE (* Attempt to convert the first n premises to side matchings. *) fun to_side_prems_n n = let val ts' = additional_matches n @ ts val side_prems' = prems |> take n |> map (to_side_prems ts') in if forall is_some side_prems' then SOME (map the side_prems') else NONE end (* Minimum number of premises for the given mode. *) val min_prems = case mode of MODE_FORWARD => 1 - length ts | MODE_FORWARD' => 1 | MODE_BACKWARD => 1 | MODE_BACKWARD1 => 2 | MODE_BACKWARD2 => 2 | MODE_RESOLVE => 0 val _ = assert (length prems >= min_prems) "get_side_prems: too few premises." val to_test = rev (0 upto (length prems - min_prems)) in (* Always succeeds at 0. *) the (get_first to_side_prems_n to_test) end (* Convert theorems of the form A1 ==> ... ==> An ==> C to A1 & ... & An ==> C. If keep_last = true, the last assumption is kept in implication form. *) fun atomize_conj_cv keep_last ct = if length (Logic.strip_imp_prems (Thm.term_of ct)) <= (if keep_last then 2 else 1) then Conv.all_conv ct else Conv.every_conv [Conv.arg_conv (atomize_conj_cv keep_last), Conv.rewr_conv UtilBase.atomize_conjL_th] ct (* Swap the last premise to become the first. *) fun swap_prem_to_front ct = let val n = length (Logic.strip_imp_prems (Thm.term_of ct)) in if n < 2 then Conv.all_conv ct else if n = 2 then Conv.rewr_conv Drule.swap_prems_eq ct else ((Conv.arg_conv swap_prem_to_front) then_conv (Conv.rewr_conv Drule.swap_prems_eq)) ct end (* Using cv, rewrite all assumptions and conclusion in ct. *) fun horn_conv cv ct = (case Thm.term_of ct of @{const Pure.imp} $ _ $ _ => - (Conv.arg1_conv (Trueprop_conv cv)) + (Conv.arg1_conv (UtilLogic.Trueprop_conv cv)) then_conv (Conv.arg_conv (horn_conv cv)) - | _ => Trueprop_conv cv) ct + | _ => UtilLogic.Trueprop_conv cv) ct (* Try to cancel terms of the form ~~A. *) -val try_nn_cancel_cv = Conv.try_conv (rewr_obj_eq UtilBase.nn_cancel_th) +val try_nn_cancel_cv = Conv.try_conv (UtilLogic.rewr_obj_eq UtilBase.nn_cancel_th) (* Post-processing of the given theorem according to mode. *) fun post_process_th ctxt mode side_count ts th = case mode of MODE_FORWARD => let val to_skip = side_count + (2 - length ts) in th |> apply_to_thm (Util.skip_n_conv to_skip (UtilLogic.to_obj_conv ctxt)) |> Util.update_name_of_thm th "" end | MODE_FORWARD' => let val cv = swap_prem_to_front then_conv (Util.skip_n_conv (2 - length ts) (UtilLogic.to_obj_conv ctxt)) in th |> apply_to_thm (Util.skip_n_conv side_count cv) |> Util.update_name_of_thm th "" end | MODE_BACKWARD => let val cv = (atomize_conj_cv false) then_conv (Conv.rewr_conv UtilBase.backward_conv_th) then_conv (horn_conv try_nn_cancel_cv) in th |> apply_to_thm (Util.skip_n_conv side_count cv) |> Util.update_name_of_thm th "@back" end | MODE_BACKWARD1 => let val cv = (atomize_conj_cv true) then_conv (Conv.rewr_conv UtilBase.backward1_conv_th) then_conv (horn_conv try_nn_cancel_cv) in th |> apply_to_thm (Util.skip_n_conv side_count cv) |> Util.update_name_of_thm th "@back1" end | MODE_BACKWARD2 => let val cv = (Conv.arg_conv (atomize_conj_cv false)) then_conv (Conv.rewr_conv UtilBase.backward2_conv_th) then_conv (horn_conv try_nn_cancel_cv) in th |> apply_to_thm (Util.skip_n_conv side_count cv) |> Util.update_name_of_thm th "@back2" end | MODE_RESOLVE => let val rewr_th = case Thm.nprems_of th - side_count of - 0 => if is_neg (concl_of' th) then UtilBase.to_contra_form_th' + 0 => if UtilLogic.is_neg (UtilLogic.concl_of' th) then UtilBase.to_contra_form_th' else UtilBase.to_contra_form_th | 1 => UtilBase.resolve_conv_th | _ => raise Fail "resolve: too many hypothesis in th." val cv = (Conv.rewr_conv rewr_th) then_conv (horn_conv try_nn_cancel_cv) in th |> apply_to_thm (Util.skip_n_conv side_count cv) |> Util.update_name_of_thm th "@res" end (* Add basic proofstep for the given theorem and mode. *) fun add_basic_prfstep_cond th mode conds thy = let val ctxt = Proof_Context.init_global thy val ctxt' = ctxt |> Variable.declare_term (Thm.prop_of th) (* Replace variable definitions, obtaining list of replacements and the new theorem. *) val (pairs, th) = th |> apply_to_thm (UtilLogic.to_obj_conv_on_horn ctxt') |> Normalizer.meta_use_vardefs |> apsnd (Util.update_name_of_thm th "") (* List of definitions used. *) fun print_def_subst (lhs, rhs) = writeln ("Apply def " ^ (Syntax.string_of_term ctxt' lhs) ^ " = " ^ (Syntax.string_of_term ctxt' rhs)) val _ = map print_def_subst pairs fun def_subst_fun cond = case cond of WithItem ("TERM", t) => WithItem ("TERM", Normalizer.def_subst pairs t) | _ => cond in if null conds andalso (mode = MODE_FORWARD orelse mode = MODE_FORWARD') andalso Property.can_add_property_update th thy then Property.add_property_update th thy else let fun is_term_cond cond = case cond of WithItem ("TERM", _) => true | _ => false fun extract_term_cond cond = case cond of WithItem ("TERM", t) => t | _ => raise Fail "extract_term_cond" (* Instantiate each element of conds with ctxt', then separate into term and other (filter and shadow) conds. *) val (term_conds, filt_conds) = conds |> map (fn cond => cond ctxt') |> filter_split is_term_cond |> apfst (map def_subst_fun) (* Get list of assumptions to be obtained from either the property table or the wellform table. *) val ts = map extract_term_cond term_conds val side_prems = get_side_prems thy mode ts th val side_count = length side_prems val th' = th |> post_process_th ctxt' mode side_count ts val (assums, concl) = th' |> UtilLogic.strip_horn' |> apfst (drop side_count) val pats = map extract_term_cond term_conds @ assums val match_descs = term_conds @ map WithFact assums val _ = assert (Util.is_pattern_list pats) "add_basic_prfstep: invalid patterns." val _ = assert (length pats > 0 andalso length pats <= 2) "add_basic_prfstep: invalid number of patterns." in (* Switch two assumptions if necessary. *) if length pats = 2 andalso not (Util.is_pattern (hd pats)) then let val _ = writeln "Switching two patterns." val swap_prems_cv = Conv.rewr_conv Drule.swap_prems_eq val th'' = if length assums = 1 then th' else th' |> apply_to_thm (Util.skip_n_conv side_count swap_prems_cv) |> Util.update_name_of_thm th' "" val swap_match_descs = [nth match_descs 1, hd match_descs] val descs = side_prems @ swap_match_descs @ filt_conds @ [GetFact (concl, th'')] in add_and_print_prfstep (Util.name_of_thm th') descs thy end else let val descs = side_prems @ match_descs @ filt_conds @ [GetFact (concl, th')] in add_and_print_prfstep (Util.name_of_thm th') descs thy end end end fun add_forward_prfstep_cond th = add_basic_prfstep_cond th MODE_FORWARD fun add_forward'_prfstep_cond th = add_basic_prfstep_cond th MODE_FORWARD' fun add_backward_prfstep_cond th = add_basic_prfstep_cond th MODE_BACKWARD fun add_backward1_prfstep_cond th = add_basic_prfstep_cond th MODE_BACKWARD1 fun add_backward2_prfstep_cond th = add_basic_prfstep_cond th MODE_BACKWARD2 fun add_resolve_prfstep_cond th = add_basic_prfstep_cond th MODE_RESOLVE fun add_forward_prfstep th = add_forward_prfstep_cond th [] fun add_forward'_prfstep th = add_forward'_prfstep_cond th [] fun add_backward_prfstep th = add_backward_prfstep_cond th [] fun add_backward1_prfstep th = add_backward1_prfstep_cond th [] fun add_backward2_prfstep th = add_backward2_prfstep_cond th [] fun add_resolve_prfstep th = add_resolve_prfstep_cond th [] fun add_rewrite_eq_rule_cond th conds thy = let val th = if Util.is_meta_eq (Thm.concl_of th) then UtilLogic.to_obj_eq_th th else th - val (lhs, _) = th |> concl_of' |> strip_conj |> hd |> dest_eq + val (lhs, _) = th |> UtilLogic.concl_of' |> UtilLogic.strip_conj |> hd |> UtilBase.dest_eq in - thy |> add_forward_prfstep_cond th (K (WithTerm lhs) :: conds) + thy |> add_forward_prfstep_cond th (K (ProofStep.WithTerm lhs) :: conds) end fun add_rewrite_iff_rule_cond th conds thy = let val th = if Util.is_meta_eq (Thm.concl_of th) then UtilLogic.to_obj_eq_iff_th th else th - val (lhs, _) = th |> concl_of' |> dest_eq - val _ = assert (fastype_of lhs = boolT) + val (lhs, _) = th |> UtilLogic.concl_of' |> UtilBase.dest_eq + val _ = assert (fastype_of lhs = UtilBase.boolT) "add_rewrite_iff: argument not of type bool." - val forward_th = th |> equiv_forward_th - val nforward_th = th |> inv_backward_th + val forward_th = th |> UtilLogic.equiv_forward_th + val nforward_th = th |> UtilLogic.inv_backward_th |> apply_to_thm (horn_conv try_nn_cancel_cv) |> Util.update_name_of_thm th "@invbackward" in thy |> add_basic_prfstep_cond forward_th MODE_FORWARD' conds |> add_basic_prfstep_cond nforward_th MODE_FORWARD' conds end fun add_rewrite_rule_cond th conds thy = let - val th = if Util.is_meta_eq (Thm.concl_of th) then to_obj_eq_th th else th - val (lhs, _) = th |> concl_of' |> strip_conj |> hd |> dest_eq + val th = if Util.is_meta_eq (Thm.concl_of th) then UtilLogic.to_obj_eq_th th else th + val (lhs, _) = th |> UtilLogic.concl_of' |> UtilLogic.strip_conj |> hd |> UtilBase.dest_eq in - if fastype_of lhs = boolT then + if fastype_of lhs = UtilBase.boolT then add_rewrite_iff_rule_cond th conds thy else add_rewrite_eq_rule_cond th conds thy end fun add_rewrite_rule_back_cond th conds = - add_rewrite_rule_cond (obj_sym_th th) conds + add_rewrite_rule_cond (UtilLogic.obj_sym_th th) conds fun add_rewrite_rule_bidir_cond th conds = (add_rewrite_rule_cond th conds) #> add_rewrite_rule_back_cond th conds fun add_rewrite_rule th = add_rewrite_rule_cond th [] fun add_rewrite_rule_back th = add_rewrite_rule_back_cond th [] fun add_rewrite_rule_bidir th = add_rewrite_rule th #> add_rewrite_rule_back th fun setup_attrib f = Attrib.add_del (Thm.declaration_attribute ( fn th => Context.mapping (f th) I)) (Thm.declaration_attribute ( fn _ => fn _ => raise Fail "del_step: not implemented.")) end (* structure ProofStepData. *) - -open ProofStepData diff --git a/thys/Auto2_HOL/property.ML b/thys/Auto2_HOL/property.ML --- a/thys/Auto2_HOL/property.ML +++ b/thys/Auto2_HOL/property.ML @@ -1,255 +1,257 @@ (* File: property.ML Author: Bohua Zhan Theory data for properties. This data consists of the following parts: - Two tables containing property update rules. - A table containing list of fields that can have properties. *) signature PROPERTY = sig val is_property: term -> bool val add_property_field_const: term -> theory -> theory val is_property_field: theory -> term -> bool val strip_property_field: theory -> term -> term list val is_property_prem: theory -> term -> bool val get_property_name: term -> string val get_property_names: term list -> string list val get_property_arg: term -> term val get_property_arg_th: thm -> cterm (* About the PropertyUpdateData table.*) val can_add_property_update: thm -> theory -> bool val add_property_update: thm -> theory -> theory val lookup_property_update: theory -> string -> thm list val lookup_property_update_fun: theory -> string -> thm list val instantiate_property_update: Proof.context -> term -> thm -> thm option end; -structure Property : PROPERTY = +functor Property( + structure Basic_UtilBase: BASIC_UTIL_BASE; + structure UtilLogic : UTIL_LOGIC + ) : PROPERTY = struct (* Rules deriving new properties of t from other properties of t. They are indexed under the names of the properties in the premises. *) structure UpdateData = Theory_Data ( type T = (thm list) Symtab.table; val empty = Symtab.empty; val merge = Symtab.merge_list Thm.eq_thm_prop ) (* Rules for deriving properties of f x_1 ... x_n from properties of x_1, ... x_n. They are indexed under the name of the head function f. *) structure UpdateFunData = Theory_Data ( type T = (thm list) Symtab.table; val empty = Symtab.empty; val merge = Symtab.merge_list Thm.eq_thm_prop ) (* Set of fields of a structure whose property can be considered as properties of the structure itself. Relevant when checking is_property_prem. *) structure FieldData = Theory_Data ( type T = unit Symtab.table; val empty = Symtab.empty; val merge = Symtab.merge (K true) ) (* Whether the term is a property predicate applied to a term. *) fun is_property t = let - val _ = assert (fastype_of t = boolT) "is_property: wrong type" + val _ = assert (fastype_of t = Basic_UtilBase.boolT) "is_property: wrong type" val (f, ts) = Term.strip_comb t in if length ts <> 1 orelse not (Term.is_Const f) then false else let val T = fastype_of (the_single ts) val (dT, _) = Term.strip_type T in - length dT = 0 andalso T <> boolT + length dT = 0 andalso T <> Basic_UtilBase.boolT end end (* Insert the following constant as a property field. *) fun add_property_field_const t thy = case Term.head_of t of Const (c, T) => let val (pTs, _) = Term.strip_type T val _ = if length pTs = 1 then () else error "Add property field: input should be a field." val _ = writeln ("Add field " ^ c ^ " as property field.") in thy |> FieldData.map (Symtab.update_new (c, ())) end | _ => error "Add property field: input should be a constant." (* Whether the term is zero or more property field constants applied to a Var term. *) fun is_property_field thy t = case t of Var _ => true | Const (c, _) $ t' => Symtab.defined (FieldData.get thy) c andalso is_property_field thy t' | _ => false (* Given a term t, return all possible ways to strip property field constants from t. For example, if t is of the form f1(f2(x)), where f1 and f2 are property constants, then the result is [f1(f2(x)), f2(x), x]. *) fun strip_property_field thy t = case t of Const (c, _) $ t' => if Symtab.defined (FieldData.get thy) c then t :: strip_property_field thy t' else [t] | _ => [t] (* Stricter condition than is_property: the argument must be a schematic variable (up to property fields). *) fun is_property_prem thy t = is_property t andalso is_property_field thy (dest_arg t) val get_property_name = Util.get_head_name fun get_property_names ts = ts |> map get_property_name |> distinct (op =) (* Return the argument of the property. *) fun get_property_arg t = dest_arg t handle Fail "dest_arg" => raise Fail "get_property_arg: t in wrong form." (* Return the argument of the property theorem as a cterm. *) fun get_property_arg_th th = - Thm.dest_arg (cprop_of' th) + Thm.dest_arg (UtilLogic.cprop_of' th) handle CTERM _ => raise Fail "get_property_carg" | Fail "dest_Trueprop" => raise Fail "get_property_carg" (* Add the given rule as a property update. The requirements on th is as follows: - The conclusion must be a property constant, with argument in the form of either ?x or f ?x1 ... ?xn. - Each premise must be a property constant on ?x (in the first case) or one of ?x1 ... ?xn (in the second case). The argument of the property in the conclusion must contain all schematic variables of the theorem. *) fun can_add_property_update th thy = let val (prems, concl) = UtilLogic.strip_horn' th in if is_property concl andalso forall (is_property_prem thy) prems then let val concl_arg = get_property_arg concl val all_vars = map Var (Term.add_vars (Thm.prop_of th) []) in if is_Var concl_arg then (* First case: check that concl_arg is the only schematic Var. *) length all_vars = 1 else (* Second case: concl_arg is of the form f ?x1 ... ?xn. *) let val args = Util.dest_args concl_arg in forall is_Var args andalso subset (op aconv) (all_vars, args) end end else false end (* Add the given theorem as a property update rule. Choose which table to add the rule to. *) fun add_property_update th thy = let val (prems, concl) = UtilLogic.strip_horn' th val _ = assert (is_property concl) "add_property_update: concl must be a property constant." val _ = assert (forall (is_property_prem thy) prems) "add_property_update: prem must be a property premise." val concl_arg = get_property_arg concl val all_vars = map Var (Term.add_vars (Thm.prop_of th) []) in if is_Var concl_arg then (* First case. Each premise must also be about ?x. Add to UpdateData table under the names of the predicates. *) let val _ = assert (length all_vars = 1) "add_property_update: extraneous Vars in th." val names = get_property_names prems val _ = writeln ("Add property rule for " ^ (Util.string_of_list I names)) in thy |> UpdateData.map ( fold (Symtab.update_list Thm.eq_thm_prop) (map (rpair th) names)) end else (* Second case. concl_arg in the form f ?x1 ... ?xn. Add to UpdateFunData table under f. *) let val (f, args) = Term.strip_comb concl_arg val c = case f of Const (c, _) => c | _ => raise Fail "add_property_update: f is not constant." val _ = assert (forall is_Var args) "add_property_update: all args of concl must be Vars." val _ = assert (subset (op aconv) (all_vars, args)) "add_property_update: extraneous Vars in th." val _ = writeln ("Add property rule for function " ^ c) in thy |> UpdateFunData.map (Symtab.update_list Thm.eq_thm_prop (c, th)) end end (* Find update rules of the form P1 x ==> ... ==> Pn x ==> P x, where c is one of P1, ... Pn. *) fun lookup_property_update thy c = Symtab.lookup_list (UpdateData.get thy) c (* Find update rules of the form A1 ==> ... ==> An ==> P(f(x1,...,xn)), where each A_i is a property on one of x_j. Here c is the name of f. *) fun lookup_property_update_fun thy c = Symtab.lookup_list (UpdateFunData.get thy) c (* Instantiate th by matching t with the argument of the conclusion of th. Return NONE if instantiation is unsuccessful (because type does not match). *) fun instantiate_property_update ctxt t th = let val (_, concl) = UtilLogic.strip_horn' th val concl_arg = get_property_arg concl val thy = Proof_Context.theory_of ctxt in if Sign.typ_instance thy (fastype_of t, fastype_of concl_arg) then let val err_str = "instantiate_property_update: cannot match with t." val inst = Pattern.first_order_match thy (concl_arg, t) fo_init handle Pattern.MATCH => raise Fail err_str in SOME (Util.subst_thm ctxt inst th) end else NONE end end (* structure Property. *) -val add_property_field_const = Property.add_property_field_const diff --git a/thys/Auto2_HOL/propertydata.ML b/thys/Auto2_HOL/propertydata.ML --- a/thys/Auto2_HOL/propertydata.ML +++ b/thys/Auto2_HOL/propertydata.ML @@ -1,278 +1,282 @@ (* File: propertydata.ML Author: Bohua Zhan Table maintaining list of properties satisfied by terms currently appearing in the proof. Under each key t, maintain a list of (id, th), where th is of the form P(t), with P a registered property constant. *) signature PROPERTY_DATA = sig val clean_resolved: box_id -> Proof.context -> Proof.context val clear_incr: Proof.context -> Proof.context val get_property_for_term: Proof.context -> term -> (box_id * thm) list val add_property_raw: box_id * thm -> Proof.context -> Proof.context val convert_property: Proof.context -> box_id * thm -> box_id * thm -> box_id * thm val get_property: Proof.context -> box_id * cterm -> (box_id * thm) list val get_property_t: Proof.context -> box_id * term -> (box_id * thm) list val apply_property_update_rule: Proof.context -> box_id -> thm option -> (box_id * thm) list val apply_property_update_on_term: Proof.context -> box_id -> term -> (box_id * thm) list val process_update_property: (box_id * thm) list -> Proof.context -> Proof.context val process_rewrite_property: box_id * thm -> Proof.context -> Proof.context val add_property: box_id * thm -> Proof.context -> Proof.context val get_new_property: Proof.context -> (box_id * thm) list end; -structure PropertyData : PROPERTY_DATA = - +functor PropertyData( + structure Basic_UtilBase: BASIC_UTIL_BASE + structure Property: PROPERTY; + structure RewriteTable: REWRITE_TABLE + structure UtilLogic: BASIC_UTIL_LOGIC + ) : PROPERTY_DATA = struct structure Data = Proof_Data ( type T = ((box_id * thm) list) Termtab.table fun init _ = Termtab.empty ) (* Remove all entries at or below id. *) fun clean_resolved id ctxt = let fun clean_property tb = tb |> Termtab.map (fn _ => filter_out (BoxID.is_eq_ancestor ctxt id o fst)) in ctxt |> Data.map clean_property end fun clear_incr ctxt = let fun clear_one infos = if exists (BoxID.has_incr_id o fst) infos then infos |> map (apfst BoxID.replace_incr_id) |> Util.max_partial (BoxID.info_eq_better ctxt) else infos in ctxt |> Data.map (Termtab.map (fn _ => clear_one)) end (* Retrieve the current list of properties for a term t. *) fun get_property_for_term ctxt t = let val property = Data.get ctxt in the_default [] (Termtab.lookup property t) end (* Add a new property. Similar to add_equiv_raw. *) fun add_property_raw (cur_info as (_, th)) ctxt = let - val t = Property.get_property_arg (prop_of' th) + val t = Property.get_property_arg (UtilLogic.prop_of' th) val props = get_property_for_term ctxt t in if exists (fn info => BoxID.info_eq_better ctxt info cur_info) props then ctxt else let val props' = props |> filter_out (BoxID.info_eq_better ctxt cur_info) |> cons cur_info in ctxt |> Data.map (Termtab.update (t, props')) end end (* Given th of the form P(s), and eq_th of the form s == t, return P(t). Merge boxes corresponding to the two theorems. *) fun convert_property ctxt (id', eq_th) (id, th) = (BoxID.merge_boxes ctxt (id, id'), - th |> apply_to_thm' (Conv.arg_conv (Conv.rewr_conv eq_th))) + th |> UtilLogic.apply_to_thm' (Conv.arg_conv (Conv.rewr_conv eq_th))) fun get_property_raw ctxt (id, prop) = let val (P, t) = Term.dest_comb prop in (get_property_for_term ctxt t) - |> filter (fn (_, th) => Term.head_of (prop_of' th) aconv P) + |> filter (fn (_, th) => Term.head_of (UtilLogic.prop_of' th) aconv P) |> BoxID.merge_box_with_info ctxt id |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) end (* Attempt to retrieve property with the given statement from the table. Here the statement is given as a certified term. *) fun get_property ctxt (id, cprop) = let val (P, t) = Term.dest_comb (Thm.term_of cprop) in if RewriteTable.in_table_raw_for_id ctxt (id, t) then get_property_raw ctxt (id, Thm.term_of cprop) else let val ct = Thm.dest_arg cprop fun process_head_rep (id', eq_th) = (get_property_raw ctxt (id, P $ Util.rhs_of eq_th)) |> map (convert_property ctxt (id', meta_sym eq_th)) in (RewriteTable.subterm_simplify_info ctxt ct) |> maps (RewriteTable.get_head_rep_with_id_th ctxt) |> maps process_head_rep |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) end end (* As above, except statement is only a term. *) fun get_property_t ctxt (id, prop) = let val (_, t) = Term.dest_comb prop in if RewriteTable.in_table_raw_for_id ctxt (id, t) then get_property_raw ctxt (id, prop) else get_property ctxt (id, Thm.cterm_of ctxt prop) end (* th is an instantiated property update rule (without schematic variables). All premises and conclusions of th should be properties. Apply this rule at id and below to get new properties. *) fun apply_property_update_rule ctxt id th_opt = case th_opt of NONE => [] | SOME th => let - val prems = map dest_Trueprop (Thm.prems_of th) + val prems = map UtilLogic.dest_Trueprop (Thm.prems_of th) in if null prems then [(id, th)] else prems |> map (pair id) |> map (get_property_raw ctxt) |> BoxID.get_all_merges_info ctxt |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) |> map (fn (id', ths) => (id', ths MRS th)) end (* Find relevant property updates for term t of the form f(x1,...x_n), from new properties of x1, ... xn. Apply these to get list of new properties. *) fun apply_property_update_on_term ctxt id t = - if fastype_of t = boolT then [] else + if fastype_of t = Basic_UtilBase.boolT then [] else case head_of t of Const (c, _) => let val thy = Proof_Context.theory_of ctxt val updt_rules = Property.lookup_property_update_fun thy c fun process_updt_rule th = th |> Property.instantiate_property_update ctxt t |> apply_property_update_rule ctxt id in maps process_updt_rule updt_rules end | _ => [] (* Work out all consequences of adding a list of properties. In addition to propagating properties along equalities, need to apply property update rules. *) fun process_update_property inits ctxt = let val thy = Proof_Context.theory_of ctxt fun eq_property ((id, th), (id', th')) = (id = id' andalso Thm.prop_of th aconv Thm.prop_of th') fun process_property (id, th) (to_process, ctxt) = let - val t = Property.get_property_arg (prop_of' th) + val t = Property.get_property_arg (UtilLogic.prop_of' th) val property_t = get_property_for_term ctxt t in if exists (fn info => BoxID.info_eq_better ctxt info (id, th)) property_t then (to_process, ctxt) else (insert eq_property (id, th) to_process, add_property_raw (id, th) ctxt) end fun update_step (to_process, ctxt) = case to_process of [] => ([], ctxt) | (id, th) :: rest => let - val t = Property.get_property_arg (prop_of' th) + val t = Property.get_property_arg (UtilLogic.prop_of' th) (* Neighbors of t. Here th: P t, th': t = t', result: P t'. *) fun process_neigh (id', eq_th) = convert_property ctxt (id', eq_th) (id, th) val new_property_neigh = map process_neigh (RewriteTable.equiv_neighs ctxt t) (* Derived properties of t. *) val ts = Property.strip_property_field thy t - val c = Property.get_property_name (prop_of' th) + val c = Property.get_property_name (UtilLogic.prop_of' th) val updt_rules = Property.lookup_property_update thy c fun process_updt_rule (t, th) = th |> Property.instantiate_property_update ctxt t |> apply_property_update_rule ctxt id val new_property_t = maps process_updt_rule (Util.all_pairs (ts, updt_rules)) (* Derived properties of parent terms of t. *) val parents_t = map (Thm.term_of) (RewriteTable.immediate_contains ctxt t) val new_property_ps = maps (apply_property_update_on_term ctxt id) parents_t in (rest, ctxt) |> fold process_property new_property_neigh |> fold process_property new_property_t |> fold process_property new_property_ps |> update_step end in ([], ctxt) |> fold process_property inits |> update_step |> snd end (* Work out all consequences of adding an equality for the property table. *) fun process_rewrite_property (id, th) ctxt = let val (t1, t2) = (Util.lhs_of th, Util.rhs_of th) (* New properties. *) val t1_property = get_property_for_term ctxt t1 val t2_property = get_property_for_term ctxt t2 val t1_newp = t2_property |> map (convert_property ctxt (id, meta_sym th)) val t2_newp = t1_property |> map (convert_property ctxt (id, th)) in ctxt |> process_update_property (t1_newp @ t2_newp) end (* First make sure t is in the table. Add property P(t) to the table, and work out any consequences. The result is a consistent property table. *) fun add_property (id, th) ctxt = let val ct = Property.get_property_arg_th th in ctxt |> RewriteTable.add_term (id, ct) |> snd |> process_update_property [(id, th)] end (* Return the list of properties that depend on prim_id, as a list of (id, th). *) fun get_new_property ctxt = (Data.get ctxt) |> Termtab.dest_list |> map snd |> filter (fn (id, _) => BoxID.has_incr_id id) end (* structure PropertyData *) diff --git a/thys/Auto2_HOL/rewrite.ML b/thys/Auto2_HOL/rewrite.ML --- a/thys/Auto2_HOL/rewrite.ML +++ b/thys/Auto2_HOL/rewrite.ML @@ -1,915 +1,915 @@ (* File: rewrite.ML Author: Bohua Zhan Maintains the congruence closure of the currently known equalities. Automatically apply transitivity and the congruence property. *) (* Order on equality theorems, by comparing the right side. *) val simp_ord = Term_Ord.term_ord o (apply2 Util.rhs_of) (* Equality on box_id * thm, comparing the right side. *) fun eq_info ((id, th), (id', th')) = (id = id' andalso Util.rhs_of th aconv Util.rhs_of th') (* Debug functions. *) fun print_info ctxt (id, th) = "(" ^ (BoxID.string_of_box_id id) ^ ", " ^ (th |> Util.rhs_of |> Syntax.string_of_term ctxt) ^ ")" fun print_infos ctxt lst = commas (map (print_info ctxt) lst) fun print_info' ctxt (id, ths) = "(" ^ (BoxID.string_of_box_id id) ^ ", " ^ (ths |> map Util.rhs_of |> Util.string_of_terms ctxt) ^ ")" fun print_infos' ctxt lst = cat_lines (map (print_info' ctxt) lst) (* Data structure for rewrite table. *) type rewrite_table = { (* terms[t] is the list of ids at which t is present. *) terms: (box_id list * cterm) Termtab.table, (* equiv[a] is list of (id, th), where th is "a == a'" under id. *) equiv: ((box_id * thm) list) Termtab.table, (* Index of reachability under equiv. *) all_equiv: ((box_id * thm) list) Termtab.table, (* contain[A] is the list of terms for which A is a direct subterm. *) contain: (cterm list) Termtab.table, (* Rewrite info of a term. *) simp: ((box_id * thm) list) Termtab.table, (* Subterm rewrite info of a term. *) subsimp: ((box_id * thm) list) Termtab.table, (* Reverse of subsimp. Head rep of a term. *) reps: ((box_id * thm) list) Termtab.table } signature REWRITE_TABLE = sig (* Construction and modification. *) val clean_resolved: box_id -> Proof.context -> Proof.context val clear_incr: Proof.context -> Proof.context (* contain table. *) val add_contain: cterm -> term -> Proof.context -> Proof.context val immediate_contains: Proof.context -> term -> cterm list (* terms table. *) val in_table_raw_ids: Proof.context -> term -> box_id list val in_table_raw_for_id: Proof.context -> box_id * term -> bool val in_table_raw: Proof.context -> term -> bool val add_term_raw: box_id * cterm -> Proof.context -> Proof.context val get_all_terms: Proof.context -> cterm list val get_all_id_terms: Proof.context -> (box_id * cterm) list (* equiv table. *) val add_equiv: box_id * thm -> Proof.context -> Proof.context val equiv_neighs: Proof.context -> term -> (box_id * thm) list val get_all_equiv: Proof.context -> term -> (box_id * thm) list (* simp table. *) val update_simp: box_id * thm -> Proof.context -> Proof.context val get_rewrite_info: Proof.context -> cterm -> (box_id * thm) list val get_rewrite: box_id -> Proof.context -> cterm -> thm (* subsimp and rep tables. *) val get_subterm_rewrite_info: Proof.context -> cterm -> (box_id * thm) list val remove_rep: box_id * thm -> Proof.context -> Proof.context val update_subsimp: box_id * thm -> Proof.context -> Proof.context val get_head_rep_info: Proof.context -> term -> (box_id * thm) list val get_head_rep: box_id -> Proof.context -> term -> thm option val get_head_rep_with_id_th: Proof.context -> box_id * thm -> (box_id * thm) list val get_cached_subterm_rewrite_info: Proof.context -> term -> (box_id * thm) list val get_cached_subterm_rewrite: box_id -> Proof.context -> cterm -> thm (* Simplification of arbitrary terms. *) val head_simplify: box_id -> Proof.context -> cterm -> thm val simplify: box_id -> Proof.context -> cterm -> thm val subterm_simplify: box_id -> Proof.context -> cterm -> thm val simp_val: box_id -> Proof.context -> cterm -> cterm val simp_val_t: box_id -> Proof.context -> term -> term val simplify_info: Proof.context -> cterm -> (box_id * thm) list val subterm_simplify_info: Proof.context -> cterm -> (box_id * thm) list val is_equiv: box_id -> Proof.context -> cterm * cterm -> bool val is_equiv_t: box_id -> Proof.context -> term * term -> bool (* Computation of head equiv. *) val get_head_equiv: Proof.context -> cterm -> (box_id * thm) list val get_head_equiv_with_t: Proof.context -> box_id * cterm -> term -> (box_id * thm) list (* Update of whole rewrite table. *) val process_update_simp: (box_id * thm) list -> Proof.context -> Proof.context val equiv_info: Proof.context -> box_id -> cterm * cterm -> (box_id * thm) list val equiv_info_t: Proof.context -> box_id -> term * term -> (box_id * thm) list val subequiv_info: Proof.context -> box_id -> cterm * cterm -> (box_id * thm) list val add_rewrite_raw: box_id -> thm -> Proof.context -> Proof.context val get_reachable_terms: bool -> Proof.context -> term list -> term list val complete_table: Proof.context -> (box_id * thm) list * Proof.context val add_term: box_id * cterm -> Proof.context -> (box_id * thm) list * Proof.context val add_term_list: (box_id * cterm) list -> Proof.context -> (box_id * thm) list * Proof.context val add_rewrite: box_id * thm -> Proof.context -> (box_id * thm) list * Proof.context (* Updating rewrite table from list of (id, th) pairs. *) val get_new_terms: Proof.context * Proof.context -> (box_id * cterm) list end; -structure RewriteTable : REWRITE_TABLE = +functor RewriteTable(UtilLogic : UTIL_LOGIC) : REWRITE_TABLE = struct (*** Initialization and modification of rewrite table ***) structure Data = Proof_Data ( type T = rewrite_table fun init _ = {terms = Termtab.empty, equiv = Termtab.empty, all_equiv = Termtab.empty, contain = Termtab.empty, simp = Termtab.empty, subsimp = Termtab.empty, reps = Termtab.empty} ) fun clean_resolved id ctxt = let val {terms, equiv, all_equiv, contain, simp, subsimp, reps} = Data.get ctxt fun clean tb id_fun = tb |> Termtab.map (fn _ => filter_out (BoxID.is_eq_ancestor ctxt id o id_fun)) fun clean_terms_1 (ids, ct) = (filter_out (BoxID.is_eq_ancestor ctxt id) ids, ct) fun clean_terms tb = tb |> Termtab.map (fn _ => clean_terms_1) in ctxt |> Data.map ( K {terms = clean_terms terms, equiv = clean equiv fst, all_equiv = clean all_equiv fst, contain = contain, simp = clean simp fst, subsimp = clean subsimp fst, reps = clean reps fst}) end (* equiv_eq_better ctxt info1 info2 means info2 is extraneous in the equiv table. *) fun equiv_eq_better ctxt (id, th) (id', th') = Util.rhs_of th aconv Util.rhs_of th' andalso BoxID.is_eq_ancestor ctxt id id' (* simp_eq_better ctxt info1 info2 means info2 is extraneous in the simp table. *) fun simp_eq_better ctxt (id, th) (id', th') = if BoxID.is_eq_ancestor ctxt id id' then let val (t, t') = (Util.rhs_of th, Util.rhs_of th') in t aconv t' orelse Term_Ord.term_ord (t, t') = LESS end else false fun clear_incr ctxt = let val {terms, equiv, all_equiv, contain, simp, subsimp, reps} = Data.get ctxt fun clear_one cmp infos = if exists (BoxID.has_incr_id o fst) infos then infos |> map (apfst BoxID.replace_incr_id) |> Util.max_partial cmp else infos fun clear cmp tb = tb |> Termtab.map (fn _ => clear_one cmp) fun clear_term_one (ids, ct) = if exists BoxID.has_incr_id ids then (ids |> map BoxID.replace_incr_id |> Util.max_partial (BoxID.is_eq_ancestor ctxt), ct) else (ids, ct) fun clear_term tb = tb |> Termtab.map (fn _ => clear_term_one) in ctxt |> Data.map ( K {terms = clear_term terms, equiv = clear (equiv_eq_better ctxt) equiv, all_equiv = clear (equiv_eq_better ctxt) all_equiv, contain = contain, simp = clear (simp_eq_better ctxt) simp, subsimp = clear (simp_eq_better ctxt) subsimp, reps = clear (equiv_eq_better ctxt) reps}) end fun map_terms f {terms, equiv, all_equiv, contain, simp, subsimp, reps} = {terms = f terms, equiv = equiv, all_equiv = all_equiv, contain = contain, simp = simp, subsimp = subsimp, reps = reps} fun map_equiv f {terms, equiv, all_equiv, contain, simp, subsimp, reps} = {terms = terms, equiv = f equiv, all_equiv = all_equiv, contain = contain, simp = simp, subsimp = subsimp, reps = reps} fun map_all_equiv f {terms, equiv, all_equiv, contain, simp, subsimp, reps} = {terms = terms, equiv = equiv, all_equiv = f all_equiv, contain = contain, simp = simp, subsimp = subsimp, reps = reps} fun map_contain f {terms, equiv, all_equiv, contain, simp, subsimp, reps} = {terms = terms, equiv = equiv, all_equiv = all_equiv, contain = f contain, simp = simp, subsimp = subsimp, reps = reps} fun map_simp f {terms, equiv, all_equiv, contain, simp, subsimp, reps} = {terms = terms, equiv = equiv, all_equiv = all_equiv, contain = contain, simp = f simp, subsimp = subsimp, reps = reps} fun map_subsimp f {terms, equiv, all_equiv, contain, simp, subsimp, reps} = {terms = terms, equiv = equiv, all_equiv = all_equiv, contain = contain, simp = simp, subsimp = f subsimp, reps = reps} fun map_reps f {terms, equiv, all_equiv, contain, simp, subsimp, reps} = {terms = terms, equiv = equiv, all_equiv = all_equiv, contain = contain, simp = simp, subsimp = subsimp, reps = f reps} (*** Basic manipulation of contain table. ***) (* Add the information that t contains subt directly. *) fun add_contain ct subt ctxt = let val {contain, ...} = Data.get ctxt val contain_subt = the_default [] (Termtab.lookup contain subt) in if member (op aconvc) contain_subt ct then ctxt else ctxt |> Data.map (map_contain (Termtab.update (subt, ct :: contain_subt))) end (* Get all terms that directly contains t. *) fun immediate_contains ctxt t = let val {contain, ...} = Data.get ctxt in the_default [] (Termtab.lookup contain t) end (*** Basic manipulation of the terms table. ***) (* Return ids in which t is directly in table. *) fun in_table_raw_ids ctxt t = let val {terms, ...} = Data.get ctxt in case Termtab.lookup terms t of NONE => [] | SOME (ids, _) => ids end (* Whether the term is present in table at the given id. *) fun in_table_raw_for_id ctxt (id, t) = exists (BoxID.is_eq_descendent ctxt id) (in_table_raw_ids ctxt t) (* Whether the term is present at any id. *) fun in_table_raw ctxt t = length (in_table_raw_ids ctxt t) > 0 (* Add t to table at the given id. *) fun add_term_raw (id, ct) ctxt = let val t = Thm.term_of ct val prev_ids = in_table_raw_ids ctxt t val ids_t = (id :: prev_ids) |> Util.max_partial (BoxID.is_eq_ancestor ctxt) in ctxt |> Data.map (map_terms (Termtab.update (t, (ids_t, ct)))) end (* Return list of all terms (at any ids) present in table. *) fun get_all_terms ctxt = let val {terms, ...} = Data.get ctxt in map (fn (_, (_, ct)) => ct) (Termtab.dest terms) end (* Return list of all terms with ids present in table. *) fun get_all_id_terms ctxt = let val {terms, ...} = Data.get ctxt in maps (fn (_, (ids, ct)) => map (rpair ct) ids) (Termtab.dest terms) end (*** Basic manipulation of equiv table. ***) fun has_equiv_eq_better ctxt infos info' = exists (fn info => equiv_eq_better ctxt info info') infos (* Add one way equivalence edge. Maintain the non-comparable property of the set of equivalence edges between any two nodes. cur_info is of form (id, t == t'). *) fun add_equiv_raw (cur_info as (_, th)) ctxt = let val {equiv, ...} = Data.get ctxt val t = Util.lhs_of th val equiv_t = the_default [] (Termtab.lookup equiv t) in if exists (fn info => equiv_eq_better ctxt info cur_info) equiv_t then ctxt else ctxt |> Data.map (map_equiv ( Termtab.update ( t, equiv_t |> filter_out (equiv_eq_better ctxt cur_info) |> cons cur_info))) end (* Add the given equivalence edge. *) fun add_equiv (id, th) ctxt = let val _ = assert (not (Util.lhs_of th aconv Util.rhs_of th)) "add_equiv: t == t'" in ctxt |> add_equiv_raw (id, th) |> add_equiv_raw (id, meta_sym th) end (* Get all edges coming from t in the equiv graph. *) fun equiv_neighs ctxt t = let val {equiv, ...} = Data.get ctxt in case Termtab.lookup equiv t of NONE => [] | SOME infos => infos end fun get_all_equiv ctxt t = let val {all_equiv, ...} = Data.get ctxt in the (Termtab.lookup all_equiv t) handle Option.Option => raise Fail "all_equiv: not found" end (*** Basic manipulation of simp table. ***) (* Update rewrite of t under id to th: t == t'. Similar to add_equiv_raw. *) fun update_simp (id, th) ctxt = let val {simp, ...} = Data.get ctxt val ct = Thm.lhs_of th val t = Thm.term_of ct val simp_t = the_default [([], Thm.reflexive ct)] (Termtab.lookup simp t) in if exists (fn info => simp_eq_better ctxt info (id, th)) simp_t then ctxt else ctxt |> Data.map (map_simp ( Termtab.update ( t, simp_t |> filter_out (simp_eq_better ctxt (id, th)) |> cons (id, th)))) end (* Get the entire rewrite info of t as stored directly in table. Return as a list of (id, th) pairs. *) fun get_rewrite_info ctxt ct = let val {simp, ...} = Data.get ctxt in case Termtab.lookup simp (Thm.term_of ct) of NONE => [([], Thm.reflexive ct)] | SOME infos => infos end (* Return the rewrite of t under a given id, as stored directly in table. Return t == t if not in table. *) fun get_rewrite id ctxt ct = case get_rewrite_info ctxt ct of [] => Thm.reflexive ct | ths' => ths' |> filter (BoxID.is_eq_descendent ctxt id o fst) |> map snd |> Util.max (rev_order o simp_ord) (*** Basic manipulation of subsimp and rep tables. ***) (* Given list of lists of simplifying infos indexed by term and then id, return simplifying infos indexed first by id and then term. *) fun merge_simp_infos ctxt lsts = let val ids = map (map fst) lsts |> BoxID.get_all_merges ctxt fun get_for_id_lst id lst = lst |> filter (BoxID.is_eq_descendent ctxt id o fst) |> map snd |> Util.max (rev_order o simp_ord) fun get_for_id id = (id, map (get_for_id_lst id) lsts) in map get_for_id ids end (* Get entire subterm rewrite info of t. Return as a list of (id, rewrite) pairs. *) fun get_subterm_rewrite_info ctxt ct = let val (cf, subs) = Drule.strip_comb ct in if null subs then [([], Thm.reflexive ct)] else (map (get_rewrite_info ctxt) subs) |> merge_simp_infos ctxt |> map (fn (id, equivs) => (id, Util.comb_equiv (cf, equivs))) |> Util.max_partial (simp_eq_better ctxt) end (* Given a head rep (id, th) (where th is reverse of a subterm rewrite), remove th from the reps table. *) fun remove_rep (id, th) ctxt = let val {reps, ...} = Data.get ctxt val t = Util.lhs_of th val reps_t = the_default [] (Termtab.lookup reps t) val _ = assert (member eq_info reps_t (id, th)) "remove_rep: not found" in ctxt |> Data.map (map_reps (Termtab.map_entry t (remove eq_info (id, th)))) end (* Update subterm rewrite of t under id to th: t == t'. Also update the reps table (reverse of subsimp table). *) fun update_subsimp (id, th) ctxt = let val {subsimp, ...} = Data.get ctxt val (t, t') = (Util.lhs_of th, Util.rhs_of th) val subsimp_t = the_default [] (Termtab.lookup subsimp t) in if exists (fn info => simp_eq_better ctxt info (id, th)) subsimp_t then ctxt else let val th' = meta_sym th val (rem, keep) = filter_split (simp_eq_better ctxt (id, th)) subsimp_t in ctxt |> Data.map (map_subsimp (Termtab.update (t, (id, th) :: keep))) |> Data.map (map_reps (Termtab.map_default (t', []) (cons (id, th')))) |> fold remove_rep (map (apsnd meta_sym) rem) end end (* Returns the list of (id, th) pairs, where th is t == t', such that t' subterm rewrites to t at box id. *) fun get_head_rep_info ctxt t = let val {reps, ...} = Data.get ctxt in case Termtab.lookup reps t of NONE => [] | SOME infos => infos end (* Assume t is subterm simplified under id. Return SOME (t == v) if v is a term in the rewrite table that is subterm rewrites to t, under the given id. If there is no such v, return NONE. *) fun get_head_rep id ctxt t = get_first (fn (id', t') => if BoxID.is_eq_ancestor ctxt id' id then SOME t' else NONE) (get_head_rep_info ctxt t) (* Returns head representations of the right side of th under id, or under more restrictive assumptions. Merge the equivalence theorems with th. Not guaranteed to be non-redundant. *) fun get_head_rep_with_id_th ctxt (id, th) = get_head_rep_info ctxt (Util.rhs_of th) |> map (BoxID.merge_eq_infos ctxt (id, th)) (* Using the subsimp table for subterm simplification. *) fun get_cached_subterm_rewrite_info ctxt t = let val {subsimp, ...} = Data.get ctxt in case Termtab.lookup subsimp t of NONE => [] | SOME infos => infos end (* Obtain the subterm simplification of t under id, or a less restrictive assumption. *) fun get_cached_subterm_rewrite id ctxt ct = case get_cached_subterm_rewrite_info ctxt (Thm.term_of ct) of [] => Thm.reflexive ct | ths' => ths' |> filter (BoxID.is_eq_descendent ctxt id o fst) |> map snd |> Util.max (rev_order o simp_ord) (*** Simplification of terms not indexed in table. ***) (* Assume t is subterm simplified under id. *) fun head_simplify id ctxt ct = case get_head_rep id ctxt (Thm.term_of ct) of NONE => Thm.reflexive ct | SOME th => Util.transitive_list [th, get_rewrite id ctxt (Thm.rhs_of th)] (* First simplify subterms, then simplify head. *) fun simplify id ctxt ct = if in_table_raw ctxt (Thm.term_of ct) then get_rewrite id ctxt ct else let val th = subterm_simplify id ctxt ct in Util.transitive_list [th, head_simplify id ctxt (Thm.rhs_of th)] end (* Simplify subterms, but not head. *) and subterm_simplify id ctxt ct = if in_table_raw ctxt (Thm.term_of ct) then get_cached_subterm_rewrite id ctxt ct else let val (cf, subs) = Drule.strip_comb ct val equivs = map (simplify id ctxt) subs in Util.comb_equiv (cf, equivs) end (* Convenient function for getting simplified value. *) fun simp_val id ctxt ct = Thm.rhs_of (simplify id ctxt ct) fun simp_val_t id ctxt t = Util.rhs_of (simplify id ctxt (Thm.cterm_of ctxt t)) (* Get all simplifications and subterm-simplifications. *) fun simplify_info ctxt ct = if in_table_raw ctxt (Thm.term_of ct) then get_rewrite_info ctxt ct else let val subterm_info = subterm_simplify_info ctxt ct fun get_rewrite_with_id_th (id', th') = (get_rewrite_info ctxt (Thm.rhs_of th')) |> map (BoxID.merge_eq_infos ctxt (id', th')) val head_info = subterm_info |> maps (get_head_rep_with_id_th ctxt) |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) |> maps get_rewrite_with_id_th in subterm_info @ head_info |> Util.max_partial (simp_eq_better ctxt) end (* Note similarity with get_subterm_rewrite_info. *) and subterm_simplify_info ctxt ct = if in_table_raw ctxt (Thm.term_of ct) then get_cached_subterm_rewrite_info ctxt (Thm.term_of ct) else let val (cf, subs) = Drule.strip_comb ct in if null subs then [([], Thm.reflexive ct)] else (map (simplify_info ctxt) subs) |> merge_simp_infos ctxt |> map (fn (id, equivs) => (id, Util.comb_equiv (cf, equivs))) |> Util.max_partial (simp_eq_better ctxt) end (* Exported equivalence function. *) fun is_equiv id ctxt (ct1, ct2) = ct1 aconvc ct2 orelse simp_val id ctxt ct1 aconvc simp_val id ctxt ct2 fun is_equiv_t id ctxt (t1, t2) = t1 aconv t2 orelse is_equiv id ctxt (Thm.cterm_of ctxt t1, Thm.cterm_of ctxt t2) (* No assumption on u. Get list of (id', u') pairs where id <= id', u' is in the table equivalent to u under id', and the head of u' is the same as that for t. If (id1, u1) and (id2, u2) are such that id2 implies id1 and all arguments of u1, u2 are equivalent under t2, then (id2, u2) is not included in the results. This function is much faster when u is in table (uses get_all_equiv directly). *) fun get_head_equiv ctxt cu = let val u = Thm.term_of cu val _ = assert (not (Term.is_open u)) "get_all_head_equivs: u is open" fun rep_eq_better (id1, th1) (id2, th2) = BoxID.is_eq_ancestor ctxt id1 id2 andalso is_equiv id2 ctxt (Thm.rhs_of th1, Thm.rhs_of th2) fun get_all_equiv_under_id (id', th') = (get_all_equiv ctxt (Util.rhs_of th')) |> map (BoxID.merge_eq_infos ctxt (id', th')) in if in_table_raw ctxt u then (* Get all equivs from table. *) get_all_equiv ctxt u else cu |> subterm_simplify_info ctxt |> maps (get_head_rep_with_id_th ctxt) |> Util.max_partial rep_eq_better |> maps get_all_equiv_under_id |> cons ([], Thm.reflexive cu) |> Util.max_partial (equiv_eq_better ctxt) end (* Get list of head equivs for cu, whose head agrees (or matches) with t. *) fun get_head_equiv_with_t ctxt (id, cu) t = let fun is_valid (_, th) = if Term.is_Var t then true else if Term.is_Var (Term.head_of t) then true else Term.aconv_untyped (Term.head_of (Util.rhs_of th), Term.head_of t) in (get_head_equiv ctxt cu) |> filter is_valid |> BoxID.merge_box_with_info ctxt id end (*** Adding rewrite rules and maintain invariants of the table. ***) (* Work out all consequences of updating simps in the rewrite table. This includes updating simp of neighboring nodes, and updating simps / subsimps of nodes for which the current node is a subterm. inits is a list of simps to be added, where each element is of the form (id, t == t') updating simp of t. In this function we maintain a list of simplifications (to_process) that have been added to the table, but whose consequences (on neighboring terms and containing terms) need to be processed. *) fun process_update_simp inits ctxt = let fun eq_update ((id, th), (id', th')) = (id = id' andalso Thm.prop_of th aconv Thm.prop_of th') (* Add the simp (id, th) for the left side of th to the table. If (id, th) is not redundant, add (id, th) to the list of simps whose consequences need to be processed. *) fun process_simp (id, th) (to_process, ctxt) = if simp_ord (th, get_rewrite id ctxt (Thm.lhs_of th)) = LESS then (insert eq_update (id, th) to_process, update_simp (id, th) ctxt) else (to_process, ctxt) (* Recompute subterm simplifications of t. Update the subsimp table as well as the simp table (in the latter case possibly adding to the to_process list. *) fun process_term ct (to_process, ctxt) = let val subsimps = get_subterm_rewrite_info ctxt ct val ctxt' = fold update_subsimp subsimps ctxt in fold process_simp subsimps (to_process, ctxt') end (* Pull the first item of to_process and work out its consequences. *) fun update_step (to_process, ctxt) = case to_process of [] => ([], ctxt) | (id, th) :: rest => let val t = Util.lhs_of th (* th: t == simp_t, th': t = t', result: t' = simp_t. *) fun process_neigh (id', th') = BoxID.merge_eq_infos ctxt (id', meta_sym th') (id, th) val new_simp = map process_neigh (equiv_neighs ctxt t) in (rest, ctxt) |> fold process_simp new_simp |> fold process_term (immediate_contains ctxt t) |> update_step end in ([], ctxt) |> fold process_simp inits |> update_step |> snd end (* Return list of ids that are descendents of id at which t1 and t2 are equiv. *) fun equiv_info ctxt id (ct1, ct2) = let val simp1 = simplify_info ctxt ct1 val simp2 = simplify_info ctxt ct2 fun compare (id1, th1) (id2, th2) = if Util.rhs_of th1 aconv Util.rhs_of th2 then [(BoxID.merge_boxes ctxt (id1, id2), Util.transitive_list [th1, meta_sym th2])] else [] in (maps (fn s1 => maps (compare s1) simp2) simp1) |> BoxID.merge_box_with_info ctxt id |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) end fun equiv_info_t ctxt id (t1, t2) = equiv_info ctxt id (Thm.cterm_of ctxt t1, Thm.cterm_of ctxt t2) (* Return list of ids that are descendents of id at which t1 and t2 are subterm equiv. *) fun subequiv_info ctxt id (ct1, ct2) = if not ((Term.head_of (Thm.term_of ct1)) aconv Term.head_of (Thm.term_of ct2)) then [] else let val simp1 = subterm_simplify_info ctxt ct1 val simp2 = subterm_simplify_info ctxt ct2 fun compare (id1, th1) (id2, th2) = if Util.rhs_of th1 aconv Util.rhs_of th2 then [(BoxID.merge_boxes ctxt (id1, id2), Util.transitive_list [th1, meta_sym th2])] else [] in (maps (fn s1 => maps (compare s1) simp2) simp1) |> BoxID.merge_box_with_info ctxt id |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) end (* Get list of terms reachable from the terms ts. If trav_contains = false, only travel along equiv edges. Otherwise also travel from t to any term containing t. *) fun get_reachable_terms trav_contains ctxt ts = let fun helper (new, all) = case new of [] => ([], all) | t :: rest => let val contains = if trav_contains then map Thm.term_of (immediate_contains ctxt t) else [] val neighs = (equiv_neighs ctxt t) |> map (fn (_, th) => Util.rhs_of th) val new' = (contains @ neighs) |> distinct (op aconv) |> subtract (op aconv) all val all' = new' @ all in helper (rest @ new', all') end in snd (helper (ts, ts)) end fun add_all_equiv_raw (cur_info as (_, th)) ctxt = let val {all_equiv, ...} = Data.get ctxt val t = Util.lhs_of th val equiv_t = the_default [] (Termtab.lookup all_equiv t) in if exists (fn info => equiv_eq_better ctxt info cur_info) equiv_t then ctxt else ctxt |> Data.map (map_all_equiv ( Termtab.update ( t, equiv_t |> filter_out (equiv_eq_better ctxt cur_info) |> cons cur_info))) end fun update_all_equiv id th ctxt = let val (ct1, ct2) = (Thm.lhs_of th, Thm.rhs_of th) val t1_equiv = get_all_equiv ctxt (Thm.term_of ct1) val t2_equiv = get_all_equiv ctxt (Thm.term_of ct2) (* eq1 is t1 == t1', eq2 is t2 == t2', add t1' == t2' and t2' == t1' to all_equiv. *) fun process_pair ((id1, eq1), (id2, eq2)) ctxt = let val id' = BoxID.merge_boxes ctxt (id, BoxID.merge_boxes ctxt (id1, id2)) val eq = Util.transitive_list [meta_sym eq1, th, eq2] in ctxt |> add_all_equiv_raw (id', eq) |> add_all_equiv_raw (id', meta_sym eq) end in fold process_pair (Util.all_pairs (t1_equiv, t2_equiv)) ctxt end (* Assume t1 and t2 are already in table. Add equiv edge th: t1 == t2 and work out all the consequences. *) fun add_rewrite_raw id th ctxt = let val (ct1, ct2) = (Thm.lhs_of th, Thm.rhs_of th) (* New simplifications. *) val t1_rinfo = get_rewrite_info ctxt ct1 val t2_rinfo = get_rewrite_info ctxt ct2 val t1_news = t2_rinfo |> map (BoxID.merge_eq_infos ctxt (id, th)) val t2_news = t1_rinfo |> map (BoxID.merge_eq_infos ctxt (id, meta_sym th)) in ctxt |> add_equiv (id, th) |> process_update_simp (t1_news @ t2_news) |> update_all_equiv id th end (* Add equiv edges so the rewrite table is consistent. That is, any two nodes that have the same simp should be connected in the equiv graph. Also return the list of equiv edges added. *) fun complete_table ctxt = let fun find_new_equiv ct = let val equivs = get_all_equiv ctxt (Thm.term_of ct) in (get_subterm_rewrite_info ctxt ct) |> maps (get_head_rep_with_id_th ctxt) |> Util.max_partial (equiv_eq_better ctxt) |> filter_out (has_equiv_eq_better ctxt equivs) end val new_equivs = (maps find_new_equiv (get_all_terms ctxt)) |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) in case new_equivs of [] => ([], ctxt) | (id, th) :: _ => (* Add bi-directional edges for terms with the same subterm simplification. Keep track of list of edges added. *) let val ctxt' = ctxt |> add_rewrite_raw id th val (edges, ctxt'') = complete_table ctxt' in ((id, th) :: edges, ctxt'') end end (* Add term t to rewrite table at box id. The primed version (used in add_rewrite_thm) add t and all its subterms to the table. The unprimed version also add new equiv edges if necessary, resulting is a consistent rewrite table. *) fun add_term' (id, ct) ctxt = if in_table_raw_for_id ctxt (id, Thm.term_of ct) then ctxt else let (* First add subterms. *) val t = Thm.term_of ct val imm_subterms = UtilLogic.list_subterms t val id_subs = map (fn t => (id, Thm.cterm_of ctxt t)) imm_subterms val ctxt' = ctxt |> fold add_term' id_subs |> add_term_raw (id, ct) in (* If t is already in table (just not at id), then we are done. Otherwise compute all indexed info for t. *) if in_table_raw ctxt t then ctxt' else let (* Compute simplification and properties of t. Note we are doing this in box [] (not id), so it does not need to be recomputed if a term is added to a more general box. *) val simps = get_subterm_rewrite_info ctxt' ct in ctxt' |> fold (add_contain ct) imm_subterms |> fold update_simp simps |> fold update_subsimp simps |> add_all_equiv_raw ([], Thm.reflexive ct) end end fun add_term (id, ct) ctxt = if in_table_raw_for_id ctxt (id, Thm.term_of ct) then ([], ctxt) else ctxt |> add_term' (id, ct) |> complete_table fun add_term_list term_infos ctxt = case term_infos of [] => ([], ctxt) | (id, ct) :: rest => let val (edges, ctxt') = add_term (id, ct) ctxt val (edges', ctxt'') = add_term_list rest ctxt' in (edges @ edges', ctxt'') end (* First make sure t1 and t2 are in table. Add t1 == t2 to the table, and work out any consequences. The result is a consistent rewrite table. *) fun add_rewrite (id, eq_th) ctxt = let (* eq_th is meta-equality when between lambda terms. *) val meta_eq = if Util.is_meta_eq (Thm.prop_of eq_th) then eq_th - else to_meta_eq eq_th + else UtilLogic.to_meta_eq eq_th val (ct1, ct2) = meta_eq |> Thm.cprop_of |> Thm.dest_equals in if is_equiv id ctxt (ct1, ct2) then ([], ctxt) else ctxt |> add_term' (id, ct1) |> add_term' (id, ct2) |> add_rewrite_raw id meta_eq |> complete_table |> apfst (cons (id, meta_eq)) end (* Returns the list of terms in ctxt' that is not in ctxt. *) fun get_new_terms (ctxt, ctxt') = let fun get_for_term ct = (in_table_raw_ids ctxt' (Thm.term_of ct)) |> filter_out (fn id => in_table_raw_for_id ctxt (id, Thm.term_of ct)) |> map (rpair ct) in maps get_for_term (get_all_terms ctxt') end end (* structure RewriteTable. *) diff --git a/thys/Auto2_HOL/status.ML b/thys/Auto2_HOL/status.ML --- a/thys/Auto2_HOL/status.ML +++ b/thys/Auto2_HOL/status.ML @@ -1,402 +1,411 @@ (* File: status.ML Author: Bohua Zhan Definition of updates and proof status. *) datatype raw_update = AddItems of {id: box_id, sc: int option, raw_items: raw_item list} | AddBoxes of {id: box_id, sc: int option, init_assum: term} | ResolveBox of {id: box_id, th: thm} | ShadowItem of {id: box_id, item: box_item} type update = {sc: int, prfstep_name: string, source: box_item list, raw_updt: raw_update} signature UPDATE = sig val target_of_update: raw_update -> box_id val replace_id_of_update: raw_update -> raw_update val string_of_raw_update: Proof.context -> raw_update -> string val thm_to_ritem: thm -> raw_item val thm_update: box_id * thm -> raw_update val thm_update_sc: int -> box_id * thm -> raw_update val apply_exists_ritems: Proof.context -> thm -> raw_item list * thm val update_info: Proof.context -> box_id -> raw_item list -> string val source_info: update -> string end; -structure Update : UPDATE = +functor Update( + structure BoxItem: BOXITEM; + structure ItemIO: ITEM_IO; + structure UtilLogic: UTIL_LOGIC; + ): UPDATE = struct fun target_of_update raw_updt = case raw_updt of AddItems {id, ...} => id | AddBoxes {id, ...} => id | ResolveBox {id, ...} => id | ShadowItem {id, ...} => id fun replace_id_of_update raw_updt = let val id = target_of_update raw_updt val _ = assert (BoxID.has_incr_id id) "replace_id_of_update" val id' = BoxID.replace_incr_id id in case raw_updt of AddItems {sc, raw_items, ...} => AddItems {id = id', sc = sc, raw_items = raw_items} | AddBoxes {sc, init_assum, ...} => AddBoxes {id = id', sc = sc, init_assum = init_assum} | ResolveBox {th, ...} => ResolveBox {id = id', th = th} | ShadowItem {item, ...} => ShadowItem {id = id', item = item} end (* Debugging output for raw update. *) fun string_of_raw_update ctxt raw_updt = case raw_updt of AddItems {id, raw_items, ...} => "Add items " ^ (Util.string_of_list ( ItemIO.string_of_raw_item ctxt) raw_items) ^ " to box " ^ (BoxID.string_of_box_id id) | AddBoxes {id, init_assum, ...} => "Add box " ^ (Syntax.string_of_term ctxt init_assum) ^ " under box " ^ (BoxID.string_of_box_id id) | ResolveBox {id, ...} => "Resolve box " ^ (BoxID.string_of_box_id id) | ShadowItem {id, item} => "Shadow item " ^ (ItemIO.string_of_item ctxt item) ^ " in box " ^ (BoxID.string_of_box_id id) (* Create raw_item from theorem. *) fun thm_to_ritem th = let val prop = Thm.prop_of th - val _ = assert (is_Trueprop prop) "thm_update: prop is not Trueprop." + val _ = assert (UtilLogic.is_Trueprop prop) "thm_update: prop is not Trueprop." val _ = assert (not (Util.has_vars prop)) "thm_to_ritem: prop contains schematic variables." in - Fact (TY_PROP, [dest_Trueprop prop], th) + Fact (TY_PROP, [UtilLogic.dest_Trueprop prop], th) end (* Create raw_update from id and theorem. *) fun thm_update (id, th) = - if Thm.prop_of th aconv pFalse then + if Thm.prop_of th aconv UtilLogic.pFalse then ResolveBox {id = id, th = th} else AddItems {id = id, sc = NONE, raw_items = [thm_to_ritem th]} fun thm_update_sc sc (id, th) = - if Thm.prop_of th aconv pFalse then + if Thm.prop_of th aconv UtilLogic.pFalse then ResolveBox {id = id, th = th} else AddItems {id = id, sc = SOME sc, raw_items = [thm_to_ritem th]} (* Apply the given existence theorem. Return a pair (ritems, th), ritems contain the new variables and the new handler, and th is the instantiated property of the new variables, which can either be processed further, or made into an ritem using thm_update. *) fun apply_exists_ritems ctxt ex_th = let - val (vars, body) = (UtilLogic.strip_exists (prop_of' ex_th)) + val (vars, body) = (UtilLogic.strip_exists (UtilLogic.prop_of' ex_th)) |> Util.to_internal_vars ctxt - |> apsnd mk_Trueprop + |> apsnd UtilLogic.mk_Trueprop in if null vars then ([], ex_th) else (map (BoxItem.var_to_fact) vars @ [Handler (vars, body, ex_th)], Thm.assume (Thm.cterm_of ctxt body)) end (* Print a list of raw items. *) fun update_info ctxt id ritems = (Util.string_of_list' (ItemIO.string_of_raw_item ctxt) (filter (not o BoxItem.is_handler_raw) ritems)) ^ " at box " ^ (BoxID.string_of_box_id id) (* Print the source of the update. *) fun source_info {prfstep_name, source, ...} = prfstep_name ^ " on " ^ (Util.string_of_list' (fn {uid, ...} => "(" ^ string_of_int uid ^ ")") source) end (* structure Update. *) structure Updates_Heap = Heap ( type elem = update fun ord ({sc = sc1, ...}, {sc = sc2, ...}) = int_ord (sc1, sc2) ) type status = { assums: term list, handlers: (box_id * (term list * term * thm)) list, items: (box_item * box_id list) Inttab.table, queue: Updates_Heap.T, resolve_th: thm option, ctxt: Proof.context } signature STATUS = sig val empty_status: Proof.context -> status val map_context: (Proof.context -> Proof.context) -> status -> status val add_handler: box_id * (term list * term * thm) -> status -> status val get_handlers: status -> (box_id * (term list * term * thm)) list val add_item: box_item -> status -> status val get_items: status -> box_item list val get_all_items_at_id: status -> box_id -> box_item list val get_num_items: status -> int val set_resolve_th: thm -> status -> status val get_resolve_th: status -> thm val add_prim_box: box_id -> term -> status -> int * status val add_resolved: box_id -> status -> status val get_init_assums: status -> box_id -> term list val get_init_assum: status -> int -> term val lookup_item: status -> int -> (box_item * box_id list) option val clear_incr: status -> status val add_shadowed: box_id * box_item -> status -> status val query_shadowed: status -> box_id -> box_item -> bool val query_removed: status -> box_item -> bool val find_fact: status -> box_id -> term -> thm option val find_ritem_exact: status -> box_id -> raw_item -> bool val invoke_handler: Proof.context -> term list * term * thm -> thm -> thm val get_on_resolve: status -> box_id -> int -> thm -> thm val find_prim_box: status -> box_id -> term -> int option val map_queue: (Updates_Heap.T -> Updates_Heap.T) -> status -> status val add_to_queue: update -> status -> status val delmin_from_queue: status -> status end; -structure Status : STATUS = +functor Status( + structure BoxItem: BOXITEM; + structure ItemIO: ITEM_IO; + structure UtilLogic: UTIL_LOGIC + structure PropertyData: PROPERTY_DATA; + structure RewriteTable: REWRITE_TABLE; + structure WellformData: WELLFORM_DATA; + ): STATUS = struct fun empty_status ctxt = {assums = [UtilLogic.term_of_bool true], handlers = [], items = Inttab.empty, queue = Updates_Heap.empty, ctxt = ctxt, resolve_th = NONE} fun map_context f {assums, handlers, items, queue, ctxt, resolve_th} = {assums = assums, handlers = handlers, items = items, queue = queue, ctxt = f ctxt, resolve_th = resolve_th} fun add_handler (id, handler) {assums, handlers, items, queue, ctxt, resolve_th} = {assums = assums, handlers = (id, handler) :: handlers, items = items, queue = queue, ctxt = ctxt, resolve_th = resolve_th} fun get_handlers {handlers, ...} = handlers fun map_items f {assums, handlers, items, queue, ctxt, resolve_th} = {assums = assums, handlers = handlers, items = f items, queue = queue, ctxt = ctxt, resolve_th = resolve_th} fun add_item (item as {uid, ...}) st = st |> map_items (Inttab.update_new (uid, (item, []))) fun get_items {items, ...} = items |> Inttab.dest |> map snd |> filter_out (fn ({id, ...}, shadow_ids) => member (op =) shadow_ids id) |> map fst fun get_items_at_id (st as {ctxt, ...}) id = st |> get_items |> filter (fn {id = id', ...} => BoxID.is_eq_ancestor ctxt id' id) fun get_all_items_at_id {items, ctxt, ...} id = items |> Inttab.dest |> map snd |> map fst |> filter (fn {id = id', ...} => BoxID.is_eq_ancestor ctxt id' id) fun get_num_items {items, ...} = length (Inttab.dest items) fun set_resolve_th th {assums, handlers, items, queue, ctxt, ...} = {assums = assums, handlers = handlers, items = items, queue = queue, ctxt = ctxt, resolve_th = SOME th} fun get_resolve_th {resolve_th, ...} = case resolve_th of NONE => raise Fail "get_resolve_th: not resolved." | SOME th => th fun add_prim_box parent_id assum {assums, handlers, items, queue, ctxt, resolve_th} = let val (id, ctxt') = BoxID.add_prim_id parent_id ctxt in (id, {assums = assums @ [assum], handlers = handlers, items = items, queue = queue, ctxt = ctxt', resolve_th = resolve_th}) end fun add_resolved id {assums, handlers, items, queue, ctxt, resolve_th} = let val ctxt' = BoxID.add_resolved id ctxt in {assums = assums, handlers = handlers, items = items, queue = queue, ctxt = ctxt' |> RewriteTable.clean_resolved id |> PropertyData.clean_resolved id |> WellformData.clean_resolved id, resolve_th = resolve_th} end fun get_init_assums {assums, ctxt, ...} id = let val parents = BoxID.get_ancestors_prim ctxt id in map (nth assums) parents end fun get_init_assum {assums, ...} prim_id = nth assums prim_id fun lookup_item {items, ...} uid = Inttab.lookup items uid fun clear_incr st = st |> map_items ( Inttab.map (fn _ => fn (item, ids) => (BoxItem.item_replace_incr item, ids))) fun add_shadowed (shadow_id, {uid, ...}) (st as {ctxt, ...}) = case lookup_item st uid of NONE => raise Fail "add_shadowed: item not found." | SOME (item, shadow_ids) => let val shadow_ids' = (shadow_id :: shadow_ids) |> Util.max_partial (BoxID.is_eq_ancestor ctxt) in st |> map_items (Inttab.update (uid, (item, shadow_ids'))) end fun query_shadowed (st as {ctxt, ...}) shadow_id {uid, ...} = case lookup_item st uid of NONE => raise Fail "query_shadowed: item not found." | SOME (_, shadow_ids) => exists (BoxID.is_eq_descendent ctxt shadow_id) shadow_ids fun query_removed st (item as {id, ...}) = query_shadowed st id item (* Try to find fact at id or above with the proposition t. Return SOME th if found. *) fun find_fact (st as {ctxt, ...}) id t = let val ct = Thm.cterm_of ctxt t val items = get_items_at_id st id val res = (WellformData.find_fact ctxt items (id, ct)) |> filter (fn (id', _) => id' = id) in case res of [] => NONE | (_, th) :: _ => SOME th end (* Find item with the exact ty_str and tname, whose id is an eq-ancestor of the given id. There are two special cases: if the ritem in question is of type EQ and PROPERTY, in which case we try to find it in the rewrite table. *) fun find_ritem_exact (st as {ctxt, ...}) id ritem = case ritem of Handler _ => false | Fact (ty_str, tname, _) => if ty_str = TY_EQ then let val (lhs, rhs) = the_pair (BoxItem.get_tname_raw ritem) in RewriteTable.is_equiv_t id ctxt (lhs, rhs) end else if ty_str = TY_PROPERTY then let val prop = the_single tname val infos = (PropertyData.get_property_t ctxt (id, prop)) |> filter (fn (id', _) => BoxID.is_eq_ancestor ctxt id' id) in not (null infos) end else let val thy = Proof_Context.theory_of ctxt val {shadow_fn, ...} = ItemIO.get_io_info thy ty_str fun eq_item {ty_str = ty_str2, tname = tname2, ...} = if ty_str <> ty_str2 then false else case shadow_fn of NONE => eq_list (op aconv) (tname, map Thm.term_of tname2) | SOME (f, _) => f ctxt id (tname, tname2) in exists eq_item (get_all_items_at_id st id) end (* Invoke a single handler (vars, t, ex_th) on the theorem th. *) fun invoke_handler ctxt (vars, t, ex_th) th = if member (op aconv) (Thm.hyps_of th) t then th |> Thm.implies_intr (Thm.cterm_of ctxt t) |> fold (UtilLogic.ex_elim ctxt) (rev vars) |> Thm.elim_implies ex_th else th (* Derive the consequence if box id is resolved, to the parent id formed by getting parent at index i. *) fun get_on_resolve (st as {ctxt, ...}) id i th = let (* First get list of handlers to invoke. *) val prim_id = nth id i val handlers = st |> get_handlers |> filter (fn (id', _) => BoxID.is_eq_ancestor ctxt [prim_id] id' andalso BoxID.is_eq_descendent ctxt id id') |> map snd val init_assum = get_init_assum st prim_id val th' = th |> fold (invoke_handler ctxt) handlers |> Thm.implies_intr (Thm.cterm_of ctxt init_assum) |> apply_to_thm UtilLogic.rewrite_from_contra_form val _ = assert (prim_id <> BoxID.home_id orelse null (Thm.hyps_of th')) "get_on_resolve: did not remove all hypothesis at box 0." in th' end (* Find a primitive box (if there is any) whose initial facts agree exactly with the given initial facts. Note we cannot yet handle new variables. *) fun find_prim_box {assums, ctxt, ...} id init_assum = let fun agree_at_id prim_id' = let val parent = BoxID.get_parent_prim ctxt prim_id' val is_equiv_t = RewriteTable.is_equiv_t parent ctxt in is_equiv_t (init_assum, nth assums prim_id') end fun can_test_id prim_id' = BoxID.is_eq_ancestor ctxt (BoxID.get_parent_prim ctxt prim_id') id val ids_to_test = filter can_test_id (0 upto (length assums - 1)) in find_first agree_at_id ids_to_test end fun map_queue f {assums, handlers, items, queue, ctxt, resolve_th} = {assums = assums, handlers = handlers, items = items, queue = f queue, ctxt = ctxt, resolve_th = resolve_th} fun add_to_queue updt = map_queue (Updates_Heap.insert updt) val delmin_from_queue = map_queue Updates_Heap.delete_min end (* structure Status *) - -val _ = Theory.setup (ItemIO.add_basic_item_io) diff --git a/thys/Auto2_HOL/util_logic.ML b/thys/Auto2_HOL/util_logic.ML --- a/thys/Auto2_HOL/util_logic.ML +++ b/thys/Auto2_HOL/util_logic.ML @@ -1,628 +1,625 @@ (* File: util_logic.ML Author: Bohua Zhan Utility functions, after fixing object logic. *) signature BASIC_UTIL_LOGIC = sig (* Terms *) val Trueprop: term val is_Trueprop: term -> bool val mk_Trueprop: term -> term val dest_Trueprop: term -> term val Trueprop_conv: conv -> conv val pFalse: term val Not: term val mk_not: term -> term val dest_not: term -> term val is_neg: term -> bool val get_neg: term -> term val get_neg': term -> term val conj: term val is_conj: term -> bool val mk_conj: term * term -> term val strip_conj: term -> term list val list_conj: term list -> term val disj: term val is_disj: term -> bool val mk_disj: term * term -> term val strip_disj: term -> term list val list_disj: term list -> term val imp: term val is_imp: term -> bool val mk_imp: term * term -> term val dest_imp: term -> term * term val strip_obj_imp: term -> term list * term val list_obj_imp: term list * term -> term val is_obj_all: term -> bool val is_ball: term -> bool val mk_obj_all: term -> term -> term val is_ex: term -> bool val is_bex: term -> bool val mk_exists: term -> term -> term val is_mem: term -> bool val mk_mem: term * term -> term (* Theorems *) val is_true_th: thm -> bool val prop_of': thm -> term val cprop_of': thm -> cterm val concl_of': thm -> term val make_trueprop_eq: thm -> thm val assume_eq: theory -> term * term -> thm val apply_to_thm': conv -> thm -> thm val to_meta_eq: thm -> thm val to_obj_eq: thm -> thm val obj_sym: thm -> thm val rewr_obj_eq: thm -> conv val conj_left_th: thm -> thm val conj_right_th: thm -> thm val equiv_forward_th: thm -> thm val equiv_backward_th: thm -> thm val inv_backward_th: thm -> thm val to_obj_eq_th: thm -> thm val to_obj_eq_iff_th: thm -> thm val obj_sym_th: thm -> thm end; signature UTIL_LOGIC = sig include BASIC_UTIL_LOGIC (* Terms *) val term_of_bool: bool -> term val bool_of_term: term -> bool (* Finding subterms *) val list_subterms: term -> term list val get_all_subterms: term -> term list val get_all_subterms_skip_if: term -> term list (* cterms *) val get_cneg: cterm -> cterm (* Theorems *) val make_neg_eq: thm -> thm val rewrite_to_contra_form: conv val rewrite_from_contra_form: conv val to_obj_conv: Proof.context -> conv val to_obj_conv_on_horn: Proof.context -> conv val to_meta_conv: Proof.context -> conv val split_conj_th: thm -> thm list val split_conj_gen_th: thm -> thm list val split_not_disj_th: thm -> thm list val strip_horn': thm -> term list * term val mk_conjs_th: thm list -> thm val ex_elim: Proof.context -> term -> thm -> thm val force_abs_form: term -> term val strip_obj_horn: term -> term list * (term list * term) val list_obj_horn: term list * (term list * term) -> term val is_ex_form_gen: term -> bool val normalize_exists: Proof.context -> conv val strip_exists: term -> term list * term (* Wrapper around common tactics. *) val prove_by_tac: (Proof.context -> int -> tactic) -> Proof.context -> thm list -> term -> thm val contra_by_tac: (Proof.context -> int -> tactic) -> Proof.context -> thm list -> thm end; -structure UtilLogic : UTIL_LOGIC = +functor UtilLogic(UtilBase: UTIL_BASE) : UTIL_LOGIC = struct (* Booleans *) -fun term_of_bool b = (if b then bTrue else bFalse) +fun term_of_bool b = (if b then UtilBase.bTrue else UtilBase.bFalse) fun bool_of_term t = - if t aconv bTrue then true - else if t aconv bFalse then false + if t aconv UtilBase.bTrue then true + else if t aconv UtilBase.bFalse then false else raise Fail "bool_of_term: unexpected t." (* Trueprop *) -val Trueprop = Const (UtilBase.Trueprop_name, boolT --> propT) +val Trueprop = Const (UtilBase.Trueprop_name, UtilBase.boolT --> propT) (* Returns whether the given term is Trueprop. *) fun is_Trueprop t = let val _ = assert (fastype_of t = propT) "is_Trueprop: wrong type" in case t of Const (c, _) $ _ => c = UtilBase.Trueprop_name | _ => false end fun mk_Trueprop P = Trueprop $ P fun dest_Trueprop t = if is_Trueprop t then dest_arg t else raise Fail "dest_Trueprop" fun Trueprop_conv cv ct = if is_Trueprop (Thm.term_of ct) then Conv.arg_conv cv ct else raise CTERM ("Trueprop_conv", [ct]) -val pFalse = Trueprop $ bFalse +val pFalse = Trueprop $ UtilBase.bFalse (* Not *) -val Not = Const (UtilBase.Not_name, boolT --> boolT) +val Not = Const (UtilBase.Not_name, UtilBase.boolT --> UtilBase.boolT) fun mk_not P = Not $ P (* Returns whether the given term is in neg form. *) fun is_neg t = let - val _ = assert (fastype_of t = boolT) "is_neg: wrong type" + val _ = assert (fastype_of t = UtilBase.boolT) "is_neg: wrong type" in case t of Const (c, _) $ _ => c = UtilBase.Not_name | _ => false end fun dest_not t = if is_neg t then dest_arg t else raise Fail "dest_not" (* Returns the negation of the given term. Avoids double negatives. *) fun get_neg t = if is_neg t then dest_not t else Not $ t (* Version of get_neg for Trueprop terms. *) fun get_neg' t = let val _ = assert (is_Trueprop t) "get_neg': input should be a Trueprop." in t |> dest_Trueprop |> get_neg |> mk_Trueprop end (* Conjunction and disjunction *) -val conj = Const (UtilBase.Conj_name, boolT --> boolT --> boolT) +val conj = Const (UtilBase.Conj_name, UtilBase.boolT --> UtilBase.boolT --> UtilBase.boolT) fun is_conj t = case t of Const (c, _) $ _ $ _ => c = UtilBase.Conj_name | _ => false fun mk_conj (t, u) = conj $ t $ u fun strip_conj t = if is_conj t then (dest_arg1 t) :: strip_conj (dest_arg t) else [t] fun list_conj ts = case ts of [] => error "list_conj" | [t] => t | t :: rest => mk_conj (t, list_conj rest) -val disj = Const (UtilBase.Disj_name, boolT --> boolT --> boolT) +val disj = Const (UtilBase.Disj_name, UtilBase.boolT --> UtilBase.boolT --> UtilBase.boolT) fun is_disj t = case t of Const (c, _) $ _ $ _ => c = UtilBase.Disj_name | _ => false fun mk_disj (t, u) = disj $ t $ u fun strip_disj t = if is_disj t then (dest_arg1 t) :: strip_disj (dest_arg t) else [t] fun list_disj ts = case ts of [] => raise Fail "list_disj" | [t] => t | t :: ts' => mk_disj (t, list_disj ts') (* Object implication *) -val imp = Const (UtilBase.Imp_name, boolT --> boolT --> boolT) +val imp = Const (UtilBase.Imp_name, UtilBase.boolT --> UtilBase.boolT --> UtilBase.boolT) fun is_imp t = case t of Const (c, _) $ _ $ _ => c = UtilBase.Imp_name | _ => false fun mk_imp (t, u) = imp $ t $ u fun dest_imp t = if is_imp t then (dest_arg1 t, dest_arg t) else raise Fail "dest_imp" (* Given t of form A1 --> ... --> An, return ([A1, ..., A(n-1)], An). *) fun strip_obj_imp t = if is_imp t then let val (As, B') = strip_obj_imp (dest_arg t) in (dest_arg1 t :: As, B') end else ([], t) fun list_obj_imp (As, C) = case As of [] => C | A :: rest => mk_imp (A, list_obj_imp (rest, C)) fun is_obj_all t = case t of Const (c, _) $ Abs _ => c = UtilBase.All_name | _ => false fun is_ball t = case t of Const (c, _) $ _ $ _ => c = UtilBase.Ball_name | _ => false fun mk_obj_all t body = let val (x, T) = Term.dest_Free t in - Const (UtilBase.All_name, (T --> boolT) --> boolT) $ Term.absfree (x, T) body + Const (UtilBase.All_name, (T --> UtilBase.boolT) --> UtilBase.boolT) $ Term.absfree (x, T) body end fun is_ex t = case t of Const (c, _) $ Abs _ => c = UtilBase.Ex_name | _ => false fun is_bex t = case t of Const (c, _) $ _ $ _ => c = UtilBase.Bex_name | _ => false fun mk_exists t body = let val (x, T) = Term.dest_Free t in - Const (UtilBase.Ex_name, (T --> boolT) --> boolT) $ Term.absfree (x, T) body + Const (UtilBase.Ex_name, (T --> UtilBase.boolT) --> UtilBase.boolT) $ Term.absfree (x, T) body end fun is_mem t = case t of Const (c, _) $ _ $ _ => c = UtilBase.Mem_name | _ => false fun mk_mem (x, A) = let val T = fastype_of x in - Const (UtilBase.Mem_name, T --> UtilBase.mk_setT T --> boolT) $ x $ A + Const (UtilBase.Mem_name, T --> UtilBase.mk_setT T --> UtilBase.boolT) $ x $ A end (* Finding subterms *) (* Get all subterms of t. *) fun list_subterms t = let fun dest_at_head t = case t of Abs (_, _, body) => dest_at_head body | _ => if Term.is_open t then t |> Term.strip_comb |> snd |> maps dest_at_head else [t] and dest t = case t of Abs _ => dest_at_head t | _ $ _ => t |> Term.strip_comb |> snd | _ => [] in dest t end (* List all (closed) subterms. Larger ones will be listed first. *) fun get_all_subterms t = t |> list_subterms |> maps get_all_subterms |> distinct (op aconv) |> cons t (* List all (closed) subterms. For terms of form if cond then yes or no, list only subterms of cond. *) fun get_all_subterms_skip_if t = if UtilBase.is_if t then t |> Util.dest_args |> hd |> get_all_subterms_skip_if |> cons t else t |> list_subterms |> maps get_all_subterms_skip_if |> distinct (op aconv) |> cons t (* cterms *) fun get_cneg ct = let val t = Thm.term_of ct - val _ = assert (fastype_of t = boolT) "get_neg: wrong type" + val _ = assert (fastype_of t = UtilBase.boolT) "get_neg: wrong type" in if is_neg t then Thm.dest_arg ct else Thm.apply UtilBase.cNot ct end (* Theorems *) -fun is_true_th th = Thm.eq_thm_prop (th, true_th) +fun is_true_th th = Thm.eq_thm_prop (th, UtilBase.true_th) fun prop_of' th = dest_Trueprop (Thm.prop_of th) fun cprop_of' th = Thm.dest_arg (Thm.cprop_of th) fun concl_of' th = dest_Trueprop (Thm.concl_of th) (* Given an equality between bools, make it an equality between props, by applying the function Trueprop to both sides. *) fun make_trueprop_eq th = Thm.combination (Thm.reflexive UtilBase.cTrueprop) th (* Assumed theorems. *) fun assume_eq thy (t1, t2) = - Thm.assume (Thm.global_cterm_of thy (mk_Trueprop (mk_eq (t1, t2)))) + Thm.assume (Thm.global_cterm_of thy (mk_Trueprop (UtilBase.mk_eq (t1, t2)))) (* Apply cv to the statement of th, skipping Trueprop. *) fun apply_to_thm' cv th = apply_to_thm (Trueprop_conv cv) th val to_meta_eq = apply_to_thm (Util.concl_conv UtilBase.to_meta_eq_cv) val to_obj_eq = apply_to_thm (Util.concl_conv UtilBase.to_obj_eq_cv) val obj_sym = apply_to_thm (Util.concl_conv UtilBase.obj_sym_cv) (* Obtain rewriting conv from obj equality. *) fun rewr_obj_eq eq_th = Conv.rewr_conv (to_meta_eq eq_th) (* Given an equality A == B, make the equality ~A == ~B. Cancel ~~ on both sides if exists. *) fun make_neg_eq th = th |> Thm.combination (Thm.reflexive UtilBase.cNot) |> apply_to_lhs (Conv.try_conv (rewr_obj_eq UtilBase.nn_cancel_th)) |> apply_to_rhs (Conv.try_conv (rewr_obj_eq UtilBase.nn_cancel_th)) (* If ct is of the form [...] ==> False, leave it unchanged. Otherwise, change [...] ==> B to [..., ~ B] ==> False and change [...] ==> ~ B to [..., B] ==> False. *) fun rewrite_to_contra_form ct = let val (_, concl) = Logic.strip_horn (Thm.term_of ct) val concl' = dest_Trueprop concl in - if concl' aconv bFalse then + if concl' aconv UtilBase.bFalse then Conv.all_conv ct else if is_neg concl' then Util.concl_conv (Conv.rewr_conv UtilBase.to_contra_form_th') ct else Util.concl_conv (Conv.rewr_conv UtilBase.to_contra_form_th) ct end (* Rewrite ct from [...] ==> A ==> False to [...] ==> ~ A and from [...] ==> ~ A ==> False to [...] ==> A. *) fun rewrite_from_contra_form ct = let val (prems, concl) = Logic.strip_horn (Thm.term_of ct) val _ = assert (concl aconv pFalse) "rewrite_from_contra_form: concl should be false." val num_prems = length prems val last_prem' = dest_Trueprop (nth prems (num_prems-1)) val to_contra_form = if is_neg last_prem' then UtilBase.to_contra_form_th else UtilBase.to_contra_form_th' in Util.concl_conv_n (num_prems-1) (Conv.rewr_conv (meta_sym to_contra_form)) ct end (* Converts ==> to --> and !! to !. *) fun to_obj_conv ctxt ct = case Thm.term_of ct of @{const Pure.imp} $ _ $ _ => Conv.every_conv [Conv.binop_conv (to_obj_conv ctxt), Conv.rewr_conv UtilBase.atomize_imp_th] ct | Const (@{const_name Pure.all}, _) $ Abs _ => Conv.every_conv [Conv.binder_conv (to_obj_conv o snd) ctxt, Conv.rewr_conv UtilBase.atomize_all_th] ct | _ => Conv.all_conv ct (* When ct is of form A1 ==> ... ==> An, apply to_obj_conv to each Ai. *) fun to_obj_conv_on_horn ctxt ct = case Thm.term_of ct of @{const Pure.imp} $ _ $ _ => Conv.every_conv [Conv.arg1_conv (to_obj_conv ctxt), Conv.arg_conv (to_obj_conv_on_horn ctxt)] ct | Const (@{const_name Pure.all}, _) $ Abs _ => Conv.binder_conv (to_obj_conv_on_horn o snd) ctxt ct | _ => Conv.all_conv ct (* Convert !! and ==> on the outermost level. Pull all !! to the front. *) fun to_meta_conv ctxt ct = let val t = Thm.term_of ct in if is_Trueprop t andalso is_obj_all (dest_Trueprop t) then Conv.every_conv [Conv.rewr_conv (meta_sym UtilBase.atomize_all_th), to_meta_conv ctxt] ct else if is_Trueprop t andalso is_ball (dest_Trueprop t) then Conv.every_conv [Conv.arg_conv (rewr_obj_eq UtilBase.Ball_def_th), to_meta_conv ctxt] ct else if is_Trueprop t andalso is_imp (dest_Trueprop t) then Conv.every_conv [Conv.rewr_conv (meta_sym UtilBase.atomize_imp_th), to_meta_conv ctxt] ct else if Logic.is_all t then Conv.binder_conv (to_meta_conv o snd) ctxt ct else if Util.is_implies t then Conv.every_conv [Conv.arg_conv (to_meta_conv ctxt), Util.swap_meta_imp_alls ctxt] ct else Conv.all_conv ct end (* Modify th using imp_th, and add postfix to name (if available). *) fun thm_RS_mod imp_th suffix th = (th RS imp_th) |> Drule.zero_var_indexes |> Util.update_name_of_thm th suffix (* From A & B, obtain A. *) val conj_left_th = thm_RS_mod UtilBase.conjunct1_th "@left" (* From A & B, obtain B. *) val conj_right_th = thm_RS_mod UtilBase.conjunct2_th "@right" (* From (A::bool) = B, obtain A ==> B. *) val equiv_forward_th = thm_RS_mod UtilBase.iffD1_th "@eqforward" (* From (A::bool) = B, obtain B ==> A. *) val equiv_backward_th = thm_RS_mod UtilBase.iffD2_th "@eqbackward" (* From (A::bool) = B, obtain ~A ==> ~B. *) val inv_backward_th = thm_RS_mod UtilBase.inv_back_th "@invbackward" (* Same as to_obj_eq, except keeping names and indices. *) fun to_obj_eq_th th = th |> to_obj_eq |> Util.update_name_of_thm th "@obj_eq" (* Same as to_obj_eq_iff, except keeping names and indices. *) fun to_obj_eq_iff_th th = th |> UtilBase.to_obj_eq_iff |> Util.update_name_of_thm th "@iff" (* Same as obj_sym, except keeping names and indices. *) fun obj_sym_th th = th |> obj_sym |> Util.update_name_of_thm th "@sym" (* Given th of form (P ==>) A1 & ... & An, return theorems (P ==>) A1, ..., (P ==>) An, where there can be zero or more premises in front. *) fun split_conj_th th = if is_conj (prop_of' th) then (th RS UtilBase.conjunct1_th) :: (split_conj_th (th RS UtilBase.conjunct2_th)) else [th] (* More general form. *) fun split_conj_gen_th th = if is_conj (prop_of' th) then maps split_conj_gen_th [th RS UtilBase.conjunct1_th, th RS UtilBase.conjunct2_th] else [th] (* Given th of form ~ (A1 | ... | An), return theorems ~ A1, ... ~ An. *) fun split_not_disj_th th = let val t = prop_of' th in if is_neg t andalso is_disj (dest_not t) then (th RS UtilBase.or_intro1_th) :: (split_not_disj_th (th RS UtilBase.or_intro2_th)) else [th] end (* Similar to Logic.strip_horn, except remove Trueprop. *) fun strip_horn' th = (Logic.strip_horn (Thm.prop_of th)) |> apfst (map dest_Trueprop) |> apsnd dest_Trueprop fun mk_conjs_th ths = case ths of [] => raise Fail "mk_conjs_th" | [th] => th | th :: rest => [th, mk_conjs_th rest] MRS UtilBase.conjI_th (* Given th of form P x ==> False, where x is the given free variable, obtain new theorem of form (EX x. P x) ==> False. The function is written so it can be applied to multiple variables with fold. For example, "fold (ex_elim ctxt) [x, y] (P x y ==> False) will give (EX y x. P x y) ==> False. *) fun ex_elim ctxt freevar th = let val thy = Proof_Context.theory_of ctxt val th' = th |> Thm.forall_intr (Thm.cterm_of ctxt freevar) val head_prem = hd (Thm.prems_of UtilBase.exE_th') val inst = Pattern.match thy (head_prem, Thm.prop_of th') fo_init val exE_inst = Util.subst_thm ctxt inst UtilBase.exE_th' in Thm.elim_implies th' exE_inst end fun force_abs_form t = case t of Abs _ => t | _ => Abs ("x", domain_type (fastype_of t), t $ Bound 0) (* Normalization of object forall expressions in horn-clause form. *) fun strip_obj_horn t = if is_obj_all t then case t of _ $ (u as Abs _) => let val (v, body) = Term.dest_abs_global u val (vars, (assum, concl)) = strip_obj_horn body in (Free v :: vars, (assum, concl)) end | f $ arg => strip_obj_horn (f $ force_abs_form arg) | _ => raise Fail "strip_obj_horn" else if is_ball t then case t of _ $ S $ (u as Abs _) => let val (v, body) = Term.dest_abs_global u val var = Free v val mem = mk_mem (var, S) val (vars, (assum, concl)) = strip_obj_horn body in (var :: vars, (mem :: assum, concl)) end | f $ S $ arg => strip_obj_horn (f $ S $ force_abs_form arg) | _ => raise Fail "strip_obj_horn" else if is_imp t then let val (vars, (assum, concl)) = strip_obj_horn (dest_arg t) in (vars, (dest_arg1 t :: assum, concl)) end else ([], ([], t)) fun list_obj_horn (vars, (As, B)) = (list_obj_imp (As, B)) |> fold mk_obj_all (rev vars) (* Whether t can be immediately rewritten into EX form. *) fun is_ex_form_gen t = is_ex t orelse is_bex t orelse (is_neg t andalso is_obj_all (dest_not t)) orelse (is_neg t andalso is_ball (dest_not t)) orelse (is_conj t andalso is_ex_form_gen (dest_arg t)) (* Rewrite A & EX v_i. B to EX v_i. A & B. *) fun swap_conj_exists ctxt ct = let val t = Thm.term_of ct in if is_conj t andalso is_ex (dest_arg t) then Conv.every_conv [rewr_obj_eq UtilBase.swap_ex_conj_th, Conv.binder_conv (swap_conj_exists o snd) ctxt] ct else Conv.all_conv ct end (* Normalize existence statement into EX v_i. A_1 & ... & A_n. This includes moving as many existence quantifiers to the left as possible. *) fun normalize_exists ctxt ct = let val t = Thm.term_of ct in if is_ex t then Conv.binder_conv (normalize_exists o snd) ctxt ct else if is_bex t then Conv.every_conv [rewr_obj_eq UtilBase.Bex_def_th, normalize_exists ctxt] ct else if is_neg t andalso is_obj_all (dest_not t) then Conv.every_conv [rewr_obj_eq UtilBase.not_all_th, normalize_exists ctxt] ct else if is_neg t andalso is_ball (dest_not t) then Conv.every_conv [Conv.arg_conv (rewr_obj_eq UtilBase.Ball_def_th), rewr_obj_eq UtilBase.not_all_th, Conv.binder_conv (K (rewr_obj_eq UtilBase.not_imp_th)) ctxt, normalize_exists ctxt] ct else if is_conj t then Conv.every_conv [Conv.arg_conv (normalize_exists ctxt), swap_conj_exists ctxt] ct else Conv.all_conv ct end (* Assume t is in normal form. *) fun strip_exists t = if is_ex t then case t of _ $ (u as Abs _) => let val (v, body) = Term.dest_abs_global u val (vars, body') = strip_exists body in (Free v :: vars, body') end | _ => raise Fail "strip_exists" else ([], t) (* Generic wrapper function. tac can be arith_tac, simp_tac, etc. *) fun prove_by_tac tac ctxt ths goal = let val goal' = Logic.list_implies (map Thm.prop_of ths, mk_Trueprop goal) in ths MRS (Goal.prove ctxt [] [] goal' (HEADGOAL o tac o #context)) end -fun contra_by_tac tac ctxt ths = prove_by_tac tac ctxt ths bFalse +fun contra_by_tac tac ctxt ths = prove_by_tac tac ctxt ths UtilBase.bFalse end (* structure UtilLogic *) - -structure Basic_UtilLogic: BASIC_UTIL_LOGIC = UtilLogic -open Basic_UtilLogic diff --git a/thys/Auto2_HOL/wfdata.ML b/thys/Auto2_HOL/wfdata.ML --- a/thys/Auto2_HOL/wfdata.ML +++ b/thys/Auto2_HOL/wfdata.ML @@ -1,381 +1,389 @@ (* File: wfdata.ML Author: Bohua Zhan Table storing wellformed-ness data during the proof. *) signature WELLFORM_DATA = sig val clean_resolved: box_id -> Proof.context -> Proof.context val clear_incr: Proof.context -> Proof.context val initialize_wellform_data: term -> Proof.context -> Proof.context (* wellform table. *) val get_wellform_for_term: Proof.context -> term -> (cterm * (box_id * thm) list) list val get_wellform_infos_for_term: Proof.context -> term -> (box_id * thm) list val convert_wellform: Proof.context -> box_id * thm -> box_id * thm -> box_id * thm val get_wellform: Proof.context -> box_id * cterm -> (box_id * thm) list val get_wellform_t: Proof.context -> box_id * term -> (box_id * thm) list val get_complete_wellform: Proof.context -> box_id * cterm -> (box_id * thm list) list val cterm_to_wfterm: Proof.context -> term list -> box_id * cterm -> (box_id * wfterm) list val term_to_wfterm: Proof.context -> term list -> box_id * term -> (box_id * wfterm) list (* Update to wellform table. *) val add_wellform_data_raw: term * (box_id * thm) -> Proof.context -> Proof.context val find_fact: Proof.context -> box_item list -> box_id * cterm -> (box_id * thm) list val complete_wellform_data_for_terms: box_item list -> term list -> Proof.context -> Proof.context val complete_wellform_data: box_item list -> Proof.context -> Proof.context (* Get changes to wellform table. *) val get_new_wellform_data: Proof.context -> (term * (box_id * thm)) list (* More queries on the wellform table.*) val simplify_info: Proof.context -> term list -> cterm -> (box_id * (wfterm * thm)) list val simplify: Proof.context -> term list -> cterm list -> box_id * wfterm -> (box_id * (wfterm * thm)) list val get_head_equiv: Proof.context -> term list -> box_id * cterm -> (box_id * (wfterm * thm)) list end; -structure WellformData : WELLFORM_DATA = +functor WellformData( + structure Basic_UtilBase : BASIC_UTIL_BASE; + structure Basic_UtilLogic : BASIC_UTIL_LOGIC; + structure ItemIO: ITEM_IO; + structure Property: PROPERTY; + structure PropertyData: PROPERTY_DATA; + structure RewriteTable: REWRITE_TABLE; + structure WfTerm: WFTERM + ) : WELLFORM_DATA = struct structure Data = Proof_Data ( type T = ((cterm * (box_id * thm) list) list) Termtab.table fun init _ = Termtab.empty ) fun clean_resolved id ctxt = let fun clean_wellform_1 (ctarget, infos) = (ctarget, filter_out (BoxID.is_eq_ancestor ctxt id o fst) infos) fun clean_wellform tb = tb |> Termtab.map (fn _ => map clean_wellform_1) in ctxt |> Data.map clean_wellform end fun clear_incr ctxt = let fun clear_one (ct, infos) = if exists (BoxID.has_incr_id o fst) infos then (ct, infos |> map (apfst BoxID.replace_incr_id) |> Util.max_partial (BoxID.info_eq_better ctxt)) else (ct, infos) in ctxt |> Data.map (Termtab.map (fn _ => map clear_one)) end (* Initialize wellform data for term t, with (currently) empty slots for each target. *) fun initialize_wellform_data t ctxt = let val wellform = Data.get ctxt val thy = Proof_Context.theory_of ctxt val targets = map (Thm.cterm_of ctxt) (WellForm.lookup_wellform_data thy t) in if Termtab.defined wellform t then ctxt else if null targets then ctxt else ctxt |> Data.map (Termtab.update (t, map (rpair []) targets)) end fun get_wellform_for_term ctxt t = let val wellform = Data.get ctxt in the_default [] (Termtab.lookup wellform t) end (* Given a term t, return the wellform data for t in a list of (id, th) pairs. *) fun get_wellform_infos_for_term ctxt t = maps snd (get_wellform_for_term ctxt t) (* Given (id, th) a wellform data for the left side of eq_th, convert to a wellform data for the right side. *) fun convert_wellform ctxt (id', eq_th) (id, th) = let val (lhs, rhs) = Logic.dest_equals (Thm.prop_of eq_th) val thy = Proof_Context.theory_of ctxt in if lhs aconv rhs then (id, th) else - case WellForm.lookup_wellform_pattern thy (Util.lhs_of eq_th, prop_of' th) of + case WellForm.lookup_wellform_pattern thy (Util.lhs_of eq_th, Basic_UtilLogic.prop_of' th) of NONE => raise Fail "convert_wellform: invalid input." | SOME (pat, data_pat) => let (* Cannot use eq_th directly. Rather, find all the equivalences again for all the subterms. *) val cargs1 = Util.dest_cargs (Thm.lhs_of eq_th) val cargs2 = Util.dest_cargs (Thm.rhs_of eq_th) val eq_ths = map (fn (ct, ct') => (RewriteTable.equiv_info ctxt id' (ct, ct')) |> filter (fn (id'', _) => id'' = id') |> map snd |> the_single) (cargs1 ~~ cargs2) val pat_args = Util.dest_args pat - val th' = apply_to_thm' ( + val th' = Basic_UtilLogic.apply_to_thm' ( Util.pattern_rewr_conv data_pat (pat_args ~~ eq_ths)) th (* Check the result is in the right form. *) val subst' = pat_args ~~ map Thm.term_of cargs2 val prop' = Term.subst_atomic subst' data_pat - val _ = assert (prop_of' th' aconv prop') + val _ = assert (Basic_UtilLogic.prop_of' th' aconv prop') "convert_wellform: invalid output." in (BoxID.merge_boxes ctxt (id, id'), th') end end (* Retrieve wellform data for the given term. May need to rewrite t up to subterm equivalence. *) fun get_wellform ctxt (id, ct) = let val t = Thm.term_of ct in if RewriteTable.in_table_raw_for_id ctxt (id, t) then (get_wellform_infos_for_term ctxt t) |> BoxID.merge_box_with_info ctxt id else let (* eq_th rewrites t to t', where t' is subterm equivalent to t and in the table. Obtain wellform data for t', then convert them into wellform data for t. *) fun process_head_rep (id', eq_th) = (get_wellform_infos_for_term ctxt (Util.rhs_of eq_th)) |> map (convert_wellform ctxt (id', meta_sym eq_th)) in (RewriteTable.subterm_simplify_info ctxt ct) |> maps (RewriteTable.get_head_rep_with_id_th ctxt) |> maps process_head_rep |> BoxID.merge_box_with_info ctxt id |> Util.max_partial (BoxID.info_eq_better ctxt) end end fun get_wellform_t ctxt (id, t) = get_wellform ctxt (id, Thm.cterm_of ctxt t) (* Retrieve the complete set of wellform data for term t, indexed by box id. *) fun get_complete_wellform ctxt (id, ct) = let val thy = Proof_Context.theory_of ctxt val t = Thm.term_of ct val targets = WellForm.lookup_wellform_data thy t in case targets of [] => [(id, [])] | _ => let val data = (get_wellform ctxt (id, ct)) - |> map (fn (id', th) => (prop_of' th, (id', th))) + |> map (fn (id', th) => (Basic_UtilLogic.prop_of' th, (id', th))) |> AList.group (op aconv) |> map snd val _ = assert (length data <= length targets) "get_complete_wellform: unexpected length data" in if length data < length targets then [] else data |> BoxID.get_all_merges_info ctxt |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) end end (* Given a regular term, construct a wellformed term by adding theorems corresponding to any combination whose head agrees with one of fheads. *) fun cterm_to_wfterm ctxt fheads (id, ct) = case Thm.term_of ct of _ $ _ => let val t = Thm.term_of ct val (cf, cargs) = Drule.strip_comb ct in if forall (fn fhead => not (Util.is_head fhead t)) fheads then [(id, WfTerm ct)] else let val wfdata = get_complete_wellform ctxt (id, ct) fun process_wfdata (id', ths) = (* For each wellform data of t at box id', retrieve the wellform data of each argment, then merge together. *) cargs |> map (pair id') |> map (cterm_to_wfterm ctxt fheads) |> BoxID.get_all_merges_info ctxt |> map (fn (id'', arg') => (id'', WfComb (cf, arg', ths))) in (maps process_wfdata wfdata) |> Util.max_partial (BoxID.id_is_eq_ancestor ctxt) end end | _ => [(id, WfTerm ct)] fun term_to_wfterm ctxt fheads (id, t) = cterm_to_wfterm ctxt fheads (id, Thm.cterm_of ctxt t) (* Add the given wellform data (id, th) for the term t to the table. *) fun add_wellform_data_raw (t, (id, th)) ctxt = let - val cprop = cprop_of' th + val cprop = Basic_UtilLogic.cprop_of' th val wellform_data = get_wellform_for_term ctxt t val infos = the (AList.lookup (op aconvc) wellform_data cprop) handle Option.Option => raise Fail "add_wellform_data_raw" in if exists (fn info => BoxID.info_eq_better ctxt info (id, th)) infos then ctxt else let val infos' = infos |> filter_out (BoxID.info_eq_better ctxt (id, th)) |> cons (id, th) in ctxt |> Data.map ( Termtab.map_entry t (AList.map_entry (op aconvc) cprop (K infos'))) end end (* Given a list of items, attempt to justify t (of type bool) using the theorems, or equalities / properties in the rewrite table. Return a list of (id, th) pairs, where the statement of th' is t. *) fun find_fact ctxt items (id, ct) = let val t = Thm.term_of ct fun match_item item = (ItemIO.match_arg ctxt (PropMatch t) item (id, fo_init)) |> map (fn ((id', _), th) => (id', th)) in - if is_eq_term t then (* Equality case *) - map (apsnd to_obj_eq) - (RewriteTable.equiv_info ctxt id (cdest_eq ct)) + if Basic_UtilBase.is_eq_term t then (* Equality case *) + map (apsnd Basic_UtilLogic.to_obj_eq) + (RewriteTable.equiv_info ctxt id (Basic_UtilBase.cdest_eq ct)) else if Property.is_property t then (* Property case *) PropertyData.get_property ctxt (id, ct) else (* General case *) maps match_item items end (* Given a list of (id, th), complete the wellform data. New wellform data could be added when: - There are new theorems in the list of (id, th). - There are new properties added. - There are new equalities added. *) fun complete_wellform_data_for_terms items ts ctxt = let fun get_for_t t = let val cur_data = get_wellform_for_term ctxt t val target_ids = RewriteTable.in_table_raw_ids ctxt t fun is_target (id, cprop) = let val prop_data = the (AList.lookup (op aconvc) cur_data cprop) in not (exists (fn (id', _) => BoxID.is_eq_ancestor ctxt id' id) prop_data) end val targets = (Util.all_pairs (target_ids, map fst cur_data)) |> filter is_target in targets |> maps (find_fact ctxt items) |> map (pair t) end val all_data = maps get_for_t ts in fold add_wellform_data_raw all_data ctxt end fun complete_wellform_data items ctxt = let val wellform = Data.get ctxt val ts = Termtab.keys wellform in complete_wellform_data_for_terms items ts ctxt end (* Return the list of new wellform data (added under id prim or lower). Return as a list of (t, (id, th)). *) fun get_new_wellform_data ctxt = let val wellform = Data.get ctxt in wellform |> Termtab.dest |> maps (fn (t, vals) => map (pair t) (maps snd vals)) |> filter (fn (_, (id, _)) => BoxID.has_incr_id id) end (* Find simplifications of ct that are well-formed. *) fun simplify_info ctxt fheads ct = let val infos = RewriteTable.simplify_info ctxt ct fun process_info (id, eq_th) = let val rhs = Thm.rhs_of eq_th val wfts = cterm_to_wfterm ctxt fheads (id, rhs) in map (fn (id', wft) => (id', (wft, eq_th))) wfts end in maps process_info infos end (* Assume cts are distinct subterms of wft, find all combinations of simplifying the terms ts. *) fun simplify ctxt fheads cts (id, wft) = if null cts then [(id, (wft, Thm.reflexive (WfTerm.cterm_of wft)))] else cts |> map (simplify_info ctxt fheads) |> BoxID.get_all_merges_info ctxt |> BoxID.merge_box_with_info ctxt id |> map (fn (id', wf_eqs) => (id', WfTerm.rewrite_on_eqs fheads wf_eqs wft)) (* Find all well-formed head equivalences of ct. Return as a list of (id, (wft, eq_th)) pairs. *) fun get_head_equiv ctxt fheads (id, ct) = let val head_equivs = ct |> RewriteTable.get_head_equiv ctxt |> BoxID.merge_box_with_info ctxt id fun process_head_equiv (id', eq_th) = let val rhs = Thm.rhs_of eq_th val wfts = cterm_to_wfterm ctxt fheads (id', rhs) in map (fn (id'', wft) => (id'', (wft, eq_th))) wfts end in maps process_head_equiv head_equivs end end (* WellformData *) diff --git a/thys/Auto2_HOL/wfterm.ML b/thys/Auto2_HOL/wfterm.ML --- a/thys/Auto2_HOL/wfterm.ML +++ b/thys/Auto2_HOL/wfterm.ML @@ -1,314 +1,319 @@ (* File: wfterm.ML Author: Bohua Zhan Definition and basic functions on Wellformed terms. *) (* A wellformed term is a cterm combined with wellform data (usually only for a particular function head. *) datatype wfterm = WfTerm of cterm | WfComb of cterm * wfterm list * thm list infix 1 then_wfconv infix 0 else_wfconv (* Conversion between wellformed terms. Produce a new wellformed term, together with an equality between corresponding terms. *) type wfconv = wfterm -> wfterm * thm signature WFTERM = sig (* Conversion between term and wfterm. *) val cterm_of: wfterm -> cterm val term_of: wfterm -> term val theory_of_wft: wfterm -> theory val wellform_ths_of: wfterm -> thm list val find_target_on_ths: thm list -> term -> thm val cterm_to_wfterm_on_ths: thm list -> term list -> cterm -> wfterm val cterm_to_wfterm_assume: term list -> cterm -> wfterm (* Wellformed version of conv. *) val strip_comb: wfterm -> cterm * wfterm list val all_conv: wfconv val no_conv: wfconv val argn_conv: int -> wfconv -> wfconv val arg_conv: wfconv -> wfconv val arg1_conv: wfconv -> wfconv val rewr_obj_eq: term list -> thm -> wfconv val then_wfconv: wfconv * wfconv -> wfconv val else_wfconv: wfconv * wfconv -> wfconv val try_conv: wfconv -> wfconv val repeat_conv: wfconv -> wfconv val first_conv: wfconv list -> wfconv val every_conv: wfconv list -> wfconv val binop_conv: wfconv -> wfconv val conv_of: conv -> wfconv (* Some debugging facilities. *) val string_of_wfterm: Proof.context -> wfterm -> string val test_wfconv: Proof.context -> term list -> wfconv -> string -> string * string -> unit (* Rewriting a wellformed expression. *) val rewrite_on_eqs: term list -> (wfterm * thm) list -> wfconv end; -structure WfTerm : WFTERM = +functor WfTerm( + structure Basic_UtilBase: BASIC_UTIL_BASE; + structure UtilLogic : UTIL_LOGIC + ) : WFTERM = struct (* Extract a regular cterm from a wellformed term. *) fun cterm_of wft = case wft of WfTerm ct => ct | WfComb (cf, args, _) => Drule.list_comb (cf, map cterm_of args) fun term_of wft = Thm.term_of (cterm_of wft) fun theory_of_wft wft = Thm.theory_of_cterm (cterm_of wft) fun wellform_ths_of wft = case wft of WfTerm _ => [] | WfComb (_, args, ths) => ths @ (maps wellform_ths_of args) fun find_target_on_ths ths prop = let - fun get_on_th th = if prop_of' th aconv prop then SOME th else NONE + fun get_on_th th = if UtilLogic.prop_of' th aconv prop then SOME th else NONE in the (get_first get_on_th ths) handle Option.Option => raise Fail "find_target_on_ths: not found" end (* Obtain a wellformed term from a cterm and a list of theorems. Traverse only through combinations whose function head agrees with one of fheads. Throws an exception if any wellform data is not found. *) fun cterm_to_wfterm_on_ths assum_ths fheads ct = case Thm.term_of ct of _ $ _ => let val thy = Thm.theory_of_cterm ct val t = Thm.term_of ct val (cf, cargs) = Drule.strip_comb ct in if forall (fn fhead => not (Util.is_head fhead t)) fheads then WfTerm ct else let val ths = map (find_target_on_ths assum_ths) (WellForm.lookup_wellform_data thy t) handle Fail "find_target_on_ths: not found" => raise Fail "cterm_to_wfterm_on_ths" val wfargs = map (cterm_to_wfterm_on_ths assum_ths fheads) cargs in WfComb (cf, wfargs, ths) end end | _ => WfTerm ct fun cterm_to_wfterm_assume fheads ct = case Thm.term_of ct of _ $ _ => let val thy = Thm.theory_of_cterm ct val t = Thm.term_of ct val (cf, cargs) = Drule.strip_comb ct in if forall (fn fhead => not (Util.is_head fhead t)) fheads then WfTerm ct else let - val ths = map (Thm.assume o Thm.global_cterm_of thy o mk_Trueprop) + val ths = map (Thm.assume o Thm.global_cterm_of thy o UtilLogic.mk_Trueprop) (WellForm.lookup_wellform_data thy t) val wfargs = map (cterm_to_wfterm_assume fheads) cargs in WfComb (cf, wfargs, ths) end end | _ => WfTerm ct fun strip_comb wft = case wft of WfTerm ct => (ct, []) | WfComb (cf, args, _) => (cf, args) fun all_conv wft = (wft, Thm.reflexive (cterm_of wft)) fun no_conv _ = raise Fail "WfTerm.no_conv" (* Apply wfcv to the nth argument of wft (counting from the left). Remember to update the wellform data at head. *) fun argn_conv n wfcv wft = case wft of WfTerm _ => raise Fail "argn_conv: applied to atom." | WfComb (cf, args, ths) => let val _ = assert (n < length args) "argn_conv: n out of bounds." val thy = Thm.theory_of_cterm cf val (wft', eq_th) = wfcv (nth args n) (* New list of arguments, and equalities between the two lists. *) val args' = take n args @ [wft'] @ drop (n + 1) args val eq_ths' = map (Thm.reflexive o cterm_of) (take n args) @ [eq_th] @ map (Thm.reflexive o cterm_of) (drop (n + 1) args) fun convert_th th = let (* Obtain pattern from theory. *) val (pat, data_pat) = the (WellForm.lookup_wellform_pattern - thy (term_of wft, prop_of' th)) + thy (term_of wft, UtilLogic.prop_of' th)) handle Option.Option => raise Fail "convert_th: invalid input." val pat_args = Util.dest_args pat in - apply_to_thm' ( + UtilLogic.apply_to_thm' ( Util.pattern_rewr_conv data_pat (pat_args ~~ eq_ths')) th end in (WfComb (cf, args', map convert_th ths), Util.comb_equiv (cf, eq_ths')) end fun arg_conv wfcv wft = let val (_, args) = Term.strip_comb (term_of wft) in argn_conv (length args - 1) wfcv wft end fun arg1_conv wfcv wft = let val (_, args) = Term.strip_comb (term_of wft) in argn_conv (length args - 2) wfcv wft end (* Apply the given theorem as a rewrite rule. th is assumed to be in the form P_1 ==> ... ==> P_m ==> l = r & Q_1 & ... & Q_n, where P_1 through P_m provide wellform data for l, and Q_1 through Q_n provide additional wellform data for r. *) fun rewr_obj_eq fheads th wft = let val thy = theory_of_wft wft val t = term_of wft val wf_ths = wellform_ths_of wft (* First need to instantiate th by matching l with the term of wft. *) - val pat_l = th |> concl_of' |> strip_conj |> hd |> dest_eq |> fst + val pat_l = th + |> UtilLogic.concl_of' + |> UtilLogic.strip_conj + |> hd + |> Basic_UtilBase.dest_eq + |> fst val inst = Pattern.first_order_match thy (pat_l, t) fo_init val th' = Util.subst_thm_thy thy inst th (* Second, find theorems justifying premises of th'. *) - val prems = map dest_Trueprop (Thm.prems_of th') + val prems = map UtilLogic.dest_Trueprop (Thm.prems_of th') val prems_th = map (find_target_on_ths wf_ths) prems handle Fail "find_target_on_ths: not found" => raise Fail "WfTerm.rewr_obj_eq" val th'' = prems_th MRS th' (* Finally, apply rewrite rule and fill in wellform data on rhs. *) val th_splits = UtilLogic.split_conj_th th'' val (eq_th, wf_ths') = (hd th_splits, tl th_splits) val c_rhs = eq_th |> Thm.cprop_of |> Thm.dest_arg |> Thm.dest_arg val wf_rhs = cterm_to_wfterm_on_ths (wf_ths @ wf_ths') fheads c_rhs in - (wf_rhs, to_meta_eq eq_th) + (wf_rhs, UtilLogic.to_meta_eq eq_th) end handle Pattern.MATCH => raise Fail "WfTerm.rewr_obj_eq failed" fun (wfcv1 then_wfconv wfcv2) wft = let val (wft', eq_th1) = wfcv1 wft val (wft'', eq_th2) = wfcv2 wft' in (wft'', Util.transitive_list [eq_th1, eq_th2]) end fun (wfcv1 else_wfconv wfcv2) wft = (wfcv1 wft handle THM _ => wfcv2 wft | CTERM _ => wfcv2 wft | TERM _ => wfcv2 wft | TYPE _ => wfcv2 wft | Fail _ => wfcv2 wft) fun try_conv wfcv = wfcv else_wfconv all_conv fun repeat_conv wfcv wft = try_conv (wfcv then_wfconv repeat_conv wfcv) wft fun first_conv wfcvs = fold_rev (curry op else_wfconv) wfcvs no_conv fun every_conv wfcvs = fold_rev (curry op then_wfconv) wfcvs all_conv fun binop_conv wfcv = arg1_conv wfcv then_wfconv arg_conv wfcv (* Transforming a regular conversion to a well-formed conversion. This should be used only when the output of the conversion is guaranteed to not have further wellform structure. *) fun conv_of cv wft = let val ct = cterm_of wft val eq_th = cv ct in (WfTerm (Thm.rhs_of eq_th), eq_th) end fun string_of_wfterm ctxt wft = case wft of WfTerm ct => "@(" ^ Syntax.string_of_term ctxt (Thm.term_of ct) ^ ")" | WfComb (ct, wfts, _) => Syntax.string_of_term ctxt (Thm.term_of ct) ^ " $ " ^ (space_implode " $ " (map (string_of_wfterm ctxt) wfts)) (* Apply wfcv on the term given by str_t, and compare the result to the term given by str_res. *) fun test_wfconv ctxt fheads wfcv err_str (str1, str2) = let val (t1, t2) = (Syntax.read_term ctxt str1, Syntax.read_term ctxt str2) val wft = cterm_to_wfterm_assume fheads (Thm.cterm_of ctxt t1) val th = snd (wfcv wft) handle Fail err => let val _ = trace_t ctxt "Input:" t1 in raise Fail (err ^ " -- " ^ err_str) end in if t1 aconv (Util.lhs_of th) andalso t2 aconv (Util.rhs_of th) then () else let val _ = trace_t ctxt "Input:" t1 val _ = trace_t ctxt "Expected:" t2 val _ = trace_t ctxt "Actual:" (Thm.prop_of th) in raise Fail err_str end end (* Given a well-formed term wft, and a list of equations to well-formed terms, rewrite all subterms of wft that are equal to the left side of one of the equations. Assume that all fheads are binary functions. *) fun rewrite_on_eqs fheads wf_eqs wft = let val t = term_of wft in if member (op aconv) fheads (Term.head_of t) then binop_conv (rewrite_on_eqs fheads wf_eqs) wft else case find_first (fn (_, eq) => t aconv (Util.lhs_of eq)) wf_eqs of NONE => all_conv wft | SOME (wft, eq_th) => (wft, eq_th) end end (* structure WfTerm *) - -fun (wfcv1 then_wfconv wfcv2) = WfTerm.then_wfconv (wfcv1, wfcv2) -fun (wfcv1 else_wfconv wfcv2) = WfTerm.else_wfconv (wfcv1, wfcv2) diff --git a/thys/Auto2_Imperative_HOL/Imperative/SepAuto.thy b/thys/Auto2_Imperative_HOL/Imperative/SepAuto.thy --- a/thys/Auto2_Imperative_HOL/Imperative/SepAuto.thy +++ b/thys/Auto2_Imperative_HOL/Imperative/SepAuto.thy @@ -1,517 +1,517 @@ (* File: SepAuto.thy Author: Bohua Zhan *) section \Separation logic\ theory SepAuto imports SepLogic_Base "HOL-Imperative_HOL.Imperative_HOL" begin text \ Separation logic for Imperative\_HOL, and setup of auto2. The development of separation logic here follows \cite{Separation_Logic_Imperative_HOL-AFP} by Lammich and Meis. \ subsection \Partial Heaps\ datatype pheap = pHeap (heapOf: heap) (addrOf: "addr set") setup \add_simple_datatype "pheap"\ fun in_range :: "(heap \ addr set) \ bool" where "in_range (h,as) \ (\a\as. a < lim h)" setup \add_rewrite_rule @{thm in_range.simps}\ text \Two heaps agree on a set of addresses.\ definition relH :: "addr set \ heap \ heap \ bool" where [rewrite]: "relH as h h' = (in_range (h, as) \ in_range (h', as) \ (\t. \a\as. refs h t a = refs h' t a \ arrays h t a = arrays h' t a))" lemma relH_D [forward]: "relH as h h' \ in_range (h, as) \ in_range (h', as)" by auto2 lemma relH_D2 [rewrite]: "relH as h h' \ a \ as \ refs h t a = refs h' t a" "relH as h h' \ a \ as \ arrays h t a = arrays h' t a" by auto2+ setup \del_prfstep_thm_eqforward @{thm relH_def}\ lemma relH_dist_union [forward]: "relH (as \ as') h h' \ relH as h h' \ relH as' h h'" by auto2 lemma relH_ref [rewrite]: "relH as h h' \ addr_of_ref r \ as \ Ref.get h r = Ref.get h' r" - by (simp add: Ref.get_def relH_def) + by (auto intro: relH_D2 arg_cong simp: Ref.get_def) lemma relH_array [rewrite]: "relH as h h' \ addr_of_array r \ as \ Array.get h r = Array.get h' r" - by (simp add: Array.get_def relH_def) + by (auto intro: relH_D2 arg_cong simp: Array.get_def) lemma relH_set_ref [resolve]: "relH {a. a < lim h \ a \ {addr_of_ref r}} h (Ref.set r x h)" by (simp add: Ref.set_def relH_def) lemma relH_set_array [resolve]: "relH {a. a < lim h \ a \ {addr_of_array r}} h (Array.set r x h)" by (simp add: Array.set_def relH_def) subsection \Assertions\ datatype assn_raw = Assn (assn_fn: "pheap \ bool") fun aseval :: "assn_raw \ pheap \ bool" where "aseval (Assn f) h = f h" setup \add_rewrite_rule @{thm aseval.simps}\ definition proper :: "assn_raw \ bool" where [rewrite]: "proper P = ( (\h as. aseval P (pHeap h as) \ in_range (h,as)) \ (\h h' as. aseval P (pHeap h as) \ relH as h h' \ in_range (h',as) \ aseval P (pHeap h' as)))" fun in_range_assn :: "pheap \ bool" where "in_range_assn (pHeap h as) \ (\a\as. a < lim h)" setup \add_rewrite_rule @{thm in_range_assn.simps}\ typedef assn = "Collect proper" @proof @have "Assn in_range_assn \ Collect proper" @qed setup \add_rewrite_rule @{thm Rep_assn_inject}\ setup \register_wellform_data ("Abs_assn P", ["proper P"])\ setup \add_prfstep_check_req ("Abs_assn P", "proper P")\ lemma Abs_assn_inverse' [rewrite]: "proper y \ Rep_assn (Abs_assn y) = y" by (simp add: Abs_assn_inverse) lemma proper_Rep_assn [forward]: "proper (Rep_assn P)" using Rep_assn by auto definition models :: "pheap \ assn \ bool" (infix "\" 50) where [rewrite_bidir]: "h \ P \ aseval (Rep_assn P) h" lemma models_in_range [resolve]: "pHeap h as \ P \ in_range (h,as)" by auto2 lemma mod_relH [forward]: "relH as h h' \ pHeap h as \ P \ pHeap h' as \ P" by auto2 instantiation assn :: one begin definition one_assn :: assn where [rewrite]: "1 \ Abs_assn (Assn (\h. addrOf h = {}))" instance .. end abbreviation one_assn :: assn ("emp") where "one_assn \ 1" lemma one_assn_rule [rewrite]: "h \ emp \ addrOf h = {}" by auto2 setup \del_prfstep_thm @{thm one_assn_def}\ instantiation assn :: times begin definition times_assn where [rewrite]: "P * Q = Abs_assn (Assn ( \h. (\as1 as2. addrOf h = as1 \ as2 \ as1 \ as2 = {} \ aseval (Rep_assn P) (pHeap (heapOf h) as1) \ aseval (Rep_assn Q) (pHeap (heapOf h) as2))))" instance .. end lemma mod_star_conv [rewrite]: "pHeap h as \ A * B \ (\as1 as2. as = as1 \ as2 \ as1 \ as2 = {} \ pHeap h as1 \ A \ pHeap h as2 \ B)" by auto2 setup \del_prfstep_thm @{thm times_assn_def}\ lemma aseval_ext [backward]: "\h. aseval P h = aseval P' h \ P = P'" apply (cases P) apply (cases P') by auto lemma assn_ext: "\h as. pHeap h as \ P \ pHeap h as \ Q \ P = Q" @proof @have "Rep_assn P = Rep_assn Q" @qed setup \add_backward_prfstep_cond @{thm assn_ext} [with_filt (order_filter "P" "Q")]\ setup \del_prfstep_thm @{thm aseval_ext}\ lemma assn_one_left: "1 * P = (P::assn)" @proof @have "\h as. pHeap h as \ P \ pHeap h as \ 1 * P" @with @have "as = {} \ as" @end @qed lemma assn_times_comm: "P * Q = Q * (P::assn)" @proof @have "\h as. pHeap h as \ P * Q \ pHeap h as \ Q * P" @with @case "pHeap h as \ P * Q" @with @obtain as1 as2 where "as = as1 \ as2" "as1 \ as2 = {}" "pHeap h as1 \ P" "pHeap h as2 \ Q" @have "as = as2 \ as1" @end @case "pHeap h as \ Q * P" @with @obtain as1 as2 where "as = as1 \ as2" "as1 \ as2 = {}" "pHeap h as1 \ Q" "pHeap h as2 \ P" @have "as = as2 \ as1" @end @end @qed lemma assn_times_assoc: "(P * Q) * R = P * (Q * (R::assn))" @proof @have "\h as. pHeap h as \ (P * Q) * R \ pHeap h as \ P * (Q * R)" @with @case "pHeap h as \ (P * Q) * R" @with @obtain as1 as2 where "as = as1 \ as2" "as1 \ as2 = {}" "pHeap h as1 \ P * Q" "pHeap h as2 \ R" @obtain as11 as12 where "as1 = as11 \ as12" "as11 \ as12 = {}" "pHeap h as11 \ P" "pHeap h as12 \ Q" @have "as = as11 \ (as12 \ as2)" @end @case "pHeap h as \ P * (Q * R)" @with @obtain as1 as2 where "as = as1 \ as2" "as1 \ as2 = {}" "pHeap h as1 \ P" "pHeap h as2 \ Q * R" @obtain as21 as22 where "as2 = as21 \ as22" "as21 \ as22 = {}" "pHeap h as21 \ Q" "pHeap h as22 \ R" @have "as = (as1 \ as21) \ as22" @end @end @qed instantiation assn :: comm_monoid_mult begin instance apply standard apply (rule assn_times_assoc) apply (rule assn_times_comm) by (rule assn_one_left) end subsubsection \Existential Quantification\ definition ex_assn :: "('a \ assn) \ assn" (binder "\\<^sub>A" 11) where [rewrite]: "(\\<^sub>Ax. P x) = Abs_assn (Assn (\h. \x. h \ P x))" lemma mod_ex_dist [rewrite]: "(h \ (\\<^sub>Ax. P x)) \ (\x. h \ P x)" by auto2 setup \del_prfstep_thm @{thm ex_assn_def}\ lemma ex_distrib_star: "(\\<^sub>Ax. P x * Q) = (\\<^sub>Ax. P x) * Q" @proof @have "\h as. pHeap h as \ (\\<^sub>Ax. P x) * Q \ pHeap h as \ (\\<^sub>Ax. P x * Q)" @with @case "pHeap h as \ (\\<^sub>Ax. P x) * Q" @with @obtain as1 as2 where "as = as1 \ as2" "as1 \ as2 = {}" "pHeap h as1 \ (\\<^sub>Ax. P x)" "pHeap h as2 \ Q" @obtain x where "pHeap h as1 \ P x" @have "pHeap h as \ P x * Q" @end @end @qed subsubsection \Pointers\ definition sngr_assn :: "'a::heap ref \ 'a \ assn" (infix "\\<^sub>r" 82) where [rewrite]: "r \\<^sub>r x = Abs_assn (Assn ( \h. Ref.get (heapOf h) r = x \ addrOf h = {addr_of_ref r} \ addr_of_ref r < lim (heapOf h)))" lemma sngr_assn_rule [rewrite]: "pHeap h as \ r \\<^sub>r x \ (Ref.get h r = x \ as = {addr_of_ref r} \ addr_of_ref r < lim h)" by auto2 setup \del_prfstep_thm @{thm sngr_assn_def}\ definition snga_assn :: "'a::heap array \ 'a list \ assn" (infix "\\<^sub>a" 82) where [rewrite]: "r \\<^sub>a x = Abs_assn (Assn ( \h. Array.get (heapOf h) r = x \ addrOf h = {addr_of_array r} \ addr_of_array r < lim (heapOf h)))" lemma snga_assn_rule [rewrite]: "pHeap h as \ r \\<^sub>a x \ (Array.get h r = x \ as = {addr_of_array r} \ addr_of_array r < lim h)" by auto2 setup \del_prfstep_thm @{thm snga_assn_def}\ subsubsection \Pure Assertions\ definition pure_assn :: "bool \ assn" ("\") where [rewrite]: "\b = Abs_assn (Assn (\h. addrOf h = {} \ b))" lemma pure_assn_rule [rewrite]: "h \ \b \ (addrOf h = {} \ b)" by auto2 setup \del_prfstep_thm @{thm pure_assn_def}\ definition top_assn :: assn ("true") where [rewrite]: "top_assn = Abs_assn (Assn in_range_assn)" lemma top_assn_rule [rewrite]: "pHeap h as \ true \ in_range (h, as)" by auto2 setup \del_prfstep_thm @{thm top_assn_def}\ setup \del_prfstep_thm @{thm models_def}\ subsubsection \Properties of assertions\ abbreviation bot_assn :: assn ("false") where "bot_assn \ \False" lemma top_assn_reduce: "true * true = true" @proof @have "\h. h \ true \ h \ true * true" @with @have "addrOf h = addrOf h \ {}" @end @qed lemma mod_pure_star_dist [rewrite]: "h \ P * \b \ (h \ P \ b)" @proof @case "h \ P \ b" @with @have "addrOf h = addrOf h \ {}" @end @qed lemma pure_conj: "\(P \ Q) = \P * \Q" by auto2 subsubsection \Entailment and its properties\ definition entails :: "assn \ assn \ bool" (infix "\\<^sub>A" 10) where [rewrite]: "(P \\<^sub>A Q) \ (\h. h \ P \ h \ Q)" lemma entails_triv: "A \\<^sub>A A" by auto2 lemma entails_true: "A \\<^sub>A true" by auto2 lemma entails_frame [backward]: "P \\<^sub>A Q \ P * R \\<^sub>A Q * R" by auto2 lemma entails_frame': "\ (A * F \\<^sub>A Q) \ A \\<^sub>A B \ \ (B * F \\<^sub>A Q)" by auto2 lemma entails_frame'': "\ (P \\<^sub>A B * F) \ A \\<^sub>A B \ \ (P \\<^sub>A A * F)" by auto2 lemma entails_equiv_forward: "P = Q \ P \\<^sub>A Q" by auto2 lemma entails_equiv_backward: "P = Q \ Q \\<^sub>A P" by auto2 lemma entailsD [forward]: "P \\<^sub>A Q \ h \ P \ h \ Q" by auto2 lemma entails_trans2: "A \\<^sub>A D * B \ B \\<^sub>A C \ A \\<^sub>A D * C" by auto2 lemma entails_pure': "\(\b \\<^sub>A Q) \ (\(emp \\<^sub>A Q) \ b)" by auto2 lemma entails_pure: "\(P * \b \\<^sub>A Q) \ (\(P \\<^sub>A Q) \ b)" by auto2 lemma entails_ex: "\((\\<^sub>Ax. P x) \\<^sub>A Q) \ (\x. \(P x \\<^sub>A Q))" by auto2 lemma entails_ex_post: "\(P \\<^sub>A (\\<^sub>Ax. Q x)) \ \x. \(P \\<^sub>A Q x)" by auto2 lemma entails_pure_post: "\(P \\<^sub>A Q * \b) \ P \\<^sub>A Q \ \b" by auto2 setup \del_prfstep_thm @{thm entails_def}\ subsection \Definition of the run predicate\ inductive run :: "'a Heap \ heap option \ heap option \ 'a \ bool" where "run c None None r" | "execute c h = None \ run c (Some h) None r" | "execute c h = Some (r, h') \ run c (Some h) (Some h') r" setup \add_case_induct_rule @{thm run.cases}\ setup \fold add_resolve_prfstep @{thms run.intros(1,2)}\ setup \add_forward_prfstep @{thm run.intros(3)}\ lemma run_complete [resolve]: "\\' r. run c \ \' (r::'a)" @proof @obtain "r::'a" where "r = r" @case "\ = None" @with @have "run c None None r" @end @case "execute c (the \) = None" @with @have "run c \ None r" @end @qed lemma run_to_execute [forward]: "run c (Some h) \' r \ if \' = None then execute c h = None else execute c h = Some (r, the \')" @proof @case_induct "run c (Some h) \' r" @qed setup \add_rewrite_rule @{thm execute_bind(1)}\ lemma runE [forward]: "run f (Some h) (Some h') r' \ run (f \ g) (Some h) \ r \ run (g r') (Some h') \ r" by auto2 setup \add_rewrite_rule @{thm Array.get_alloc}\ setup \add_rewrite_rule @{thm Ref.get_alloc}\ setup \add_rewrite_rule_bidir @{thm Array.length_def}\ subsection \Definition of hoare triple, and the frame rule.\ definition new_addrs :: "heap \ addr set \ heap \ addr set" where [rewrite]: "new_addrs h as h' = as \ {a. lim h \ a \ a < lim h'}" definition hoare_triple :: "assn \ 'a Heap \ ('a \ assn) \ bool" ("<_>/ _/ <_>") where [rewrite]: "

c \ (\h as \ r. pHeap h as \ P \ run c (Some h) \ r \ (\ \ None \ pHeap (the \) (new_addrs h as (the \)) \ Q r \ relH {a . a < lim h \ a \ as} h (the \) \ lim h \ lim (the \)))" lemma hoare_tripleD [forward]: "

c \ run c (Some h) \ r \ \as. pHeap h as \ P \ (\ \ None \ pHeap (the \) (new_addrs h as (the \)) \ Q r \ relH {a . a < lim h \ a \ as} h (the \) \ lim h \ lim (the \))" by auto2 setup \del_prfstep_thm_eqforward @{thm hoare_triple_def}\ abbreviation hoare_triple' :: "assn \ 'r Heap \ ('r \ assn) \ bool" ("<_> _ <_>\<^sub>t") where "

c \<^sub>t \

c <\r. Q r * true>" theorem frame_rule [backward]: "

c \

c <\x. Q x * R>" @proof @have "\h as \ r. pHeap h as \ P * R \ run c (Some h) \ r \ (\ \ None \ pHeap (the \) (new_addrs h as (the \)) \ Q r * R \ relH {a . a < lim h \ a \ as} h (the \) \ lim h \ lim (the \))" @with @obtain as1 as2 where "as = as1 \ as2" "as1 \ as2 = {}" "pHeap h as1 \ P \ pHeap h as2 \ R" @have "relH as2 h (the \)" @have "new_addrs h as (the \) = new_addrs h as1 (the \) \ as2" @end @qed text \This is the last use of the definition of separating conjunction.\ setup \del_prfstep_thm @{thm mod_star_conv}\ theorem bind_rule: "

f \ \x. g x \

f \ g " @proof @have "\h as \ r. pHeap h as \ P \ run (f \ g) (Some h) \ r \ (\ \ None \ pHeap (the \) (new_addrs h as (the \)) \ R r \ relH {a . a < lim h \ a \ as} h (the \) \ lim h \ lim (the \))" @with \ \First step from h to h'\ @obtain \' r' where "run f (Some h) \' r'" @obtain h' where "\' = Some h'" @let "as' = new_addrs h as h'" @have "pHeap h' as' \ Q r'" \ \Second step from h' to h''\ @have "run (g r') (Some h') \ r" @obtain h'' where "\ = Some h''" @let "as'' = new_addrs h' as' h''" @have "pHeap h'' as'' \ R r" @have "as'' = new_addrs h as h''" @end @qed text \Actual statement used:\ lemma bind_rule': "

f \ \

f \ g \ \x. \ g x " using bind_rule by blast lemma pre_rule': "\

f \ P \\<^sub>A P' \ \ f " @proof @have "P * R \\<^sub>A P' * R" @qed lemma pre_rule'': "

f \ P' \\<^sub>A P * R \ f <\x. Q x * R>" @proof @have "

f <\x. Q x * R>" @qed lemma pre_ex_rule: "\ <\\<^sub>Ax. P x> f \ (\x. \

f )" by auto2 lemma pre_pure_rule: "\

b> f \ \

f \ b" by auto2 lemma pre_pure_rule': "\ <\b> f \ \ f \ b" by auto2 lemma post_rule: "

f \ \x. Q x \\<^sub>A R x \

f " by auto2 setup \fold del_prfstep_thm [@{thm entailsD}, @{thm entails_frame}, @{thm frame_rule}]\ text \Actual statement used:\ lemma post_rule': "

f \ \

f \ \x. \ (Q x \\<^sub>A R x)" using post_rule by blast lemma norm_pre_pure_iff: "

b> c \ (b \

c )" by auto2 lemma norm_pre_pure_iff2: "<\b> c \ (b \ c )" by auto2 subsection \Hoare triples for atomic commands\ text \First, those that do not modify the heap.\ setup \add_rewrite_rule @{thm execute_assert(1)}\ lemma assert_rule: "<\(R x)> assert R x <\r. \(r = x)>" by auto2 lemma execute_return' [rewrite]: "execute (return x) h = Some (x, h)" by (metis comp_eq_dest_lhs execute_return) lemma return_rule: " return x <\r. \(r = x)>" by auto2 setup \add_rewrite_rule @{thm execute_nth(1)}\ lemma nth_rule: "\<^sub>a xs * \(i < length xs)> Array.nth a i <\r. a \\<^sub>a xs * \(r = xs ! i)>" by auto2 setup \add_rewrite_rule @{thm execute_len}\ lemma length_rule: "\<^sub>a xs> Array.len a <\r. a \\<^sub>a xs * \(r = length xs)>" by auto2 setup \add_rewrite_rule @{thm execute_lookup}\ lemma lookup_rule: "

\<^sub>r x> !p <\r. p \\<^sub>r x * \(r = x)>" by auto2 setup \add_rewrite_rule @{thm execute_freeze}\ lemma freeze_rule: "\<^sub>a xs> Array.freeze a <\r. a \\<^sub>a xs * \(r = xs)>" by auto2 text \Next, the update rules.\ setup \add_rewrite_rule @{thm Ref.lim_set}\ lemma Array_lim_set [rewrite]: "lim (Array.set p xs h) = lim h" by (simp add: Array.set_def) setup \fold add_rewrite_rule [@{thm Ref.get_set_eq}, @{thm Array.get_set_eq}]\ setup \add_rewrite_rule @{thm Array.update_def}\ setup \add_rewrite_rule @{thm execute_upd(1)}\ lemma upd_rule: "\<^sub>a xs * \(i < length xs)> Array.upd i x a <\r. a \\<^sub>a list_update xs i x * \(r = a)>" by auto2 setup \add_rewrite_rule @{thm execute_update}\ lemma update_rule: "

\<^sub>r y> p := x <\r. p \\<^sub>r x>" by auto2 text \Finally, the allocation rules.\ lemma lim_set_gen [rewrite]: "lim (h\lim := l\) = l" by simp -lemma Array_alloc_def' [rewrite]: +lemma Array_alloc_def' [rewrite]: "Array.alloc xs h = (let l = lim h; r = Array l in (r, (Array.set r xs (h\lim := l + 1\))))" by (simp add: Array.alloc_def) setup \fold add_rewrite_rule [ @{thm addr_of_array.simps}, @{thm addr_of_ref.simps}, @{thm Ref.alloc_def}]\ lemma refs_on_Array_set [rewrite]: "refs (Array.set p xs h) t i = refs h t i" by (simp add: Array.set_def) lemma arrays_on_Ref_set [rewrite]: "arrays (Ref.set p x h) t i = arrays h t i" by (simp add: Ref.set_def) lemma refs_on_Array_alloc [rewrite]: "refs (snd (Array.alloc xs h)) t i = refs h t i" by (metis (no_types, lifting) Array.alloc_def refs_on_Array_set select_convs(2) snd_conv surjective update_convs(3)) lemma arrays_on_Ref_alloc [rewrite]: "arrays (snd (Ref.alloc x h)) t i = arrays h t i" by (metis (no_types, lifting) Ref.alloc_def arrays_on_Ref_set select_convs(1) sndI surjective update_convs(3)) lemma arrays_on_Array_alloc [rewrite]: "i < lim h \ arrays (snd (Array.alloc xs h)) t i = arrays h t i" by (smt Array.alloc_def Array.set_def addr_of_array.simps fun_upd_apply less_or_eq_imp_le linorder_not_less simps(1) snd_conv surjective update_convs(1) update_convs(3)) lemma refs_on_Ref_alloc [rewrite]: "i < lim h \ refs (snd (Ref.alloc x h)) t i = refs h t i" by (smt Ref.alloc_def Ref.set_def addr_of_ref.simps fun_upd_apply less_or_eq_imp_le linorder_not_less select_convs(2) simps(6) snd_conv surjective update_convs(3)) setup \add_rewrite_rule @{thm execute_new}\ lemma new_rule: " Array.new n x <\r. r \\<^sub>a replicate n x>" by auto2 setup \add_rewrite_rule @{thm execute_of_list}\ lemma of_list_rule: " Array.of_list xs <\r. r \\<^sub>a xs>" by auto2 setup \add_rewrite_rule @{thm execute_ref}\ lemma ref_rule: " ref x <\r. r \\<^sub>r x>" by auto2 setup \fold del_prfstep_thm [ @{thm sngr_assn_rule}, @{thm snga_assn_rule}, @{thm pure_assn_rule}, @{thm top_assn_rule}, @{thm mod_pure_star_dist}, @{thm one_assn_rule}, @{thm hoare_triple_def}, @{thm mod_ex_dist}]\ setup \del_simple_datatype "pheap"\ subsection \Definition of procedures\ text \ASCII abbreviations for ML files.\ abbreviation (input) ex_assn_ascii :: "('a \ assn) \ assn" (binder "EXA" 11) where "ex_assn_ascii \ ex_assn" abbreviation (input) models_ascii :: "pheap \ assn \ bool" (infix "|=" 50) where "h |= P \ h \ P" ML_file "sep_util.ML" ML \ structure AssnMatcher = AssnMatcher(SepUtil) structure SepLogic = SepLogic(SepUtil) val add_assn_matcher = AssnMatcher.add_assn_matcher val add_entail_matcher = AssnMatcher.add_entail_matcher val add_forward_ent_prfstep = SepLogic.add_forward_ent_prfstep val add_rewrite_ent_rule = SepLogic.add_rewrite_ent_rule val add_hoare_triple_prfstep = SepLogic.add_hoare_triple_prfstep \ setup \AssnMatcher.add_assn_matcher_proofsteps\ setup \SepLogic.add_sep_logic_proofsteps\ ML_file "sep_steps_test.ML" attribute_setup forward_ent = \setup_attrib add_forward_ent_prfstep\ attribute_setup rewrite_ent = \setup_attrib add_rewrite_ent_rule\ attribute_setup hoare_triple = \setup_attrib add_hoare_triple_prfstep\ setup \fold add_hoare_triple_prfstep [ @{thm assert_rule}, @{thm update_rule}, @{thm nth_rule}, @{thm upd_rule}, @{thm return_rule}, @{thm ref_rule}, @{thm lookup_rule}, @{thm new_rule}, @{thm of_list_rule}, @{thm length_rule}, @{thm freeze_rule}]\ text \Some simple tests\ theorem " ref x <\r. r \\<^sub>r x>" by auto2 theorem "\<^sub>r x> ref x <\r. a \\<^sub>r x * r \\<^sub>r x>" by auto2 theorem "\<^sub>r x> (!a) <\r. a \\<^sub>r x * \(r = x)>" by auto2 theorem "\<^sub>r x * b \\<^sub>r y> (!a) <\r. a \\<^sub>r x * b \\<^sub>r y * \(r = x)>" by auto2 theorem "\<^sub>r x * b \\<^sub>r y> (!b) <\r. a \\<^sub>r x * b \\<^sub>r y * \(r = y)>" by auto2 theorem "\<^sub>r x> do { a := y; !a } <\r. a \\<^sub>r y * \(r = y)>" by auto2 theorem "\<^sub>r x> do { a := y; a := z; !a } <\r. a \\<^sub>r z * \(r = z)>" by auto2 theorem "\<^sub>r x> do { y \ !a; ref y} <\r. a \\<^sub>r x * r \\<^sub>r x>" by auto2 theorem " return x <\r. \(r = x)>" by auto2 end diff --git a/thys/Auto2_Imperative_HOL/Imperative/assn_matcher.ML b/thys/Auto2_Imperative_HOL/Imperative/assn_matcher.ML --- a/thys/Auto2_Imperative_HOL/Imperative/assn_matcher.ML +++ b/thys/Auto2_Imperative_HOL/Imperative/assn_matcher.ML @@ -1,324 +1,325 @@ (* File: assn_matcher.ML Author: Bohua Zhan Matching of assertions. *) (* Given arguments ctxt (pat, t) (id, inst), match pat with t. Assume pat is not a product. Produce t ==> pat(s) or t ==> pat(s) * t' for those in AssnMatchData. Produce pat(s) or pat(s) * t' ==> t for those in AssnInvMatchData. *) type assn_matcher = Proof.context -> term * cterm -> id_inst -> id_inst_th list signature ASSN_MATCHER = sig val add_assn_matcher: assn_matcher -> theory -> theory val assn_match_term: Proof.context -> term * cterm -> id_inst -> id_inst_th list val assn_match_all: Proof.context -> term * cterm -> id_inst -> id_inst_th list val assn_match_strict: Proof.context -> term * cterm -> id_inst -> id_inst_th list val triv_assn_matcher: assn_matcher val emp_assn_matcher: assn_matcher val true_assn_matcher: assn_matcher val add_entail_matcher: thm -> theory -> theory val assn_match_single: Proof.context -> term * cterm -> id_inst -> id_inst_th list val add_assn_matcher_proofsteps: theory -> theory end (* Due to strange interaction between functors and Theory_Data, this must be put outside. *) structure MatchData = Theory_Data ( type T = (assn_matcher * serial) list val empty = [] val merge = Library.merge (eq_snd op =) ) functor AssnMatcher(SepUtil: SEP_UTIL): ASSN_MATCHER = struct open SepUtil +structure ACUtil = Auto2_HOL_Extra_Setup.ACUtil (* Matching in the forward direction *) fun add_assn_matcher matcher = MatchData.map (cons (matcher, serial ())) (* Assume pat is not in the form A * B. Match pat with one or more terms of t. Return theorem of form t ==> pat(s) * t'. *) fun assn_match_term ctxt (pat, ct) (id, inst) = let val _ = assert (not (Term.is_Var pat)) "assn_match_term: pat should not be Var." val thy = Proof_Context.theory_of ctxt fun apply_matcher matcher = matcher ctxt (pat, ct) (id, inst) (* th must be an entailment, and the right side must be pat(s), pat(s) * t', or t' * pat(s). *) fun process_res ((id, inst'), th) = let val _ = assert (is_entail (prop_of' th)) "assn_match_term" val (_, rhs) = th |> prop_of' |> dest_entail val exp_rhs = Util.subst_term_norm inst' pat in if rhs aconv exp_rhs then ((id, inst'), th |> apply_to_entail_r mult_emp_right) else if UtilArith.is_times rhs andalso dest_arg1 rhs aconv exp_rhs then ((id, inst'), th) else if UtilArith.is_times rhs andalso dest_arg rhs aconv exp_rhs then ((id, inst'), th |> apply_to_entail_r (ACUtil.comm_cv assn_ac_info)) else raise Fail "assn_match_term" end in (maps (apply_matcher o #1) (MatchData.get thy)) |> map process_res end (* Match each term of pat with some term in t. Returns t ==> pat(s) * t'. *) fun assn_match_all ctxt (pat, ct) (id, inst) = if UtilArith.is_times pat then let val (A, B) = Util.dest_binop_args pat val insts = assn_match_all ctxt (A, ct) (id, inst) (* th is t ==> A(s) * t'. Match B(s) with t', with result t' ==> B(s) * t''. Produce t ==> (A(s) * B(s)) * t'' *) fun process_inst ((id', inst'), th) = let val ct' = th |> cprop_of' |> cdest_entail |> snd |> Thm.dest_arg val B' = Util.subst_term_norm inst' B val insts' = assn_match_all ctxt (B', ct') (id', inst') (* th' is t' ==> B(s) * t''. *) fun process_inst' ((id'', inst''), th') = let val res = ([th, th'] MRS entails_trans2_th) |> apply_to_entail_r ( ACUtil.assoc_sym_cv assn_ac_info) in ((id'', inst''), res) end in map process_inst' insts' end in maps process_inst insts end else assn_match_term ctxt (pat, ct) (id, inst) (* Guarantees that every term in t is matched. Returns t ==> pat(s). *) fun assn_match_strict ctxt (pat, ct) (id, inst) = let val inst = assn_match_all ctxt (pat, ct) (id, inst) fun process_inst ((id', inst'), th) = let val rhs = th |> prop_of' |> dest_entail |> snd val _ = assert (UtilArith.is_times rhs andalso dest_arg1 rhs aconv Util.subst_term_norm inst' pat) "assn_match_strict" in if dest_arg rhs aconv emp then [((id', inst'), th |> apply_to_entail_r reduce_emp_right)] else [] end in maps process_inst inst end (* Specific assertion matchers *) (* Matcher using the theorem A ==> A. *) fun triv_assn_matcher ctxt (pat, ct) (id, inst) = if pat aconv emp then [] (* leave to emp_assn_matcher *) else let val cts = ACUtil.cdest_ac assn_ac_info ct fun match_i i = let val ct' = nth cts i - val insts = Matcher.rewrite_match ctxt (pat, ct') (id, inst) + val insts = Auto2_Setup.Matcher.rewrite_match ctxt (pat, ct') (id, inst) (* eq_th is of form pat(inst') == t'. *) fun process_inst ((id', inst'), eq_th) = let val th = entail_triv_th ctxt (Thm.term_of ct) val cv = Conv.every_conv [ ACUtil.move_outmost assn_ac_info (Thm.term_of ct'), ACUtil.ac_last_conv assn_ac_info (Conv.rewr_conv (meta_sym eq_th))] in ((id', inst'), th |> apply_to_entail_r cv) end in map process_inst insts end in maps match_i (0 upto (length cts - 1)) end (* Consider the case where pat = emp. Return t ==> emp * t. *) fun emp_assn_matcher ctxt (pat, ct) (id, inst) = if not (pat aconv emp) then [] else [((id, inst), ct |> Thm.term_of |> entail_triv_th ctxt |> apply_to_entail_r mult_emp_left)] (* If pat = true, match all of t. Return t ==> emp * true. *) fun true_assn_matcher ctxt (pat, ct) (id, inst) = if not (pat aconv assn_true) then [] else [((id, inst), ct |> Thm.term_of |> entail_true_th ctxt |> apply_to_entail_r mult_emp_left)] (* We now consider the case of generating a matcher from an entailment theorem of a particular form. Given an entailment A ==> B, where B is of the form f ?xs pat_r, where f is a constant, and pat_r may contain additional schematic variables. Attempt to find a term of form f xs r within t, for the input term r, by matching the pattern A. For each match, return the implication t ==> f xs r or t ==> t' * f xs r. This function serves as the first step of entail_matcher. *) fun entail_matcher' entail_th ctxt r ct id = let (* Match pat_r with r. *) val pat_r = entail_th |> prop_of' |> dest_entail |> snd |> dest_arg - val inst_r = Matcher.rewrite_match ctxt (pat_r, Thm.cterm_of ctxt r) (id, fo_init) + val inst_r = Auto2_Setup.Matcher.rewrite_match ctxt (pat_r, Thm.cterm_of ctxt r) (id, fo_init) (* For each match, recursively match the instantiated version of A (named pat here) with t. *) fun process_inst_r ((id', inst'), eq_th) = let val entail_th' = Util.subst_thm ctxt inst' entail_th val pat = entail_th' |> prop_of' |> dest_arg1 val matches = assn_match_all ctxt (pat, ct) (id', fo_init) (* th is of form t ==> pat(s) * t'. Convert to t ==> t' * pat(s). Then use entailment theorem to convert to t ==> t' * B. Finally, convert the argument in B to the given r. *) fun process_match ((id'', _), th) = let val cv = eq_th |> Conv.rewr_conv |> Util.argn_conv 1 |> ACUtil.ac_last_conv assn_ac_info val th' = th |> apply_to_entail_r (ACUtil.comm_cv assn_ac_info) in (id'', ([th', entail_th'] MRS entails_trans2_th) |> apply_to_entail_r cv) end in map process_match matches end in maps process_inst_r inst_r end (* Given entailment theorem A ==> B, with same condition as in entail_matcher', attempt to match pat with t, and return t ==> t' * pat(s). For any matching to be performed, pat must be in the form f pat_xs r, where pat_xs may contain schematic variables, but r cannot. First, find f xs r using entail_matcher', then match pat_xs with xs. *) fun entail_matcher entail_th ctxt (pat, ct) (id, inst) = let val (f, args) = Term.strip_comb pat val pat_f = entail_th |> prop_of' |> dest_entail |> snd |> Term.head_of in if not (Term.aconv_untyped (f, pat_f)) orelse Util.has_vars (nth args 1) then [] else let val (pat_xs, r) = the_pair args val matches = entail_matcher' entail_th ctxt r ct id fun process_res (id', th) = let val xs = th |> cprop_of' |> Thm.dest_arg |> ACUtil.cdest_ac assn_ac_info |> List.last |> Drule.strip_comb |> snd |> hd - val insts = Matcher.rewrite_match ctxt (pat_xs, xs) (id', inst) + val insts = Auto2_Setup.Matcher.rewrite_match ctxt (pat_xs, xs) (id', inst) fun process_inst ((id'', inst'), eq_th) = let val cv = eq_th |> meta_sym |> Conv.rewr_conv |> Conv.arg1_conv |> ACUtil.ac_last_conv assn_ac_info in ((id'', inst'), th |> apply_to_entail_r cv) end in map process_inst insts end in maps process_res matches end end fun add_entail_matcher th = let val (pat_f, pat_args) = th |> prop_of' |> dest_entail |> snd |> Term.strip_comb val _ = assert (length pat_args = 2 andalso Term.is_Const pat_f) "add_entail_matcher: th must be in form A ==> f ?xs pat_r." in add_assn_matcher (entail_matcher th) end (* Matching in the backward direction *) (* Given a pattern pat, write t in the form pat(inst) * t'. *) fun assn_match_single ctxt (pat, ct) (id, inst) = let val cts = ACUtil.cdest_ac assn_ac_info ct fun match_i i = let val ct' = nth cts i val t' = Thm.term_of ct' - val insts = Matcher.rewrite_match ctxt (pat, ct') (id, inst) + val insts = Auto2_Setup.Matcher.rewrite_match ctxt (pat, ct') (id, inst) (* eq_th is of form pat(inst) == t'. *) fun process_inst ((id', inst'), eq_th) = let val eq_th' = if length cts = 1 then eq_th |> meta_sym |> apply_to_rhs mult_emp_right else Conv.every_conv [ ACUtil.move_outmost assn_ac_info t', Conv.arg_conv (Conv.rewr_conv (meta_sym eq_th)), ACUtil.comm_cv assn_ac_info] ct in ((id', inst'), eq_th') end in map process_inst insts end in maps match_i (0 upto (length cts - 1)) end val add_assn_matcher_proofsteps = fold add_assn_matcher [ triv_assn_matcher, emp_assn_matcher, true_assn_matcher ] end (* structure AssnMatcher. *) diff --git a/thys/Auto2_Imperative_HOL/Imperative/list_matcher_test.ML b/thys/Auto2_Imperative_HOL/Imperative/list_matcher_test.ML --- a/thys/Auto2_Imperative_HOL/Imperative/list_matcher_test.ML +++ b/thys/Auto2_Imperative_HOL/Imperative/list_matcher_test.ML @@ -1,106 +1,106 @@ (* File: list_matcher_test.ML Author: Bohua Zhan Unit test for matching of assertions on linked lists. *) local open SepUtil val ctxt = @{context} fun eq_info_list (l1, l2) = length l1 = length l2 andalso eq_set (eq_pair (op =) (op aconv)) (l1, l2) fun print_term_info ctxt (id, t) = "(" ^ (BoxID.string_of_box_id id) ^ ", " ^ (Syntax.string_of_term ctxt t) ^ ")" fun print_term_infos ctxt lst = commas (map (print_term_info ctxt) lst) fun add_rewrite id (str1, str2) ctxt = let val (t1, t2) = (Syntax.read_term ctxt str1, Syntax.read_term ctxt str2) val thy = Proof_Context.theory_of ctxt val th = assume_eq thy (t1, t2) in - ctxt |> RewriteTable.add_rewrite (id, th) + ctxt |> Auto2_Setup.RewriteTable.add_rewrite (id, th) end val ts = [@{term "[p::nat node ref option, n1, s1, s2]"}, @{term "[pp::nat node ref, pp']"}, @{term "[x::nat]"}, @{term "[xs::nat list, ys, zs]"}] val ctxt' = ctxt |> fold Proof_Context.augment ts - |> RewriteTable.add_term ([], Thm.cterm_of ctxt (Free ("P", assnT))) |> snd + |> Auto2_Setup.RewriteTable.add_term ([], Thm.cterm_of ctxt (Free ("P", assnT))) |> snd |> add_rewrite [] ("n1", "None::nat node ref option") |> snd |> add_rewrite [] ("s1", "Some pp") |> snd |> add_rewrite [] ("s2", "Some pp'") |> snd |> add_rewrite [] ("zs", "x # xs") |> snd in fun test test_fun str (pat_str, t_str, info_str) = let val pat = Proof_Context.read_term_pattern ctxt' pat_str val t = Syntax.read_term ctxt' t_str val ct = Thm.cterm_of ctxt t val info = map (apsnd (Syntax.read_term ctxt')) info_str val info' = (test_fun ctxt' (pat, ct) ([], fo_init)) |> map (apsnd prop_of') val infol = info' |> map snd |> map dest_arg1 val infor = info' |> map (fn ((id, _), t) => (id, dest_arg t)) in if forall (fn t' => t' aconv t) infol andalso eq_info_list (info, infor) then () else let val _ = tracing ("Expected: " ^ print_term_infos ctxt' info) val _ = tracing ("Actual: " ^ print_term_infos ctxt' infor) in raise Fail str end end val test_assn_term_matcher = let val test_data = [ ("os_list ?xs n1", "emp", [([], "os_list [] n1 * emp")]), ("os_list ?xs n1", "sngr_assn a 1", [([], "os_list [] n1 * sngr_assn a 1")]), ("os_list ?xs s1", "sngr_assn pp (Node x r) * os_list xs r", [([], "os_list (x # xs) s1 * emp")]), ("os_list (?x # ?xs) s1", "sngr_assn pp (Node x r) * os_list xs r", [([], "os_list (x # xs) s1 * emp")]), ("os_list ys s1", "sngr_assn pp (Node x r) * os_list xs r", []), ("os_list zs s1", "sngr_assn pp (Node x r) * os_list xs r", [([], "os_list zs s1 * emp")]), ("os_list zs s1", "sngr_assn pp (Node x r) * os_list xs r * os_list ys p", [([], "os_list zs s1 * os_list ys p")]) ] in map (test AssnMatcher.assn_match_term "test_assn_term_matcher") test_data end val test_list_prop_matcher = let val test_data = [ ("os_list ?xs s1", "sngr_assn pp (Node x r) * os_list xs r", [([], "os_list (x # xs) s1")]), ("os_list zs s1", "sngr_assn pp (Node x r) * os_list xs r", [([], "os_list zs s1")]), ("os_list ?xs p * os_list ?ys q", "os_list ys q * os_list xs p", [([], "os_list xs p * os_list ys q")]), ("os_list ?xs p", "os_list xs p * os_list ys q", []), ("sngr_assn a 1", "sngr_assn a 1", [([], "sngr_assn a 1")]), ("sngr_assn ?a 1 * os_list ?ys q * os_list ?zs r", "os_list ys q * os_list zs r * sngr_assn a 1", [([], "sngr_assn a 1 * os_list ys q * os_list zs r")]) ] in map (test AssnMatcher.assn_match_strict "test_list_prop_matcher") test_data end end diff --git a/thys/Auto2_Imperative_HOL/Imperative/sep_steps.ML b/thys/Auto2_Imperative_HOL/Imperative/sep_steps.ML --- a/thys/Auto2_Imperative_HOL/Imperative/sep_steps.ML +++ b/thys/Auto2_Imperative_HOL/Imperative/sep_steps.ML @@ -1,824 +1,824 @@ (* File: sep_steps.ML Author: Bohua Zhan Proof steps for separation logic. *) signature SEP_LOGIC = sig val normalize_hoare_goal_cv: Proof.context -> conv val normalize_entail_goal_cv: Proof.context -> conv val get_proc_def: theory -> term -> thm list val update_hoare_triple: thm -> theory -> theory val get_hoare_triples: theory -> string -> thm list val is_bind_cmd: term -> bool val get_first_cmd: term -> term val TY_CODE_POS: string val TY_ENTAIL: string val is_neg_hoare_triple: term -> bool val is_neg_entail: term -> bool val norm_precond: Proof.context -> conv val norm_entail_conds: Proof.context -> conv val is_implies_item: box_item -> bool val hoare_goal_update: Proof.context -> box_id * thm -> raw_update val entail_goal_update: Proof.context -> box_id * thm -> raw_update val init_entail: proofstep val entails_resolve: proofstep val init_pos: proofstep val add_forward_ent_prfstep: thm -> theory -> theory val add_backward_ent_prfstep: thm -> theory -> theory val add_rewrite_ent_rule: thm -> theory -> theory val rewrite_pos: proofstep val extract_pure_hoare_cv: conv val match_hoare_th: box_id -> Proof.context -> thm -> thm -> box_item -> raw_update list val match_hoare_prop: proofstep val match_hoare_disj: proofstep val match_assn_pure: proofstep val hoare_create_case: proofstep val entail_pure: proofstep val entail_create_case: proofstep val hoare_triple: proofstep val contract_hoare_cv: Proof.context -> conv val add_hoare_triple_prfstep: thm -> theory -> theory val add_sep_logic_proofsteps: theory -> theory end; functor SepLogic(SepUtil: SEP_UTIL) : SEP_LOGIC = struct open SepUtil structure AssnMatcher = AssnMatcher(SepUtil) (* Normalize a Hoare triple goal. *) fun normalize_hoare_goal_cv' ctxt ct = let val t = Thm.term_of ct val (P, _, _) = t |> dest_not |> dest_hoare_triple in if is_pure_assn P then rewr_obj_eq pre_pure_rule_th' ct else if UtilArith.is_times P andalso is_pure_assn (dest_arg P) then Conv.every_conv [rewr_obj_eq pre_pure_rule_th, Conv.arg1_conv (normalize_hoare_goal_cv' ctxt)] ct else if is_ex_assn P then Conv.every_conv [ rewr_obj_eq pre_ex_rule_th, Conv.binder_conv (normalize_hoare_goal_cv' o snd) ctxt] ct else Conv.all_conv ct end fun normalize_hoare_goal_cv ctxt ct = if is_ex (Thm.term_of ct) then Conv.binder_conv (normalize_hoare_goal_cv o snd) ctxt ct else Conv.every_conv [ Conv.arg_conv (Util.argn_conv 0 (SepUtil.normalize_assn_cv ctxt)), normalize_hoare_goal_cv' ctxt] ct fun normalize_entail_goal_cv' ctxt ct = let val t = Thm.term_of ct val (P, _) = t |> dest_not |> dest_entail in if is_pure_assn P then rewr_obj_eq entails_pure_th' ct else if UtilArith.is_times P andalso is_pure_assn (dest_arg P) then Conv.every_conv [rewr_obj_eq entails_pure_th, Conv.arg1_conv (normalize_entail_goal_cv' ctxt)] ct else if is_ex_assn P then Conv.every_conv [ rewr_obj_eq entails_ex_th, Conv.binder_conv (normalize_entail_goal_cv' o snd) ctxt] ct else Conv.all_conv ct end fun normalize_entail_goal_cv ctxt ct = if is_ex (Thm.term_of ct) then Conv.binder_conv (normalize_entail_goal_cv o snd) ctxt ct else Conv.every_conv [ Conv.arg_conv (Conv.binop_conv (SepUtil.normalize_assn_cv ctxt)), normalize_entail_goal_cv' ctxt] ct (* Data maintained for each imperative command. *) structure Data = Theory_Data ( type T = thm Symtab.table; val empty = Symtab.empty; val merge = Symtab.merge (op Thm.eq_thm_prop) ) (* Add the given theorem as a Hoare triple. Replace previous Hoare triples for this theorem if any exist. *) fun update_hoare_triple hoare_th thy = let val (_, c, _) = dest_hoare_triple (prop_of' hoare_th) val nm = Util.get_head_name c in thy |> Data.map (Symtab.update (nm, hoare_th)) end (* Obtain list of Hoare triples for the given command *) fun get_hoare_triples thy nm = the_list (Symtab.lookup (Data.get thy) nm) fun get_proc_def thy t = if is_bind_cmd t then [] else let val nm = Util.get_head_name t in if null (get_hoare_triples thy nm) then - Unfolding.get_unfold_thms_by_name thy nm + Auto2_HOL_Extra_Setup.Unfolding.get_unfold_thms_by_name thy nm else [] end fun get_first_cmd c = if is_bind_cmd c then dest_arg1 c else c fun extract_return_var t = if is_bind_cmd t then case dest_arg t of Abs (x, T, _) => if x = "uu_" then Free ("u",T) (* no assigned name *) else Free (x,T) (* regular assigned name *) | c => Free ("r", Term.domain_type (fastype_of c)) else raise Fail "extract_return_var" (* CODE_POS item stores a Hoare triple goal, indicating the current position in the program. *) val TY_CODE_POS = "CODE_POS" (* ENTAIL item stores an entailment goal, usually indicating the end of the program. *) val TY_ENTAIL = "ENTAIL" fun is_neg_hoare_triple t = is_neg t andalso is_hoare_triple (dest_not t) fun is_neg_entail t = is_neg t andalso is_entail (dest_not t) fun norm_precond ctxt ct = Util.argn_conv 0 (SepUtil.normalize_assn_cv ctxt) ct fun norm_entail_conds ctxt ct = Conv.binop_conv (SepUtil.normalize_assn_cv ctxt) ct fun is_implies_item item = Util.is_implies (Thm.prop_of (#prop item)) fun normalize_let ctxt th = let val rewr_one = Conv.first_conv [Conv.rewr_conv @{thm Let_def}, rewr_obj_eq @{thm case_prod_beta'}] val cv = Conv.every_conv [ Conv.top_conv (K (Conv.try_conv rewr_one)) ctxt, Thm.beta_conversion true] in apply_to_thm' cv th end fun hoare_goal_update ctxt (id, th) = if Util.is_implies (Thm.prop_of th) then AddItems {id = id, sc = SOME 1, raw_items = [Fact (TY_CODE_POS, [Thm.prop_of th], th)]} else let val (ex_ritems, res_th) = th |> apply_to_thm' (normalize_hoare_goal_cv ctxt) - |> Update.apply_exists_ritems ctxt + |> Auto2_Setup.Update.apply_exists_ritems ctxt val (res_th', rest) = - res_th |> UtilLogic.split_conj_gen_th + res_th |> Auto2_Setup.UtilLogic.split_conj_gen_th |> filter_split (is_neg_hoare_triple o prop_of') val _ = assert (length res_th' = 1) "hoare_goal_update" val res_th' = res_th' |> the_single |> apply_to_thm' (Conv.arg_conv (norm_precond ctxt)) |> normalize_let ctxt - val ritems = ex_ritems @ map Update.thm_to_ritem rest @ + val ritems = ex_ritems @ map Auto2_Setup.Update.thm_to_ritem rest @ [Fact (TY_CODE_POS, [prop_of' res_th'], res_th')] in AddItems {id = id, sc = SOME 1, raw_items = ritems} end fun entail_goal_update ctxt (id, th) = if Util.is_implies (Thm.prop_of th) then AddItems {id = id, sc = SOME 1, raw_items = [Fact (TY_ENTAIL, [Thm.prop_of th], th)]} else let val (ex_ritems, res_th) = th |> apply_to_thm' (normalize_entail_goal_cv ctxt) - |> Update.apply_exists_ritems ctxt + |> Auto2_Setup.Update.apply_exists_ritems ctxt val (res_th', rest) = - res_th |> UtilLogic.split_conj_gen_th + res_th |> Auto2_Setup.UtilLogic.split_conj_gen_th |> filter_split (is_neg_entail o prop_of') val _ = assert (length res_th' = 1) "entail_goal_update" val res_th' = res_th' |> the_single |> apply_to_thm' (Conv.arg_conv (norm_entail_conds ctxt)) |> normalize_let ctxt - val ritems = ex_ritems @ map Update.thm_to_ritem rest @ + val ritems = ex_ritems @ map Auto2_Setup.Update.thm_to_ritem rest @ [Fact (TY_ENTAIL, [prop_of' res_th'], res_th')] in AddItems {id = id, sc = SOME 1, raw_items = ritems} end fun init_entail_fn ctxt item = if not (BoxID.has_incr_id (#id item)) then [] else let val {id, prop, ...} = item in [entail_goal_update ctxt (id, prop)] end val init_entail = {name = "sep.init_entail", args = [TypedMatch (TY_PROP, get_neg (entail_t $ Var (("A",0), assnT) $ Var (("B",0), assnT)))], func = OneStep init_entail_fn} (* Apply entailment to the pre-condition P of P ==>_A Q. *) fun forward_ent_prfstep_fn ent_th ctxt item = if is_implies_item item then [] else let val (A, _) = dest_entail (prop_of' ent_th) val {id, prop, ...} = item val (P, _) = prop |> prop_of' |> dest_not |> dest_entail val cP = Thm.cterm_of ctxt P val insts = (AssnMatcher.assn_match_single ctxt (A, cP) (id, fo_init)) |> filter (BoxID.has_incr_id o fst o fst) fun process_inst ((id', _), eq_th) = (* eq_th is P == pat(inst) * P' *) let val prop' = prop |> apply_to_thm' ( Conv.arg_conv (Conv.arg1_conv (Conv.rewr_conv eq_th))) val prop'' = [prop', ent_th] MRS entails_frame_th' in [entail_goal_update ctxt (id', prop''), ShadowItem {id = id', item = item}] end in if null insts then [] else process_inst (hd insts) end fun forward_ent_prfstep ent_th = {name = Util.name_of_thm ent_th ^ "@ent", args = [TypedUniv TY_ENTAIL], func = OneStep (forward_ent_prfstep_fn ent_th)} fun backward_ent_prfstep_fn ent_th ctxt item = if is_implies_item item then [] else let val (_, B) = dest_entail (prop_of' ent_th) val {id, prop, ...} = item val (_, Q) = prop |> prop_of' |> dest_not |> dest_entail val cQ = Thm.cterm_of ctxt Q val insts = (AssnMatcher.assn_match_single ctxt (B, cQ) (id, fo_init)) |> filter (BoxID.has_incr_id o fst o fst) fun process_inst ((id', _), eq_th) = (* eq_th is P == pat(inst) * P' *) let val prop' = prop |> apply_to_thm' ( Conv.arg_conv (Conv.arg_conv (Conv.rewr_conv eq_th))) val prop'' = [prop', ent_th] MRS entails_frame_th'' in [entail_goal_update ctxt (id', prop''), ShadowItem {id = id', item = item}] end in if null insts then [] else process_inst (hd insts) end fun backward_ent_prfstep ent_th = {name = Util.name_of_thm ent_th ^ "@entback", args = [TypedUniv TY_ENTAIL], func = OneStep (backward_ent_prfstep_fn ent_th)} fun group_pure_assn ct = let val t = Thm.term_of ct in if is_pure_assn t then mult_emp_left ct else if UtilArith.is_times t then if has_pure_assn (dest_arg1 t) then Conv.every_conv [ Conv.arg1_conv group_pure_assn, - ACUtil.assoc_cv assn_ac_info, + Auto2_HOL_Extra_Setup.ACUtil.assoc_cv assn_ac_info, Conv.arg_conv (rewr_obj_eq (obj_sym pure_conj_th))] ct else Conv.all_conv ct else Conv.all_conv ct end fun make_sch_th ctxt th = case Thm.prop_of th of Const (@{const_name Pure.all}, _) $ Abs (nm, T, _) => let val var = Var ((nm,0),T) in Thm.forall_elim (Thm.cterm_of ctxt var) th end | _ => raise Fail "make_sch_th" fun entails_norm_ex ctxt th = let val t = prop_of' th val (_, Q) = t |> dest_not |> dest_entail in if is_ex_assn Q then (th RS entails_ex_post_th) - |> apply_to_thm (UtilLogic.to_meta_conv ctxt) + |> apply_to_thm (Auto2_Setup.UtilLogic.to_meta_conv ctxt) |> make_sch_th ctxt |> entails_norm_ex ctxt else th end (* Solve an entailment. *) fun entails_resolve_fn ctxt item = if is_implies_item item then [] else let val {id, prop, ...} = item val prop = entails_norm_ex ctxt prop val (P, Q) = dest_entail (get_neg (prop_of' prop)) val cP = Thm.cterm_of ctxt P val Q' = strip_pure_assn Q val insts = (AssnMatcher.assn_match_strict ctxt (Q', cP) (id, fo_init)) |> filter (BoxID.has_incr_id o fst o fst) fun process_inst ((id', _), mod_th) = if has_pure_assn Q then let val prop' = prop |> apply_to_thm' ( Conv.arg_conv (Conv.arg_conv group_pure_assn)) val res = [prop', mod_th] MRS entails_pure_post_th in - Update.thm_update (id', res) + Auto2_Setup.Update.thm_update (id', res) end else - Update.thm_update (id', [prop, mod_th] MRS @{thm contra_triv}) + Auto2_Setup.Update.thm_update (id', [prop, mod_th] MRS @{thm contra_triv}) in map process_inst insts end val entails_resolve = {name = "sep.entails_resolve", args = [TypedUniv TY_ENTAIL], func = OneStep entails_resolve_fn} (* Initialize CODE_POS item from a Hoare triple goal. *) fun init_pos_fn ctxt item = let val {id, prop, ...} = item val thy = Proof_Context.theory_of ctxt val (_, c, _) = dest_hoare_triple (get_neg (prop_of' prop)) val proc_defs = get_proc_def thy c fun process_proc_def proc_def = let val (lhs, _) = proc_def |> prop_of' |> dest_eq val cc = Thm.cterm_of ctxt c - val insts = (Matcher.rewrite_match ctxt (lhs, cc) (id, fo_init)) + val insts = (Auto2_Setup.Matcher.rewrite_match ctxt (lhs, cc) (id, fo_init)) |> filter (BoxID.has_incr_id o fst o fst) fun process_inst ((id', _), eq_th) = let val cv = Conv.every_conv [ Conv.rewr_conv (meta_sym eq_th), rewr_obj_eq proc_def] val th = apply_to_thm' (Conv.arg_conv (Conv.arg1_conv cv)) prop in hoare_goal_update ctxt (id', th) end in map process_inst insts end in if null proc_defs then if BoxID.has_incr_id id then [hoare_goal_update ctxt (id, prop)] else [] else maps process_proc_def proc_defs end val init_pos = {name = "sep.init_pos", args = [TypedMatch (TY_PROP, get_neg hoare_triple_pat)], func = OneStep init_pos_fn} fun forward_hoare_prfstep_fn ent_th ctxt item = if is_implies_item item then [] else let val (A, _) = dest_entail (prop_of' ent_th) val {id, prop, ...} = item val (P, _, _) = prop |> prop_of' |> dest_not |> dest_hoare_triple val cP = Thm.cterm_of ctxt P val insts = (AssnMatcher.assn_match_single ctxt (A, cP) (id, fo_init)) |> filter (BoxID.has_incr_id o fst o fst) fun process_inst ((id', _), eq_th) = let val prop' = prop |> apply_to_thm' ( Conv.arg_conv (Util.argn_conv 0 (Conv.rewr_conv eq_th))) val prop'' = [prop', ent_th] MRS pre_rule_th' in [hoare_goal_update ctxt (id', prop''), ShadowItem {id = id', item = item}] end in if null insts then [] else process_inst (hd insts) end fun forward_hoare_prfstep ent_th = {name = Util.name_of_thm ent_th ^ "@hoare_ent", args = [TypedUniv TY_CODE_POS], func = OneStep (forward_hoare_prfstep_fn ent_th)} fun add_forward_ent_prfstep ent_th thy = let val name = Util.name_of_thm ent_th val ctxt = Proof_Context.init_global thy val _ = writeln ("Add forward entailment " ^ name ^ "\n" ^ Syntax.string_of_term ctxt (prop_of' ent_th)) in thy |> add_prfstep (forward_ent_prfstep ent_th) |> add_prfstep (forward_hoare_prfstep ent_th) end fun add_backward_ent_prfstep ent_th thy = let val name = Util.name_of_thm ent_th val ctxt = Proof_Context.init_global thy val _ = writeln ("Add backward entailment " ^ name ^ "\n" ^ Syntax.string_of_term ctxt (prop_of' ent_th)) in add_prfstep (backward_ent_prfstep ent_th) thy end fun add_rewrite_ent_rule th thy = let val forward_th = (th RS entails_equiv_forward_th) |> Drule.zero_var_indexes |> Util.update_name_of_thm th "@forward" val backward_th = (th RS entails_equiv_backward_th) |> Drule.zero_var_indexes |> Util.update_name_of_thm th "@backward" in thy |> add_forward_ent_prfstep forward_th |> add_backward_ent_prfstep backward_th end (* Rewrite the first command of a Hoare triple. *) fun rewr_first_cmd eq_th ct = let val (_, c, _) = ct |> Thm.term_of |> dest_hoare_triple in if is_bind_cmd c then Conv.arg1_conv (Conv.arg1_conv (rewr_obj_eq eq_th)) ct else Conv.arg1_conv (rewr_obj_eq eq_th) ct end (* Apply rewrite to the first command in CODE_POS. *) fun rewrite_pos_fn ctxt item1 item2 = if is_implies_item item1 then [] else let val {id = id1, prop = th, ...} = item1 val {id = id2, prop = eq_th, ...} = item2 val (_, c, _) = th |> prop_of' |> dest_not |> dest_hoare_triple val c' = get_first_cmd c val (c1, _) = eq_th |> prop_of' |> dest_eq val id = BoxID.merge_boxes ctxt (id1, id2) in if not (BoxID.has_incr_id id) then [] else if c1 aconv c' then let val th' = th |> apply_to_thm' (Conv.arg_conv (rewr_first_cmd eq_th)) in [hoare_goal_update ctxt (id, th'), ShadowItem {id = id, item = item1}] end else [] end val rewrite_pos = {name = "sep.rewrite_pos", args = [TypedUniv TY_CODE_POS, TypedMatch (TY_EQ, heap_eq_pat)], func = TwoStep rewrite_pos_fn} (* Extract the pure pre-conditions from a Hoare triple fact. *) fun extract_pure_hoare_cv ct = let val (P, _, _) = ct |> Thm.term_of |> dest_hoare_triple in if is_pure_assn P then rewr_obj_eq norm_pre_pure_iff2_th ct else if UtilArith.is_times P andalso is_pure_assn (dest_arg P) then Conv.every_conv [rewr_obj_eq norm_pre_pure_iff_th, Conv.arg_conv extract_pure_hoare_cv] ct else Conv.all_conv ct end (* Use a Hoare triple to advance a step in CODE_POS. *) fun match_hoare_th id ctxt hoare_th goal item = let val (P, c, _) = goal |> prop_of' |> dest_not |> dest_hoare_triple val c' = get_first_cmd c val (P', pat, _) = dest_hoare_triple (prop_of' hoare_th) val cc = Thm.cterm_of ctxt c' - val insts = Matcher.rewrite_match ctxt (pat, cc) (id, fo_init) + val insts = Auto2_Setup.Matcher.rewrite_match ctxt (pat, cc) (id, fo_init) fun process_inst ((id', inst), eq_th) = let val P'' = P' |> strip_pure_assn |> Util.subst_term_norm inst val cP = Thm.cterm_of ctxt P val insts' = (AssnMatcher.assn_match_all ctxt (P'', cP) (id', inst)) |> filter (BoxID.has_incr_id o fst o fst) fun process_inst' ((id'', inst'), ent_th) = let val hoare_th = (Util.subst_thm ctxt inst' hoare_th) |> apply_to_thm' (Conv.arg1_conv (Conv.rewr_conv eq_th)) |> apply_to_thm' extract_pure_hoare_cv - |> apply_to_thm (UtilLogic.to_meta_conv ctxt) + |> apply_to_thm (Auto2_Setup.UtilLogic.to_meta_conv ctxt) val hoare_th' = [hoare_th, ent_th] MRS pre_rule_th'' in if is_bind_cmd c then let val return_var = extract_return_var c val th' = [hoare_th', goal] MRS bind_rule_th' val (_, (assms, concl)) = th' |> Thm.prop_of |> Util.strip_meta_horn val concl' = concl |> dest_Trueprop |> Util.rename_abs_term [return_var] |> mk_Trueprop val t' = Util.list_meta_horn ([], (assms, concl')) val th' = Thm.renamed_prop t' th' in [hoare_goal_update ctxt (id'', th'), ShadowItem {id = id'', item = item}] end else [entail_goal_update ctxt (id'', [hoare_th', goal] MRS post_rule_th'), ShadowItem {id = id'', item = item}] end in if null insts' then [] else process_inst' (hd insts') end in if null insts then [] else process_inst (hd insts) end (* Match with PROP or DISJ items that are Hoare triples. In this function, we assume item1 is a Hoare triple (and item2 is the CODE_POS item). *) fun match_hoare_prop_fn ctxt item1 item2 = if is_implies_item item2 then [] else let val {id = id1, prop = hoare_th, ...} = item1 val {id = id2, prop = goal, ...} = item2 val id = BoxID.merge_boxes ctxt (id1, id2) in match_hoare_th id ctxt hoare_th goal item2 end val match_hoare_prop = {name = "sep.match_hoare_prop", args = [TypedMatch (TY_PROP, hoare_triple_pat), TypedUniv TY_CODE_POS], func = TwoStep match_hoare_prop_fn} (* For DISJ items, check that it is a Hoare triple. *) fun match_hoare_disj_fn ctxt item1 item2 = if is_implies_item item2 then [] else let val {tname, ...} = item1 - val (_, csubs) = Logic_ProofSteps.dest_tname_of_disj tname + val (_, csubs) = Auto2_Setup.Logic_ProofSteps.dest_tname_of_disj tname val subs = map Thm.term_of csubs in if length subs > 1 then [] else if not (is_hoare_triple (the_single subs)) then [] else match_hoare_prop_fn ctxt item1 item2 end val match_hoare_disj = {name = "sep.match_hoare_disj", - args = [TypedUniv TY_DISJ, TypedUniv TY_CODE_POS], + args = [TypedUniv Auto2_Setup.Logic_ProofSteps.TY_DISJ, TypedUniv TY_CODE_POS], func = TwoStep match_hoare_disj_fn} (* Match a MATCH_POS item with hoare triple

(b)> c with proposition b, resulting in a new MATCH_POS item (shadowing the original one) with hoare triple

c . Only work in the case where there are no schematic variables in b. *) fun match_assn_pure_fn ctxt item1 item2 = let val {id, prop, ...} = item1 in if Util.is_implies (Thm.prop_of prop) then let val (A, _) = Logic.dest_implies (Thm.prop_of prop) val pat = PropMatch (dest_Trueprop A) - val insts = (ItemIO.match_arg ctxt pat item2 (id, fo_init)) + val insts = (Auto2_Setup.ItemIO.match_arg ctxt pat item2 (id, fo_init)) |> filter (BoxID.has_incr_id o fst o fst) fun process_inst ((id', _), th) = [hoare_goal_update ctxt (id', th RS prop), ShadowItem {id = id', item = item1}] in maps process_inst insts end else [] end val match_assn_pure = {name = "sep.match_assn_pure", args = [TypedUniv TY_CODE_POS, PropMatch (@{term_pat "?b::bool"})], func = TwoStep match_assn_pure_fn} fun hoare_create_case_fn _ item = let val {id, prop, ...} = item in if not (BoxID.has_incr_id id) then [] else if Util.is_implies (Thm.prop_of prop) then let val (A, _) = Logic.dest_implies (Thm.prop_of prop) in [AddBoxes {id = id, sc = SOME 1, init_assum = get_neg' A}] end else [] end val hoare_create_case = {name = "sep.hoare_create_case", args = [TypedUniv TY_CODE_POS], func = OneStep hoare_create_case_fn} fun entail_pure_fn ctxt item1 item2 = let val {id, prop, ...} = item1 in if Util.is_implies (Thm.prop_of prop) then let val (A, _) = Logic.dest_implies (Thm.prop_of prop) val pat = PropMatch (dest_Trueprop A) - val insts = (ItemIO.match_arg ctxt pat item2 (id, fo_init)) + val insts = (Auto2_Setup.ItemIO.match_arg ctxt pat item2 (id, fo_init)) |> filter (BoxID.has_incr_id o fst o fst) fun process_inst ((id', _), th) = [entail_goal_update ctxt (id', th RS prop), ShadowItem {id = id', item = item1}] in maps process_inst insts end else [] end val entail_pure = {name = "sep.entail_pure", args = [TypedUniv TY_ENTAIL, PropMatch (@{term_pat "?b::bool"})], func = TwoStep entail_pure_fn} fun entail_create_case_fn _ item = let val {id, prop, ...} = item in if not (BoxID.has_incr_id id) then [] else if Util.is_implies (Thm.prop_of prop) then let val (A, _) = Logic.dest_implies (Thm.prop_of prop) in [AddBoxes {id = id, sc = SOME 1, init_assum = get_neg' A}] end else [] end val entail_create_case = {name = "sep.entail_create_case", args = [TypedUniv TY_ENTAIL], func = OneStep entail_create_case_fn} (* Matching CODE_POS with an existing Hoare triple. *) fun hoare_triple_fn ctxt item = if is_implies_item item then [] else let val thy = Proof_Context.theory_of ctxt val {id, prop = goal, ...} = item val (_, c, _) = goal |> prop_of' |> dest_not |> dest_hoare_triple val hoare_ths = c |> get_first_cmd |> Util.get_head_name |> get_hoare_triples thy in maps (fn hoare_th => match_hoare_th id ctxt hoare_th goal item) hoare_ths end val hoare_triple = {name = "sep.hoare_triple", args = [TypedUniv TY_CODE_POS], func = OneStep hoare_triple_fn} (* Contract a Hoare triple to

c form. *) fun contract_hoare_cv' ct = if is_imp (Thm.term_of ct) then Conv.every_conv [Conv.arg_conv contract_hoare_cv', rewr_obj_eq (obj_sym norm_pre_pure_iff_th)] ct else Conv.all_conv ct fun contract_hoare_cv ctxt ct = Conv.every_conv [contract_hoare_cv', norm_precond ctxt] ct (* Given hoare_th of the form ?c , produce proofstep matching item1 with CODE_POS (?h, ?c) and item2 with proposition ?h |= ?P * ?Ru. *) fun add_hoare_triple_prfstep hoare_th thy = let val name = Util.name_of_thm hoare_th val ctxt = Proof_Context.init_global thy val hoare_th' = - hoare_th |> apply_to_thm (UtilLogic.to_obj_conv ctxt) + hoare_th |> apply_to_thm (Auto2_Setup.UtilLogic.to_obj_conv ctxt) |> apply_to_thm' (contract_hoare_cv ctxt) |> Util.update_name_of_thm hoare_th "" val _ = writeln ("Add Hoare triple " ^ name ^ "\n" ^ Syntax.string_of_term ctxt (prop_of' hoare_th')) in thy |> update_hoare_triple hoare_th' end fun code_pos_terms ts = let val t = the_single ts in if fastype_of t = propT then [] else let val (P, c, _) = t |> dest_not |> dest_hoare_triple in SepUtil.assn_rewr_terms P @ [get_first_cmd c] end end fun entail_terms ts = let val t = the_single ts in if fastype_of t = propT then [] else let val (P, Q) = t |> dest_not |> dest_entail in maps SepUtil.assn_rewr_terms [P, Q] end end val add_sep_logic_proofsteps = - fold ItemIO.add_item_type [ + fold Auto2_Setup.ItemIO.add_item_type [ (TY_CODE_POS, SOME code_pos_terms, NONE, NONE), (TY_ENTAIL, SOME entail_terms, NONE, NONE) ] #> fold add_prfstep [ init_entail, entails_resolve, init_pos, rewrite_pos, match_assn_pure, hoare_triple, match_hoare_disj, match_hoare_prop, hoare_create_case, entail_pure, entail_create_case ] end (* structure SepLogic *) diff --git a/thys/Auto2_Imperative_HOL/Imperative/sep_util.ML b/thys/Auto2_Imperative_HOL/Imperative/sep_util.ML --- a/thys/Auto2_Imperative_HOL/Imperative/sep_util.ML +++ b/thys/Auto2_Imperative_HOL/Imperative/sep_util.ML @@ -1,203 +1,205 @@ (* File: sep_util.ML Author: Bohua Zhan Utility functions for separation logic. Implements the interface defined in sep_util_base.ML. *) structure SepUtil : SEP_UTIL = struct +structure ACUtil = Auto2_HOL_Extra_Setup.ACUtil + val assnT = @{typ assn} val emp = @{term emp} val assn_true = @{term true} val assn_ac_info = Nat_Util.times_ac_on_typ @{theory} assnT fun is_true_assn t = case t of Const (@{const_name top_assn}, _) => true | _ => false val entail_t = @{term entails} fun is_entail t = case t of Const (@{const_name entails}, _) $ _ $ _ => true | _ => false (* Deconstruct A ==>_A B into (A, B). *) fun dest_entail t = case t of Const (@{const_name entails}, _) $ A $ B => (A, B) | _ => raise Fail "dest_entail: unexpected t." fun cdest_entail ct = case Thm.term_of ct of Const (@{const_name entails}, _) $ _ $ _ => (Thm.dest_arg1 ct, Thm.dest_arg ct) | _ => raise Fail "dest_entail: unexpected t." fun is_ex_assn t = case t of Const (@{const_name ex_assn}, _) $ _ => true | _ => false (* Whether t is of the form \(b). *) fun is_pure_assn t = case t of Const (@{const_name pure_assn}, _) $ _ => true | _ => false (* Given t of form t1 * ... * tn, check whether any of them is of the form \(b). *) fun has_pure_assn t = exists is_pure_assn (ACUtil.dest_ac assn_ac_info t) (* Given t of form t1 * ... * tn, remove those ti that are pure assertions and return the product of the remaining terms. *) fun strip_pure_assn t = if UtilArith.is_times t andalso is_pure_assn (dest_arg t) then strip_pure_assn (dest_arg1 t) else if is_pure_assn t then emp else t val hoare_triple_pat = @{term_pat " ?c "} val heap_eq_pat = @{term_pat "(?c1::?'a Heap) = ?c2"} fun is_hoare_triple t = case t of Const (@{const_name hoare_triple}, _) $ _ $ _ $ _ => true | _ => false fun dest_hoare_triple t = case t of Const (@{const_name hoare_triple}, _) $ P $ c $ Q => (P, c, Q) | _ => raise Fail "dest_hoare_triple" fun is_bind_cmd c = case c of Const (@{const_name bind}, _) $ _ $ _ => true | _ => false (* Convert A to emp * A *) val mult_emp_left = rewr_obj_eq (obj_sym @{thm mult_1}) (* Convert A to A * emp *) val mult_emp_right = rewr_obj_eq (obj_sym @{thm mult_1_right}) (* Convert A * emp to A *) val reduce_emp_right = rewr_obj_eq @{thm mult_1_right} (* Given A of type assnT, return the theorem A ==> A. *) fun entail_triv_th ctxt A = let val thy = Proof_Context.theory_of ctxt val inst = Pattern.first_order_match thy (Var (("A", 0), assnT), A) fo_init in Util.subst_thm ctxt inst @{thm entails_triv} end (* Given A of type assnT, return the theorem A ==> true. *) fun entail_true_th ctxt A = let val thy = Proof_Context.theory_of ctxt val inst = Pattern.first_order_match thy (Var (("A", 0), assnT), A) fo_init in Util.subst_thm ctxt inst @{thm entails_true} end (* Given theorem A ==> B and a conversion cv, apply cv to B *) val apply_to_entail_r = apply_to_thm' o Conv.arg_conv val pre_pure_rule_th = @{thm pre_pure_rule} val pre_pure_rule_th' = @{thm pre_pure_rule'} val pre_ex_rule_th = @{thm pre_ex_rule} val entails_pure_th = @{thm entails_pure} val entails_pure_th' = @{thm entails_pure'} val entails_ex_th = @{thm entails_ex} val entails_frame_th' = @{thm entails_frame'} val entails_frame_th'' = @{thm entails_frame''} val pure_conj_th = @{thm pure_conj} val entails_ex_post_th = @{thm entails_ex_post} val entails_pure_post_th = @{thm entails_pure_post} val pre_rule_th' = @{thm pre_rule'} val pre_rule_th'' = @{thm pre_rule''} val bind_rule_th' = @{thm bind_rule'} val post_rule_th' = @{thm post_rule'} val entails_equiv_forward_th = @{thm entails_equiv_forward} val entails_equiv_backward_th = @{thm entails_equiv_backward} val norm_pre_pure_iff_th = @{thm norm_pre_pure_iff} val norm_pre_pure_iff2_th = @{thm norm_pre_pure_iff2} val entails_trans2_th = @{thm entails_trans2} (* Extra functions *) fun is_case_prod t = case t of Const (@{const_name case_prod}, _) $ _ $ _ => true | _ => false fun sort_by t = case t of Const (@{const_name pure_assn}, _) $ _ => 2 | Const (@{const_name top_assn}, _) => 1 | _ => 0 fun pure_ord (t, s) = if sort_by t = 0 andalso sort_by s = 0 then Term_Ord.term_ord (t, s) = LESS else sort_by t < sort_by s fun normalize_times_cv ctxt ct = let val (A, B) = Util.dest_binop_args (Thm.term_of ct) in if is_ex_assn A then Conv.every_conv [rewr_obj_eq (obj_sym @{thm ex_distrib_star}), Conv.binder_conv (normalize_times_cv o snd) ctxt] ct else if is_ex_assn B then Conv.every_conv [ACUtil.comm_cv assn_ac_info, normalize_times_cv ctxt] ct else Conv.every_conv [ ACUtil.normalize_au assn_ac_info, ACUtil.normalize_comm_gen assn_ac_info pure_ord, ACUtil.norm_combine assn_ac_info is_true_assn (rewr_obj_eq @{thm top_assn_reduce})] ct end (* Normalization function for assertions. This function pulls all EX_A to the front, then apply AC-rules to the inside, putting all pure assertions on the right. *) fun normalize_assn_cv ctxt ct = let val t = Thm.term_of ct in if is_ex_assn t then Conv.binder_conv (normalize_assn_cv o snd) ctxt ct else if UtilArith.is_times t then Conv.every_conv [Conv.binop_conv (normalize_assn_cv ctxt), normalize_times_cv ctxt] ct else if is_pure_assn t andalso is_conj (dest_arg t) then Conv.every_conv [rewr_obj_eq @{thm pure_conj}, normalize_assn_cv ctxt] ct else if is_case_prod t then Conv.every_conv [rewr_obj_eq @{thm case_prod_beta}, normalize_assn_cv ctxt] ct else Conv.all_conv ct end (* Rewrite terms for an assertion *) fun assn_rewr_terms P = P |> ACUtil.dest_ac assn_ac_info |> filter_out is_pure_assn |> maps (snd o Term.strip_comb) |> distinct (op aconv) end (* structure SepUtil *)