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