| author | wenzelm | 
| Thu, 21 Mar 2024 21:03:06 +0100 | |
| changeset 79961 | 2b9205301ff5 | 
| parent 78801 | 42ae6e0ecfd4 | 
| child 81954 | 6f2bcdfa9a19 | 
| permissions | -rw-r--r-- | 
| 63627 | 1 | (* Title: HOL/Analysis/measurable.ML | 
| 50387 | 2 | Author: Johannes Hölzl <hoelzl@in.tum.de> | 
| 3 | ||
| 4 | Measurability prover. | |
| 5 | *) | |
| 6 | ||
| 7 | signature MEASURABLE = | |
| 8 | sig | |
| 59048 | 9 | type preprocessor = thm -> Proof.context -> (thm list * Proof.context) | 
| 10 | ||
| 50387 | 11 | datatype level = Concrete | Generic | 
| 12 | ||
| 59047 | 13 | val dest_thm_attr : attribute context_parser | 
| 59048 | 14 | val cong_thm_attr : attribute context_parser | 
| 59047 | 15 | val measurable_thm_attr : bool * (bool * level) -> attribute | 
| 53043 
8cbfbeb566a4
more standard attribute_setup / method_setup -- export key ML operations instead of parsers;
 wenzelm parents: 
51717diff
changeset | 16 | |
| 59048 | 17 | val add_del_cong_thm : bool -> thm -> Context.generic -> Context.generic ; | 
| 18 | ||
| 19 | val get_all : Context.generic -> thm list | |
| 20 | val get_dest : Context.generic -> thm list | |
| 21 | val get_cong : Context.generic -> thm list | |
| 22 | ||
| 50387 | 23 | val measurable_tac : Proof.context -> thm list -> tactic | 
| 24 | ||
| 78801 | 25 | val proc : Simplifier.proc | 
| 50387 | 26 | |
| 59048 | 27 | val add_preprocessor : string -> preprocessor -> Context.generic -> Context.generic | 
| 28 | val del_preprocessor : string -> Context.generic -> Context.generic | |
| 29 | val add_local_cong : thm -> Proof.context -> Proof.context | |
| 59353 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 30 | |
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 31 | val prepare_facts : Proof.context -> thm list -> (thm list * Proof.context) | 
| 78801 | 32 | end | 
| 50387 | 33 | |
| 34 | structure Measurable : MEASURABLE = | |
| 35 | struct | |
| 36 | ||
| 77898 | 37 | type preprocessor = thm -> Proof.context -> thm list * Proof.context | 
| 59048 | 38 | |
| 50387 | 39 | datatype level = Concrete | Generic; | 
| 40 | ||
| 77898 | 41 | type measurable_thm = thm * (bool * level); | 
| 42 | ||
| 43 | fun eq_measurable_thm ((th1, d1): measurable_thm, (th2, d2): measurable_thm) = | |
| 59047 | 44 | d1 = d2 andalso Thm.eq_thm_prop (th1, th2) ; | 
| 45 | ||
| 77899 | 46 | fun merge_preprocessors (xs: (string * preprocessor) list, ys) = | 
| 77898 | 47 | xs @ (filter (fn (name, _) => is_none (find_first (fn (name', _) => name' = name) xs)) ys) | 
| 59048 | 48 | |
| 50387 | 49 | structure Data = Generic_Data | 
| 50 | ( | |
| 77898 | 51 | type T = | 
| 77899 | 52 | measurable_thm Item_Net.T * | 
| 53 | (*dest_thms*) thm Item_Net.T * | |
| 54 | (*cong_thms*) thm Item_Net.T * | |
| 55 | (string * preprocessor) list | |
| 77898 | 56 | val empty: T = | 
| 77901 | 57 | (Item_Net.init eq_measurable_thm (single o Thm.full_prop_of o fst), Thm.item_net, Thm.item_net, []) | 
| 77898 | 58 | fun merge | 
| 77899 | 59 | ((measurable_thms1, dest_thms1, cong_thms1, preprocessors1), | 
| 60 | (measurable_thms2, dest_thms2, cong_thms2, preprocessors2)) : T = | |
| 61 | (Item_Net.merge (measurable_thms1, measurable_thms2), | |
| 62 | Item_Net.merge (dest_thms1, dest_thms2), | |
| 63 | Item_Net.merge (cong_thms1, cong_thms2), | |
| 64 | merge_preprocessors (preprocessors1, preprocessors2)) | |
| 50387 | 65 | ); | 
| 66 | ||
| 77899 | 67 | val map_measurable_thms = Data.map o @{apply 4(1)}
 | 
| 68 | val map_dest_thms = Data.map o @{apply 4(2)}
 | |
| 69 | val map_cong_thms = Data.map o @{apply 4(3)}
 | |
| 70 | val map_preprocessors = Data.map o @{apply 4(4)}
 | |
| 50387 | 71 | |
| 77899 | 72 | val debug = Attrib.setup_config_bool \<^binding>\<open>measurable_debug\<close> (K false) | 
| 73 | val split = Attrib.setup_config_bool \<^binding>\<open>measurable_split\<close> (K true) | |
| 50387 | 74 | |
| 59992 | 75 | fun generic_add_del map : attribute context_parser = | 
| 59047 | 76 | Scan.lift | 
| 77 | (Args.add >> K Item_Net.update || Args.del >> K Item_Net.remove || Scan.succeed Item_Net.update) >> | |
| 77900 
42214742b44a
proper Thm.trim_context / Thm.transfer for context data;
 wenzelm parents: 
77899diff
changeset | 78 | (fn f => Thm.declaration_attribute (map o f o Thm.trim_context)) | 
| 59047 | 79 | |
| 80 | val dest_thm_attr = generic_add_del map_dest_thms | |
| 59048 | 81 | val cong_thm_attr = generic_add_del map_cong_thms | 
| 82 | ||
| 59047 | 83 | fun del_thm th net = | 
| 84 | let | |
| 77901 | 85 | val thms = net |> Item_Net.content |> filter (fn (th', _) => Thm.eq_thm_prop (th, th')) | 
| 59047 | 86 | in fold Item_Net.remove thms net end ; | 
| 87 | ||
| 88 | fun measurable_thm_attr (do_add, d) = Thm.declaration_attribute | |
| 77900 
42214742b44a
proper Thm.trim_context / Thm.transfer for context data;
 wenzelm parents: 
77899diff
changeset | 89 | (map_measurable_thms o (if do_add then Item_Net.update o rpair d else del_thm) o Thm.trim_context) | 
| 50387 | 90 | |
| 77900 
42214742b44a
proper Thm.trim_context / Thm.transfer for context data;
 wenzelm parents: 
77899diff
changeset | 91 | fun get_dest context = map (Thm.transfer'' context) (Item_Net.content (#2 (Data.get context))); | 
| 
42214742b44a
proper Thm.trim_context / Thm.transfer for context data;
 wenzelm parents: 
77899diff
changeset | 92 | fun get_cong context = map (Thm.transfer'' context) (Item_Net.content (#3 (Data.get context))); | 
| 
42214742b44a
proper Thm.trim_context / Thm.transfer for context data;
 wenzelm parents: 
77899diff
changeset | 93 | |
| 
42214742b44a
proper Thm.trim_context / Thm.transfer for context data;
 wenzelm parents: 
77899diff
changeset | 94 | val add_cong = map_cong_thms o Item_Net.update o Thm.trim_context; | 
| 
42214742b44a
proper Thm.trim_context / Thm.transfer for context data;
 wenzelm parents: 
77899diff
changeset | 95 | val del_cong = map_cong_thms o Item_Net.remove o Thm.trim_context; | 
| 59048 | 96 | fun add_del_cong_thm true = add_cong | 
| 97 | | add_del_cong_thm false = del_cong | |
| 98 | ||
| 77899 | 99 | fun add_preprocessor name f = map_preprocessors (fn xs => xs @ [(name, f)]) | 
| 100 | fun del_preprocessor name = map_preprocessors (filter (fn (n, _) => n <> name)) | |
| 59048 | 101 | val add_local_cong = Context.proof_map o add_cong | 
| 102 | ||
| 77899 | 103 | val get_preprocessors = Context.Proof #> Data.get #> #4; | 
| 59048 | 104 | |
| 50387 | 105 | fun is_too_generic thm = | 
| 106 | let | |
| 59582 | 107 | val concl = Thm.concl_of thm | 
| 50387 | 108 | val concl' = HOLogic.dest_Trueprop concl handle TERM _ => concl | 
| 109 | in is_Var (head_of concl') end | |
| 110 | ||
| 77900 
42214742b44a
proper Thm.trim_context / Thm.transfer for context data;
 wenzelm parents: 
77899diff
changeset | 111 | fun get_thms context = | 
| 
42214742b44a
proper Thm.trim_context / Thm.transfer for context data;
 wenzelm parents: 
77899diff
changeset | 112 | map (apfst (Thm.transfer'' context)) (Item_Net.content (#1 (Data.get context))); | 
| 58965 | 113 | |
| 77900 
42214742b44a
proper Thm.trim_context / Thm.transfer for context data;
 wenzelm parents: 
77899diff
changeset | 114 | val get_all = get_thms #> map fst; | 
| 50387 | 115 | |
| 59048 | 116 | fun debug_tac ctxt msg f = | 
| 117 | if Config.get ctxt debug then print_tac ctxt (msg ()) THEN f else f | |
| 50387 | 118 | |
| 119 | fun nth_hol_goal thm i = | |
| 59582 | 120 | HOLogic.dest_Trueprop (Logic.strip_imp_concl (strip_all_body (nth (Thm.prems_of thm) (i - 1)))) | 
| 50387 | 121 | |
| 122 | fun dest_measurable_fun t = | |
| 123 | (case t of | |
| 74632 | 124 | \<^Const_>\<open>Set.member _ for f \<^Const_>\<open>measurable _ _ for _ _\<close>\<close> => f | 
| 50387 | 125 |   | _ => raise (TERM ("not a measurability predicate", [t])))
 | 
| 126 | ||
| 59582 | 127 | fun not_measurable_prop n thm = | 
| 128 | if length (Thm.prems_of thm) < n then false | |
| 129 | else | |
| 130 | (case nth_hol_goal thm n of | |
| 74632 | 131 | \<^Const_>\<open>Set.member _ for _ \<^Const_>\<open>sets _ for _\<close>\<close> => false | 
| 132 | | \<^Const_>\<open>Set.member _ for _ \<^Const_>\<open>measurable _ _ for _ _\<close>\<close> => false | |
| 59582 | 133 | | _ => true) | 
| 134 | handle TERM _ => true; | |
| 50387 | 135 | |
| 136 | fun indep (Bound i) t b = i < b orelse t <= i | |
| 137 | | indep (f $ t) top bot = indep f top bot andalso indep t top bot | |
| 138 | | indep (Abs (_,_,t)) top bot = indep t (top + 1) (bot + 1) | |
| 139 | | indep _ _ _ = true; | |
| 140 | ||
| 59837 | 141 | fun cnt_prefixes ctxt (Abs (n, T, t)) = | 
| 142 | let | |
| 69597 | 143 | fun is_countable ty = Sign.of_sort (Proof_Context.theory_of ctxt) (ty, \<^sort>\<open>countable\<close>) | 
| 50387 | 144 | fun cnt_walk (Abs (ns, T, t)) Ts = | 
| 145 | map (fn (t', t'') => (Abs (ns, T, t'), t'')) (cnt_walk t (T::Ts)) | |
| 146 | | cnt_walk (f $ g) Ts = let | |
| 147 | val n = length Ts - 1 | |
| 148 | in | |
| 149 | map (fn (f', t) => (f' $ g, t)) (cnt_walk f Ts) @ | |
| 150 | map (fn (g', t) => (f $ g', t)) (cnt_walk g Ts) @ | |
| 151 | (if is_countable (type_of1 (Ts, g)) andalso loose_bvar1 (g, n) | |
| 152 | andalso indep g n 0 andalso g <> Bound n | |
| 153 | then [(f $ Bound (n + 1), incr_boundvars (~ n) g)] | |
| 154 | else []) | |
| 155 | end | |
| 156 | | cnt_walk _ _ = [] | |
| 157 | in map (fn (t1, t2) => let | |
| 158 | val T1 = type_of1 ([T], t2) | |
| 159 | val T2 = type_of1 ([T], t) | |
| 160 | in ([SOME (Abs (n, T1, Abs (n, T, t1))), NONE, NONE, SOME (Abs (n, T, t2))], | |
| 161 | [SOME T1, SOME T, SOME T2]) | |
| 162 | end) (cnt_walk t [T]) | |
| 163 | end | |
| 164 | | cnt_prefixes _ _ = [] | |
| 165 | ||
| 59353 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 166 | fun apply_dests thm dests = | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 167 | let | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 168 | fun apply thm th' = | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 169 | let | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 170 | val th'' = thm RS th' | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 171 | in [th''] @ loop th'' end | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 172 | handle (THM _) => [] | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 173 | and loop thm = | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 174 | flat (map (apply thm) dests) | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 175 | in | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 176 |     [thm] @ ([thm RS @{thm measurable_compose_rev}] handle (THM _) => []) @ loop thm
 | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 177 | end | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 178 | |
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 179 | fun prepare_facts ctxt facts = | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 180 | let | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 181 | val dests = get_dest (Context.Proof ctxt) | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 182 | fun prep_dest thm = | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 183 | (if is_too_generic thm then [] else apply_dests thm dests) ; | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 184 |     val preprocessors = (("std", prep_dest #> pair) :: get_preprocessors ctxt) ;
 | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 185 | fun preprocess_thm (thm, raw) = | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 186 | if raw then pair [thm] else fold_map (fn (_, proc) => proc thm) preprocessors #>> flat | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 187 | |
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 188 | fun sel lv (th, (raw, lv')) = if lv = lv' then SOME (th, raw) else NONE ; | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 189 | fun get lv = ctxt |> Context.Proof |> get_thms |> rev |> map_filter (sel lv) ; | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 190 | val pre_thms = map (Simplifier.norm_hhf ctxt #> rpair false) facts @ get Concrete @ get Generic | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 191 | |
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 192 | val (thms, ctxt) = fold_map preprocess_thm pre_thms ctxt |>> flat | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 193 | in (thms, ctxt) end | 
| 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 194 | |
| 59047 | 195 | fun measurable_tac ctxt facts = | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50387diff
changeset | 196 | let | 
| 59048 | 197 | fun debug_fact msg thm () = | 
| 61877 
276ad4354069
renamed Pretty.str_of to Pretty.unformatted_string_of to emphasize its meaning;
 wenzelm parents: 
61424diff
changeset | 198 | msg ^ " " ^ Pretty.unformatted_string_of (Syntax.pretty_term ctxt (Thm.prop_of thm)) | 
| 59048 | 199 | |
| 200 | fun IF' c t i = COND (c i) (t i) no_tac | |
| 201 | ||
| 202 | fun r_tac msg = | |
| 203 | if Config.get ctxt debug | |
| 204 | then FIRST' o | |
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59353diff
changeset | 205 | map (fn thm => resolve_tac ctxt [thm] | 
| 59353 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 206 | THEN' K (debug_tac ctxt (debug_fact (msg ^ " resolved using") thm) all_tac)) | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59353diff
changeset | 207 | else resolve_tac ctxt | 
| 59048 | 208 | |
| 59353 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 hoelzl parents: 
59048diff
changeset | 209 | val (thms, ctxt) = prepare_facts ctxt facts | 
| 50387 | 210 | |
| 74632 | 211 | fun is_sets_eq \<^Const_>\<open>HOL.eq _ for | 
| 212 | \<^Const_>\<open>sets _ for _\<close> \<^Const_>\<open>sets _ for _\<close>\<close> = true | |
| 213 | | is_sets_eq \<^Const_>\<open>HOL.eq _ for | |
| 214 | \<^Const_>\<open>measurable _ _ for _ _\<close> \<^Const_>\<open>measurable _ _ for _ _\<close>\<close> = true | |
| 59048 | 215 | | is_sets_eq _ = false | 
| 216 | ||
| 217 | val cong_thms = get_cong (Context.Proof ctxt) @ | |
| 59582 | 218 | filter (fn thm => Thm.concl_of thm |> HOLogic.dest_Trueprop |> is_sets_eq handle TERM _ => false) facts | 
| 59048 | 219 | |
| 220 | fun sets_cong_tac i = | |
| 221 |       Subgoal.FOCUS (fn {context = ctxt', prems = prems, ...} => (
 | |
| 222 | let | |
| 223 | val ctxt'' = Simplifier.add_prems prems ctxt' | |
| 224 | in | |
| 74632 | 225 |           r_tac "cong intro" [@{lemma "A = B \<Longrightarrow> x \<in> B \<Longrightarrow> x \<in> A" by simp}]
 | 
| 59048 | 226 | THEN' SOLVED' (fn i => REPEAT_DETERM ( | 
| 227 |               ((r_tac "cong solve" (cong_thms @ [@{thm refl}])
 | |
| 59582 | 228 | ORELSE' IF' (fn i => fn thm => Thm.nprems_of thm > i) | 
| 59048 | 229 | (SOLVED' (asm_full_simp_tac ctxt''))) i))) | 
| 230 | end) 1) ctxt i | |
| 231 | THEN flexflex_tac ctxt | |
| 232 | ||
| 233 | val simp_solver_tac = | |
| 234 | IF' not_measurable_prop (debug_tac ctxt (K "simp ") o SOLVED' (asm_full_simp_tac ctxt)) | |
| 235 | ||
| 236 | val split_countable_tac = | |
| 237 |       Subgoal.FOCUS (fn {context = ctxt, ...} => SUBGOAL (fn (t, i) =>
 | |
| 238 | let | |
| 239 | val f = dest_measurable_fun (HOLogic.dest_Trueprop t) | |
| 240 | fun inst (ts, Ts) = | |
| 60807 | 241 | Thm.instantiate' | 
| 242 | (map (Option.map (Thm.ctyp_of ctxt)) Ts) | |
| 243 | (map (Option.map (Thm.cterm_of ctxt)) ts) | |
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59353diff
changeset | 244 |               @{thm measurable_compose_countable}
 | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
60807diff
changeset | 245 | in r_tac "case_prod countable" (cnt_prefixes ctxt f |> map inst) i end | 
| 59048 | 246 | handle TERM _ => no_tac) 1) | 
| 50387 | 247 | |
| 248 | val splitter = if Config.get ctxt split then split_countable_tac ctxt else K no_tac | |
| 249 | ||
| 59048 | 250 | val single_step_tac = | 
| 251 | simp_solver_tac | |
| 252 | ORELSE' r_tac "step" thms | |
| 253 | ORELSE' splitter | |
| 254 | ORELSE' (CHANGED o sets_cong_tac) | |
| 255 | ORELSE' (K (debug_tac ctxt (K "backtrack") no_tac)) | |
| 50387 | 256 | |
| 59048 | 257 | in debug_tac ctxt (K "start") (REPEAT (single_step_tac 1)) end; | 
| 50387 | 258 | |
| 78801 | 259 | fun proc ctxt redex = | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50387diff
changeset | 260 | let | 
| 59582 | 261 | val t = HOLogic.mk_Trueprop (Thm.term_of redex); | 
| 50387 | 262 |     fun tac {context = ctxt, prems = _ } =
 | 
| 59047 | 263 | SOLVE (measurable_tac ctxt (Simplifier.prems_of ctxt)); | 
| 73551 | 264 |   in \<^try>\<open>Goal.prove ctxt [] [] t tac RS @{thm Eq_TrueI}\<close> end;
 | 
| 50387 | 265 | |
| 266 | end | |
| 267 |