| author | wenzelm | 
| Wed, 05 Feb 2025 12:00:14 +0100 | |
| changeset 82083 | d72a4ecf8c20 | 
| parent 81953 | 02d9844ff892 | 
| child 82834 | 0b2acd437db4 | 
| permissions | -rw-r--r-- | 
| 29265 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 wenzelm parents: 
26939diff
changeset | 1 | (* Title: FOLP/simp.ML | 
| 0 | 2 | Author: Tobias Nipkow | 
| 3 | Copyright 1993 University of Cambridge | |
| 4 | ||
| 5 | FOLP version of... | |
| 6 | ||
| 7 | Generic simplifier, suitable for most logics. (from Provers) | |
| 8 | ||
| 9 | This version allows instantiation of Vars in the subgoal, since the proof | |
| 10 | term must change. | |
| 11 | *) | |
| 12 | ||
| 13 | signature SIMP_DATA = | |
| 14 | sig | |
| 15 | val case_splits : (thm * string) list | |
| 16 | val dest_red : term -> term * term * term | |
| 17 | val mk_rew_rules : thm -> thm list | |
| 18 | val norm_thms : (thm*thm) list (* [(?x>>norm(?x), norm(?x)>>?x), ...] *) | |
| 19 | val red1 : thm (* ?P>>?Q ==> ?P ==> ?Q *) | |
| 20 | val red2 : thm (* ?P>>?Q ==> ?Q ==> ?P *) | |
| 21 | val refl_thms : thm list | |
| 22 | val subst_thms : thm list (* [ ?a>>?b ==> ?P(?a) ==> ?P(?b), ...] *) | |
| 23 | val trans_thms : thm list | |
| 24 | end; | |
| 25 | ||
| 26 | ||
| 60644 | 27 | infix 4 addcongs delrews delcongs setauto; | 
| 0 | 28 | |
| 29 | signature SIMP = | |
| 30 | sig | |
| 31 | type simpset | |
| 32 | val empty_ss : simpset | |
| 33 | val addcongs : simpset * thm list -> simpset | |
| 60644 | 34 | val addrew : Proof.context -> thm -> simpset -> simpset | 
| 0 | 35 | val delcongs : simpset * thm list -> simpset | 
| 36 | val delrews : simpset * thm list -> simpset | |
| 37 | val dest_ss : simpset -> thm list * thm list | |
| 60645 
2affe7e97a34
eliminated spurious warning/tracing messages -- avoid Display.string_of_thm_without_context;
 wenzelm parents: 
60644diff
changeset | 38 | val print_ss : Proof.context -> simpset -> unit | 
| 60774 | 39 | val setauto : simpset * (Proof.context -> int -> tactic) -> simpset | 
| 60646 | 40 | val ASM_SIMP_CASE_TAC : Proof.context -> simpset -> int -> tactic | 
| 41 | val ASM_SIMP_TAC : Proof.context -> simpset -> int -> tactic | |
| 60754 | 42 | val CASE_TAC : Proof.context -> simpset -> int -> tactic | 
| 60646 | 43 | val SIMP_CASE2_TAC : Proof.context -> simpset -> int -> tactic | 
| 44 | val SIMP_THM : Proof.context -> simpset -> thm -> thm | |
| 45 | val SIMP_TAC : Proof.context -> simpset -> int -> tactic | |
| 46 | val SIMP_CASE_TAC : Proof.context -> simpset -> int -> tactic | |
| 60789 | 47 | val mk_congs : Proof.context -> string list -> thm list | 
| 48 | val mk_typed_congs : Proof.context -> (string * string) list -> thm list | |
| 0 | 49 | (* temporarily disabled: | 
| 50 | val extract_free_congs : unit -> thm list | |
| 51 | *) | |
| 32740 | 52 | val tracing : bool Unsynchronized.ref | 
| 0 | 53 | end; | 
| 54 | ||
| 32449 | 55 | functor SimpFun (Simp_data: SIMP_DATA) : SIMP = | 
| 0 | 56 | struct | 
| 57 | ||
| 19805 | 58 | local open Simp_data in | 
| 0 | 59 | |
| 60 | (*For taking apart reductions into left, right hand sides*) | |
| 61 | val lhs_of = #2 o dest_red; | |
| 62 | val rhs_of = #3 o dest_red; | |
| 63 | ||
| 64 | (*** Indexing and filtering of theorems ***) | |
| 65 | ||
| 22360 
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
 wenzelm parents: 
21287diff
changeset | 66 | fun eq_brl ((b1 : bool, th1), (b2, th2)) = b1 = b2 andalso Thm.eq_thm_prop (th1, th2); | 
| 0 | 67 | |
| 68 | (*insert a thm in a discrimination net by its lhs*) | |
| 33339 | 69 | fun lhs_insert_thm th net = | 
| 59582 | 70 | Net.insert_term eq_brl (lhs_of (Thm.concl_of th), (false,th)) net | 
| 0 | 71 | handle Net.INSERT => net; | 
| 72 | ||
| 73 | (*match subgoal i against possible theorems in the net. | |
| 74 | Similar to match_from_nat_tac, but the net does not contain numbers; | |
| 75 | rewrite rules are not ordered.*) | |
| 60756 | 76 | fun net_tac ctxt net = | 
| 77 | SUBGOAL(fn (prem, i) => | |
| 78 | resolve_tac ctxt (Net.unify_term net (Logic.strip_assums_concl prem)) i); | |
| 0 | 79 | |
| 80 | (*match subgoal i against possible theorems indexed by lhs in the net*) | |
| 60756 | 81 | fun lhs_net_tac ctxt net = | 
| 32449 | 82 | SUBGOAL(fn (prem,i) => | 
| 60756 | 83 | biresolve_tac ctxt (Net.unify_term net | 
| 19805 | 84 | (lhs_of (Logic.strip_assums_concl prem))) i); | 
| 0 | 85 | |
| 81953 
02d9844ff892
tuned: prefer Thm.prem_of, which differes wrt. exceptions that are not handled here;
 wenzelm parents: 
74525diff
changeset | 86 | fun goal_concl i thm = Logic.strip_assums_concl (Thm.prem_of thm i); | 
| 0 | 87 | |
| 88 | fun lhs_of_eq i thm = lhs_of(goal_concl i thm) | |
| 89 | and rhs_of_eq i thm = rhs_of(goal_concl i thm); | |
| 90 | ||
| 91 | fun var_lhs(thm,i) = | |
| 92 | let fun var(Var _) = true | |
| 93 | | var(Abs(_,_,t)) = var t | |
| 94 | | var(f$_) = var f | |
| 95 | | var _ = false; | |
| 96 | in var(lhs_of_eq i thm) end; | |
| 97 | ||
| 98 | fun contains_op opns = | |
| 36692 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 haftmann parents: 
36603diff
changeset | 99 | let fun contains(Const(s,_)) = member (op =) opns s | | 
| 0 | 100 | contains(s$t) = contains s orelse contains t | | 
| 101 | contains(Abs(_,_,t)) = contains t | | |
| 102 | contains _ = false; | |
| 103 | in contains end; | |
| 104 | ||
| 105 | fun may_match(match_ops,i) = contains_op match_ops o lhs_of_eq i; | |
| 106 | ||
| 107 | val (normI_thms,normE_thms) = split_list norm_thms; | |
| 108 | ||
| 109 | (*Get the norm constants from norm_thms*) | |
| 110 | val norms = | |
| 32449 | 111 | let fun norm thm = | 
| 59582 | 112 | case lhs_of (Thm.concl_of thm) of | 
| 1459 | 113 | Const(n,_)$_ => n | 
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
30190diff
changeset | 114 | | _ => error "No constant in lhs of a norm_thm" | 
| 0 | 115 | in map norm normE_thms end; | 
| 116 | ||
| 117 | fun lhs_is_NORM(thm,i) = case lhs_of_eq i thm of | |
| 36692 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 haftmann parents: 
36603diff
changeset | 118 | Const(s,_)$_ => member (op =) norms s | _ => false; | 
| 0 | 119 | |
| 60756 | 120 | fun refl_tac ctxt = resolve_tac ctxt refl_thms; | 
| 0 | 121 | |
| 122 | fun find_res thms thm = | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
30190diff
changeset | 123 | let fun find [] = error "Check Simp_Data" | 
| 6969 | 124 | | find(th::thms) = thm RS th handle THM _ => find thms | 
| 0 | 125 | in find thms end; | 
| 126 | ||
| 127 | val mk_trans = find_res trans_thms; | |
| 128 | ||
| 129 | fun mk_trans2 thm = | |
| 130 | let fun mk[] = error"Check transitivity" | |
| 6969 | 131 | | mk(t::ts) = (thm RSN (2,t)) handle THM _ => mk ts | 
| 0 | 132 | in mk trans_thms end; | 
| 133 | ||
| 134 | (*Applies tactic and returns the first resulting state, FAILS if none!*) | |
| 4271 
3a82492e70c5
changed Pure/Sequence interface -- isatool fixseq;
 wenzelm parents: 
3537diff
changeset | 135 | fun one_result(tac,thm) = case Seq.pull(tac thm) of | 
| 15531 | 136 | SOME(thm',_) => thm' | 
| 137 |       | NONE => raise THM("Simplifier: could not continue", 0, [thm]);
 | |
| 0 | 138 | |
| 60756 | 139 | fun res1 ctxt (thm,thms,i) = one_result (resolve_tac ctxt thms i,thm); | 
| 0 | 140 | |
| 141 | ||
| 142 | (**** Adding "NORM" tags ****) | |
| 143 | ||
| 144 | (*get name of the constant from conclusion of a congruence rule*) | |
| 32449 | 145 | fun cong_const cong = | 
| 59582 | 146 | case head_of (lhs_of (Thm.concl_of cong)) of | 
| 1459 | 147 | Const(c,_) => c | 
| 148 | | _ => "" (*a placeholder distinct from const names*); | |
| 0 | 149 | |
| 150 | (*true if the term is an atomic proposition (no ==> signs) *) | |
| 19805 | 151 | val atomic = null o Logic.strip_assums_hyp; | 
| 0 | 152 | |
| 153 | (*ccs contains the names of the constants possessing congruence rules*) | |
| 154 | fun add_hidden_vars ccs = | |
| 21078 | 155 | let fun add_hvars tm hvars = case tm of | 
| 44121 | 156 | Abs(_,_,body) => Misc_Legacy.add_term_vars(body,hvars) | 
| 32449 | 157 | | _$_ => let val (f,args) = strip_comb tm | 
| 1459 | 158 | in case f of | 
| 74301 | 159 | Const(c,_) => | 
| 21078 | 160 | if member (op =) ccs c | 
| 161 | then fold_rev add_hvars args hvars | |
| 44121 | 162 | else Misc_Legacy.add_term_vars (tm, hvars) | 
| 163 | | _ => Misc_Legacy.add_term_vars (tm, hvars) | |
| 1459 | 164 | end | 
| 165 | | _ => hvars; | |
| 0 | 166 | in add_hvars end; | 
| 167 | ||
| 168 | fun add_new_asm_vars new_asms = | |
| 21078 | 169 | let fun itf (tm, at) vars = | 
| 44121 | 170 | if at then vars else Misc_Legacy.add_term_vars(tm,vars) | 
| 1459 | 171 | fun add_list(tm,al,vars) = let val (_,tml) = strip_comb tm | 
| 172 | in if length(tml)=length(al) | |
| 21078 | 173 | then fold_rev itf (tml ~~ al) vars | 
| 1459 | 174 | else vars | 
| 175 | end | |
| 176 | fun add_vars (tm,vars) = case tm of | |
| 177 | Abs (_,_,body) => add_vars(body,vars) | |
| 178 | | r$s => (case head_of tm of | |
| 74301 | 179 | Const(c,_) => (case AList.lookup (op =) new_asms c of | 
| 15531 | 180 | NONE => add_vars(r,add_vars(s,vars)) | 
| 181 | | SOME(al) => add_list(tm,al,vars)) | |
| 1459 | 182 | | _ => add_vars(r,add_vars(s,vars))) | 
| 183 | | _ => vars | |
| 0 | 184 | in add_vars end; | 
| 185 | ||
| 186 | ||
| 60756 | 187 | fun add_norms ctxt (congs,ccs,new_asms) thm = | 
| 0 | 188 | let val thm' = mk_trans2 thm; | 
| 189 | (* thm': [?z -> l; Prems; r -> ?t] ==> ?z -> ?t *) | |
| 59582 | 190 | val nops = Thm.nprems_of thm' | 
| 0 | 191 | val lhs = rhs_of_eq 1 thm' | 
| 192 | val rhs = lhs_of_eq nops thm' | |
| 59582 | 193 | val asms = tl(rev(tl(Thm.prems_of thm'))) | 
| 21078 | 194 | val hvars = fold_rev (add_hidden_vars ccs) (lhs::rhs::asms) [] | 
| 0 | 195 | val hvars = add_new_asm_vars new_asms (rhs,hvars) | 
| 21078 | 196 | fun it_asms asm hvars = | 
| 1459 | 197 | if atomic asm then add_new_asm_vars new_asms (asm,hvars) | 
| 44121 | 198 | else Misc_Legacy.add_term_frees(asm,hvars) | 
| 21078 | 199 | val hvars = fold_rev it_asms asms hvars | 
| 0 | 200 | val hvs = map (#1 o dest_Var) hvars | 
| 3537 | 201 | fun norm_step_tac st = st |> | 
| 32449 | 202 | (case head_of(rhs_of_eq 1 st) of | 
| 60756 | 203 | Var(ixn,_) => if member (op =) hvs ixn then refl_tac ctxt 1 | 
| 204 | else resolve_tac ctxt normI_thms 1 ORELSE refl_tac ctxt 1 | |
| 205 | | Const _ => resolve_tac ctxt normI_thms 1 ORELSE | |
| 206 | resolve_tac ctxt congs 1 ORELSE refl_tac ctxt 1 | |
| 207 | | Free _ => resolve_tac ctxt congs 1 ORELSE refl_tac ctxt 1 | |
| 208 | | _ => refl_tac ctxt 1) | |
| 3537 | 209 | val add_norm_tac = DEPTH_FIRST (has_fewer_prems nops) norm_step_tac | 
| 15531 | 210 | val SOME(thm'',_) = Seq.pull(add_norm_tac thm') | 
| 0 | 211 | in thm'' end; | 
| 212 | ||
| 60756 | 213 | fun add_norm_tags ctxt congs = | 
| 0 | 214 | let val ccs = map cong_const congs | 
| 33317 | 215 | val new_asms = filter (exists not o #2) | 
| 59582 | 216 | (ccs ~~ (map (map atomic o Thm.prems_of) congs)); | 
| 60756 | 217 | in add_norms ctxt (congs,ccs,new_asms) end; | 
| 0 | 218 | |
| 60756 | 219 | fun normed_rews ctxt congs = | 
| 60644 | 220 | let | 
| 60756 | 221 | val add_norms = add_norm_tags ctxt congs | 
| 222 | fun normed thm = | |
| 59170 | 223 | let | 
| 60644 | 224 | val ctxt' = Variable.declare_thm thm ctxt; | 
| 59170 | 225 | in Variable.tradeT (K (map (add_norms o mk_trans) o maps mk_rew_rules)) ctxt [thm] end | 
| 60644 | 226 | in normed end; | 
| 0 | 227 | |
| 60756 | 228 | fun NORM ctxt norm_lhs_tac = EVERY' [resolve_tac ctxt [red2], norm_lhs_tac, refl_tac ctxt]; | 
| 0 | 229 | |
| 230 | val trans_norms = map mk_trans normE_thms; | |
| 231 | ||
| 232 | ||
| 233 | (* SIMPSET *) | |
| 234 | ||
| 235 | datatype simpset = | |
| 60774 | 236 |         SS of {auto_tac: Proof.context -> int -> tactic,
 | 
| 1459 | 237 | congs: thm list, | 
| 238 | cong_net: thm Net.net, | |
| 60644 | 239 | mk_simps: Proof.context -> thm -> thm list, | 
| 1459 | 240 | simps: (thm * thm list) list, | 
| 241 | simp_net: thm Net.net} | |
| 0 | 242 | |
| 60774 | 243 | val empty_ss = SS{auto_tac= K (K no_tac), congs=[], cong_net=Net.empty,
 | 
| 60756 | 244 | mk_simps = fn ctxt => normed_rews ctxt [], simps=[], simp_net=Net.empty}; | 
| 0 | 245 | |
| 246 | (** Insertion of congruences and rewrites **) | |
| 247 | ||
| 248 | (*insert a thm in a thm net*) | |
| 60645 
2affe7e97a34
eliminated spurious warning/tracing messages -- avoid Display.string_of_thm_without_context;
 wenzelm parents: 
60644diff
changeset | 249 | fun insert_thm th net = | 
| 59582 | 250 | Net.insert_term Thm.eq_thm_prop (Thm.concl_of th, th) net | 
| 60645 
2affe7e97a34
eliminated spurious warning/tracing messages -- avoid Display.string_of_thm_without_context;
 wenzelm parents: 
60644diff
changeset | 251 | handle Net.INSERT => net; | 
| 0 | 252 | |
| 60645 
2affe7e97a34
eliminated spurious warning/tracing messages -- avoid Display.string_of_thm_without_context;
 wenzelm parents: 
60644diff
changeset | 253 | val insert_thms = fold_rev insert_thm; | 
| 0 | 254 | |
| 60644 | 255 | fun addrew ctxt thm (SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}) =
 | 
| 70478 | 256 | let val thms = map Thm.trim_context (mk_simps ctxt thm) | 
| 0 | 257 | in SS{auto_tac=auto_tac,congs=congs, cong_net=cong_net, mk_simps=mk_simps,
 | 
| 21078 | 258 | simps = (thm,thms)::simps, simp_net = insert_thms thms simp_net} | 
| 0 | 259 | end; | 
| 260 | ||
| 261 | fun op addcongs(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}, thms) =
 | |
| 70478 | 262 | let val congs' = map Thm.trim_context thms @ congs; | 
| 0 | 263 | in SS{auto_tac=auto_tac, congs= congs',
 | 
| 21078 | 264 | cong_net= insert_thms (map mk_trans thms) cong_net, | 
| 60756 | 265 | mk_simps = fn ctxt => normed_rews ctxt congs', simps=simps, simp_net=simp_net} | 
| 0 | 266 | end; | 
| 267 | ||
| 268 | (** Deletion of congruences and rewrites **) | |
| 269 | ||
| 270 | (*delete a thm from a thm net*) | |
| 60645 
2affe7e97a34
eliminated spurious warning/tracing messages -- avoid Display.string_of_thm_without_context;
 wenzelm parents: 
60644diff
changeset | 271 | fun delete_thm th net = | 
| 59582 | 272 | Net.delete_term Thm.eq_thm_prop (Thm.concl_of th, th) net | 
| 60645 
2affe7e97a34
eliminated spurious warning/tracing messages -- avoid Display.string_of_thm_without_context;
 wenzelm parents: 
60644diff
changeset | 273 | handle Net.DELETE => net; | 
| 0 | 274 | |
| 60645 
2affe7e97a34
eliminated spurious warning/tracing messages -- avoid Display.string_of_thm_without_context;
 wenzelm parents: 
60644diff
changeset | 275 | val delete_thms = fold_rev delete_thm; | 
| 0 | 276 | |
| 277 | fun op delcongs(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}, thms) =
 | |
| 22360 
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
 wenzelm parents: 
21287diff
changeset | 278 | let val congs' = fold (remove Thm.eq_thm_prop) thms congs | 
| 0 | 279 | in SS{auto_tac=auto_tac, congs= congs',
 | 
| 21078 | 280 | cong_net= delete_thms (map mk_trans thms) cong_net, | 
| 60756 | 281 | mk_simps= fn ctxt => normed_rews ctxt congs', simps=simps, simp_net=simp_net} | 
| 0 | 282 | end; | 
| 283 | ||
| 33245 | 284 | fun delrew thm (SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}) =
 | 
| 0 | 285 | let fun find((p as (th,ths))::ps',ps) = | 
| 22360 
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
 wenzelm parents: 
21287diff
changeset | 286 | if Thm.eq_thm_prop(thm,th) then (ths,ps@ps') else find(ps',p::ps) | 
| 60645 
2affe7e97a34
eliminated spurious warning/tracing messages -- avoid Display.string_of_thm_without_context;
 wenzelm parents: 
60644diff
changeset | 287 | | find([],simps') = ([], simps') | 
| 0 | 288 | val (thms,simps') = find(simps,[]) | 
| 289 | in SS{auto_tac=auto_tac, congs=congs, cong_net=cong_net, mk_simps=mk_simps,
 | |
| 21078 | 290 | simps = simps', simp_net = delete_thms thms simp_net } | 
| 0 | 291 | end; | 
| 292 | ||
| 33245 | 293 | fun ss delrews thms = fold delrew thms ss; | 
| 0 | 294 | |
| 295 | ||
| 296 | fun op setauto(SS{congs,cong_net,mk_simps,simps,simp_net,...}, auto_tac) =
 | |
| 297 |     SS{auto_tac=auto_tac, congs=congs, cong_net=cong_net, mk_simps=mk_simps,
 | |
| 298 | simps=simps, simp_net=simp_net}; | |
| 299 | ||
| 300 | ||
| 301 | (** Inspection of a simpset **) | |
| 302 | ||
| 303 | fun dest_ss(SS{congs,simps,...}) = (congs, map #1 simps);
 | |
| 304 | ||
| 60645 
2affe7e97a34
eliminated spurious warning/tracing messages -- avoid Display.string_of_thm_without_context;
 wenzelm parents: 
60644diff
changeset | 305 | fun print_ss ctxt (SS{congs,simps,...}) =
 | 
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
30190diff
changeset | 306 | writeln (cat_lines | 
| 61268 | 307 | (["Congruences:"] @ map (Thm.string_of_thm ctxt) congs @ | 
| 308 | ["Rewrite Rules:"] @ map (Thm.string_of_thm ctxt o #1) simps)); | |
| 0 | 309 | |
| 310 | ||
| 311 | (* Rewriting with conditionals *) | |
| 312 | ||
| 313 | val (case_thms,case_consts) = split_list case_splits; | |
| 314 | val case_rews = map mk_trans case_thms; | |
| 315 | ||
| 316 | fun if_rewritable ifc i thm = | |
| 317 | let val tm = goal_concl i thm | |
| 1459 | 318 | fun nobound(Abs(_,_,tm),j,k) = nobound(tm,j,k+1) | 
| 319 | | nobound(s$t,j,k) = nobound(s,j,k) andalso nobound(t,j,k) | |
| 320 | | nobound(Bound n,j,k) = n < k orelse k+j <= n | |
| 321 | | nobound(_) = true; | |
| 322 | fun check_args(al,j) = forall (fn t => nobound(t,j,0)) al | |
| 323 | fun find_if(Abs(_,_,tm),j) = find_if(tm,j+1) | |
| 324 | | find_if(tm as s$t,j) = let val (f,al) = strip_comb tm in | |
| 325 | case f of Const(c,_) => if c=ifc then check_args(al,j) | |
| 326 | else find_if(s,j) orelse find_if(t,j) | |
| 327 | | _ => find_if(s,j) orelse find_if(t,j) end | |
| 328 | | find_if(_) = false; | |
| 0 | 329 | in find_if(tm,0) end; | 
| 330 | ||
| 60754 | 331 | fun IF1_TAC ctxt cong_tac i = | 
| 32449 | 332 | let fun seq_try (ifth::ifths,ifc::ifcs) thm = | 
| 60754 | 333 | (COND (if_rewritable ifc i) (DETERM(resolve_tac ctxt [ifth] i)) | 
| 1512 | 334 | (seq_try(ifths,ifcs))) thm | 
| 335 | | seq_try([],_) thm = no_tac thm | |
| 336 | and try_rew thm = (seq_try(case_rews,case_consts) ORELSE one_subt) thm | |
| 1459 | 337 | and one_subt thm = | 
| 59582 | 338 | let val test = has_fewer_prems (Thm.nprems_of thm + 1) | 
| 32449 | 339 | fun loop thm = | 
| 340 | COND test no_tac | |
| 60756 | 341 | ((try_rew THEN DEPTH_FIRST test (refl_tac ctxt i)) | 
| 342 | ORELSE (refl_tac ctxt i THEN loop)) thm | |
| 1512 | 343 | in (cong_tac THEN loop) thm end | 
| 344 | in COND (may_match(case_consts,i)) try_rew no_tac end; | |
| 0 | 345 | |
| 60754 | 346 | fun CASE_TAC ctxt (SS{cong_net,...}) i =
 | 
| 60756 | 347 | let val cong_tac = net_tac ctxt cong_net i | 
| 348 | in NORM ctxt (IF1_TAC ctxt cong_tac) i end; | |
| 349 | ||
| 0 | 350 | |
| 351 | (* Rewriting Automaton *) | |
| 352 | ||
| 353 | datatype cntrl = STOP | MK_EQ | ASMS of int | SIMP_LHS | REW | REFL | TRUE | |
| 1459 | 354 | | PROVE | POP_CS | POP_ARTR | IF; | 
| 22578 | 355 | |
| 0 | 356 | fun simp_refl([],_,ss) = ss | 
| 357 | | simp_refl(a'::ns,a,ss) = if a'=a then simp_refl(ns,a,SIMP_LHS::REFL::ss) | |
| 1459 | 358 | else simp_refl(ns,a,ASMS(a)::SIMP_LHS::REFL::POP_ARTR::ss); | 
| 0 | 359 | |
| 360 | (** Tracing **) | |
| 361 | ||
| 32740 | 362 | val tracing = Unsynchronized.ref false; | 
| 0 | 363 | |
| 364 | (*Replace parameters by Free variables in P*) | |
| 365 | fun variants_abs ([],P) = P | |
| 366 | | variants_abs ((a,T)::aTs, P) = | |
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74408diff
changeset | 367 | variants_abs (aTs, #2 (Term.dest_abs_global (Abs (a,T,P)))); | 
| 0 | 368 | |
| 369 | (*Select subgoal i from proof state; substitute parameters, for printing*) | |
| 370 | fun prepare_goal i st = | |
| 81953 
02d9844ff892
tuned: prefer Thm.prem_of, which differes wrt. exceptions that are not handled here;
 wenzelm parents: 
74525diff
changeset | 371 | let val subgi = Thm.prem_of st i | 
| 19805 | 372 | val params = rev (Logic.strip_params subgi) | 
| 373 | in variants_abs (params, Logic.strip_assums_concl subgi) end; | |
| 0 | 374 | |
| 375 | (*print lhs of conclusion of subgoal i*) | |
| 60646 | 376 | fun pr_goal_lhs ctxt i st = | 
| 377 | writeln (Syntax.string_of_term ctxt (lhs_of (prepare_goal i st))); | |
| 0 | 378 | |
| 379 | (*print conclusion of subgoal i*) | |
| 60646 | 380 | fun pr_goal_concl ctxt i st = | 
| 381 | writeln (Syntax.string_of_term ctxt (prepare_goal i st)) | |
| 0 | 382 | |
| 383 | (*print subgoals i to j (inclusive)*) | |
| 60646 | 384 | fun pr_goals ctxt (i,j) st = | 
| 0 | 385 | if i>j then () | 
| 60646 | 386 | else (pr_goal_concl ctxt i st; pr_goals ctxt (i+1,j) st); | 
| 0 | 387 | |
| 388 | (*Print rewrite for tracing; i=subgoal#, n=number of new subgoals, | |
| 389 | thm=old state, thm'=new state *) | |
| 60646 | 390 | fun pr_rew ctxt (i,n,thm,thm',not_asms) = | 
| 0 | 391 | if !tracing | 
| 392 | then (if not_asms then () else writeln"Assumption used in"; | |
| 60646 | 393 | pr_goal_lhs ctxt i thm; writeln"->"; pr_goal_lhs ctxt (i+n) thm'; | 
| 394 | if n>0 then (writeln"Conditions:"; pr_goals ctxt (i, i+n-1) thm') | |
| 0 | 395 | else (); | 
| 396 | writeln"" ) | |
| 397 | else (); | |
| 398 | ||
| 399 | (* Skip the first n hyps of a goal, and return the rest in generalized form *) | |
| 74301 | 400 | fun strip_varify(\<^Const_>\<open>Pure.imp for H B\<close>, n, vs) = | 
| 1459 | 401 | if n=0 then subst_bounds(vs,H)::strip_varify(B,0,vs) | 
| 402 | else strip_varify(B,n-1,vs) | |
| 74301 | 403 | | strip_varify(\<^Const_>\<open>Pure.all _ for \<open>Abs(_,T,t)\<close>\<close>, n, vs) = | 
| 1459 | 404 |         strip_varify(t,n,Var(("?",length vs),T)::vs)
 | 
| 0 | 405 | | strip_varify _ = []; | 
| 406 | ||
| 60646 | 407 | fun execute ctxt (ss,if_fl,auto_tac,cong_tac,net,i,thm) = | 
| 408 | let | |
| 0 | 409 | |
| 410 | fun simp_lhs(thm,ss,anet,ats,cs) = | |
| 411 | if var_lhs(thm,i) then (ss,thm,anet,ats,cs) else | |
| 60756 | 412 | if lhs_is_NORM(thm,i) then (ss, res1 ctxt (thm,trans_norms,i), anet,ats,cs) | 
| 4271 
3a82492e70c5
changed Pure/Sequence interface -- isatool fixseq;
 wenzelm parents: 
3537diff
changeset | 413 | else case Seq.pull(cong_tac i thm) of | 
| 15531 | 414 | SOME(thm',_) => | 
| 59582 | 415 | let val ps = Thm.prems_of thm | 
| 416 | and ps' = Thm.prems_of thm'; | |
| 1459 | 417 | val n = length(ps')-length(ps); | 
| 42364 | 418 | val a = length(Logic.strip_assums_hyp(nth ps (i - 1))) | 
| 33955 | 419 | val l = map (length o Logic.strip_assums_hyp) (take n (drop (i-1) ps')); | 
| 1459 | 420 | in (simp_refl(rev(l),a,REW::ss),thm',anet,ats,cs) end | 
| 15531 | 421 | | NONE => (REW::ss,thm,anet,ats,cs); | 
| 0 | 422 | |
| 423 | (*NB: the "Adding rewrites:" trace will look strange because assumptions | |
| 424 | are represented by rules, generalized over their parameters*) | |
| 425 | fun add_asms(ss,thm,a,anet,ats,cs) = | |
| 81953 
02d9844ff892
tuned: prefer Thm.prem_of, which differes wrt. exceptions that are not handled here;
 wenzelm parents: 
74525diff
changeset | 426 | let val As = strip_varify(Thm.prem_of thm i, a, []); | 
| 60646 | 427 | val thms = map (Thm.trivial o Thm.cterm_of ctxt) As; | 
| 32952 | 428 | val new_rws = maps mk_rew_rules thms; | 
| 429 | val rwrls = map mk_trans (maps mk_rew_rules thms); | |
| 33339 | 430 | val anet' = fold_rev lhs_insert_thm rwrls anet; | 
| 60645 
2affe7e97a34
eliminated spurious warning/tracing messages -- avoid Display.string_of_thm_without_context;
 wenzelm parents: 
60644diff
changeset | 431 | in (ss,thm,anet',anet::ats,cs) end; | 
| 0 | 432 | |
| 4271 
3a82492e70c5
changed Pure/Sequence interface -- isatool fixseq;
 wenzelm parents: 
3537diff
changeset | 433 | fun rew(seq,thm,ss,anet,ats,cs, more) = case Seq.pull seq of | 
| 15531 | 434 | SOME(thm',seq') => | 
| 59582 | 435 | let val n = (Thm.nprems_of thm') - (Thm.nprems_of thm) | 
| 60646 | 436 | in pr_rew ctxt (i,n,thm,thm',more); | 
| 1459 | 437 | if n=0 then (SIMP_LHS::ss, thm', anet, ats, cs) | 
| 438 | else ((replicate n PROVE) @ (POP_CS::SIMP_LHS::ss), | |
| 439 | thm', anet, ats, (ss,thm,anet,ats,seq',more)::cs) | |
| 440 | end | |
| 15531 | 441 | | NONE => if more | 
| 60756 | 442 | then rew((lhs_net_tac ctxt anet i THEN assume_tac ctxt i) thm, | 
| 1459 | 443 | thm,ss,anet,ats,cs,false) | 
| 444 | else (ss,thm,anet,ats,cs); | |
| 0 | 445 | |
| 446 | fun try_true(thm,ss,anet,ats,cs) = | |
| 60774 | 447 | case Seq.pull(auto_tac ctxt i thm) of | 
| 15531 | 448 | SOME(thm',_) => (ss,thm',anet,ats,cs) | 
| 449 | | NONE => let val (ss0,thm0,anet0,ats0,seq,more)::cs0 = cs | |
| 1459 | 450 | in if !tracing | 
| 451 | then (writeln"*** Failed to prove precondition. Normal form:"; | |
| 60646 | 452 | pr_goal_concl ctxt i thm; writeln"") | 
| 1459 | 453 | else (); | 
| 454 | rew(seq,thm0,ss0,anet0,ats0,cs0,more) | |
| 455 | end; | |
| 0 | 456 | |
| 457 | fun if_exp(thm,ss,anet,ats,cs) = | |
| 60754 | 458 | case Seq.pull (IF1_TAC ctxt (cong_tac i) i thm) of | 
| 15531 | 459 | SOME(thm',_) => (SIMP_LHS::IF::ss,thm',anet,ats,cs) | 
| 460 | | NONE => (ss,thm,anet,ats,cs); | |
| 0 | 461 | |
| 462 | fun step(s::ss, thm, anet, ats, cs) = case s of | |
| 60756 | 463 | MK_EQ => (ss, res1 ctxt (thm,[red2],i), anet, ats, cs) | 
| 1459 | 464 | | ASMS(a) => add_asms(ss,thm,a,anet,ats,cs) | 
| 465 | | SIMP_LHS => simp_lhs(thm,ss,anet,ats,cs) | |
| 60756 | 466 | | REW => rew(net_tac ctxt net i thm,thm,ss,anet,ats,cs,true) | 
| 467 | | REFL => (ss, res1 ctxt (thm,refl_thms,i), anet, ats, cs) | |
| 468 | | TRUE => try_true(res1 ctxt (thm,refl_thms,i),ss,anet,ats,cs) | |
| 1459 | 469 | | PROVE => (if if_fl then MK_EQ::SIMP_LHS::IF::TRUE::ss | 
| 470 | else MK_EQ::SIMP_LHS::TRUE::ss, thm, anet, ats, cs) | |
| 471 | | POP_ARTR => (ss,thm,hd ats,tl ats,cs) | |
| 472 | | POP_CS => (ss,thm,anet,ats,tl cs) | |
| 473 | | IF => if_exp(thm,ss,anet,ats,cs); | |
| 0 | 474 | |
| 475 | fun exec(state as (s::ss, thm, _, _, _)) = | |
| 1459 | 476 | if s=STOP then thm else exec(step(state)); | 
| 0 | 477 | |
| 478 | in exec(ss, thm, Net.empty, [], []) end; | |
| 479 | ||
| 480 | ||
| 60646 | 481 | fun EXEC_TAC ctxt (ss,fl) (SS{auto_tac,cong_net,simp_net,...}) =
 | 
| 60756 | 482 | let val cong_tac = net_tac ctxt cong_net | 
| 32449 | 483 | in fn i => | 
| 1512 | 484 | (fn thm => | 
| 59582 | 485 | if i <= 0 orelse Thm.nprems_of thm < i then Seq.empty | 
| 60646 | 486 | else Seq.single(execute ctxt (ss,fl,auto_tac,cong_tac,simp_net,i,thm))) | 
| 60774 | 487 | THEN TRY(auto_tac ctxt i) | 
| 0 | 488 | end; | 
| 489 | ||
| 60646 | 490 | fun SIMP_TAC ctxt = EXEC_TAC ctxt ([MK_EQ,SIMP_LHS,REFL,STOP],false); | 
| 491 | fun SIMP_CASE_TAC ctxt = EXEC_TAC ctxt ([MK_EQ,SIMP_LHS,IF,REFL,STOP],false); | |
| 0 | 492 | |
| 60646 | 493 | fun ASM_SIMP_TAC ctxt = EXEC_TAC ctxt ([ASMS(0),MK_EQ,SIMP_LHS,REFL,STOP],false); | 
| 494 | fun ASM_SIMP_CASE_TAC ctxt = EXEC_TAC ctxt ([ASMS(0),MK_EQ,SIMP_LHS,IF,REFL,STOP],false); | |
| 0 | 495 | |
| 60646 | 496 | fun SIMP_CASE2_TAC ctxt = EXEC_TAC ctxt ([MK_EQ,SIMP_LHS,IF,REFL,STOP],true); | 
| 0 | 497 | |
| 60646 | 498 | fun REWRITE ctxt (ss,fl) (SS{auto_tac,cong_net,simp_net,...}) =
 | 
| 60756 | 499 | let val cong_tac = net_tac ctxt cong_net | 
| 0 | 500 | in fn thm => let val state = thm RSN (2,red1) | 
| 60646 | 501 | in execute ctxt (ss,fl,auto_tac,cong_tac,simp_net,1,state) end | 
| 0 | 502 | end; | 
| 503 | ||
| 60646 | 504 | fun SIMP_THM ctxt = REWRITE ctxt ([ASMS(0),SIMP_LHS,IF,REFL,STOP],false); | 
| 0 | 505 | |
| 506 | ||
| 507 | (* Compute Congruence rules for individual constants using the substition | |
| 508 | rules *) | |
| 509 | ||
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
33955diff
changeset | 510 | val subst_thms = map Drule.export_without_context subst_thms; | 
| 0 | 511 | |
| 512 | ||
| 513 | fun exp_app(0,t) = t | |
| 514 | | exp_app(i,t) = exp_app(i-1,t $ Bound (i-1)); | |
| 515 | ||
| 74301 | 516 | fun exp_abs(\<^Type>\<open>fun T1 T2\<close>,t,i) = | 
| 1459 | 517 |         Abs("x"^string_of_int i,T1,exp_abs(T2,t,i+1))
 | 
| 0 | 518 | | exp_abs(T,t,i) = exp_app(i,t); | 
| 519 | ||
| 520 | fun eta_Var(ixn,T) = exp_abs(T,Var(ixn,T),0); | |
| 521 | ||
| 522 | ||
| 523 | fun Pinst(f,fT,(eq,eqT),k,i,T,yik,Ts) = | |
| 524 | let fun xn_list(x,n) = | |
| 33063 | 525 | let val ixs = map_range (fn i => (x^(radixstring(26,"a",i)),0)) (n - 1); | 
| 33955 | 526 | in ListPair.map eta_Var (ixs, take (n+1) Ts) end | 
| 0 | 527 |     val lhs = list_comb(f,xn_list("X",k-1))
 | 
| 528 |     val rhs = list_comb(f,xn_list("X",i-1) @ [Bound 0] @ yik)
 | |
| 529 | in Abs("", T, Const(eq,[fT,fT]--->eqT) $ lhs $ rhs) end;
 | |
| 530 | ||
| 60789 | 531 | fun find_subst ctxt T = | 
| 0 | 532 | let fun find (thm::thms) = | 
| 81953 
02d9844ff892
tuned: prefer Thm.prem_of, which differes wrt. exceptions that are not handled here;
 wenzelm parents: 
74525diff
changeset | 533 | let val (Const(_,cT), va, vb) = dest_red(Thm.prem_of thm 1); | 
| 59582 | 534 | val [P] = subtract (op =) [va, vb] (Misc_Legacy.add_term_vars (Thm.concl_of thm, [])); | 
| 1459 | 535 | val eqT::_ = binder_types cT | 
| 60789 | 536 | in if Sign.typ_instance (Proof_Context.theory_of ctxt) (T,eqT) then SOME(thm,va,vb,P) | 
| 1459 | 537 | else find thms | 
| 538 | end | |
| 15531 | 539 | | find [] = NONE | 
| 0 | 540 | in find subst_thms end; | 
| 541 | ||
| 60789 | 542 | fun mk_cong ctxt (f,aTs,rT) (refl,eq) = | 
| 16931 | 543 | let val k = length aTs; | 
| 60790 | 544 | fun ri((subst,va as Var(a,Ta),vb as Var(b,Tb), Var (P, _)),i,si,T,yik) = | 
| 545 |         let val cx = Thm.cterm_of ctxt (eta_Var(("X"^si,0),T))
 | |
| 60789 | 546 | val cb = Thm.cterm_of ctxt vb | 
| 60790 | 547 |             val cy = Thm.cterm_of ctxt (eta_Var(("Y"^si,0),T))
 | 
| 548 | val cp = Thm.cterm_of ctxt (Pinst(f,rT,eq,k,i,T,yik,aTs)) | |
| 549 | in infer_instantiate ctxt [(a,cx),(b,cy),(P,cp)] subst end; | |
| 0 | 550 | fun mk(c,T::Ts,i,yik) = | 
| 1459 | 551 | let val si = radixstring(26,"a",i) | 
| 60789 | 552 | in case find_subst ctxt T of | 
| 15531 | 553 |              NONE => mk(c,Ts,i-1,eta_Var(("X"^si,0),T)::yik)
 | 
| 554 | | SOME s => let val c' = c RSN (2,ri(s,i,si,T,yik)) | |
| 1459 | 555 |                        in mk(c',Ts,i-1,eta_Var(("Y"^si,0),T)::yik) end
 | 
| 556 | end | |
| 0 | 557 | | mk(c,[],_,_) = c; | 
| 558 | in mk(refl,rev aTs,k-1,[]) end; | |
| 559 | ||
| 60789 | 560 | fun mk_cong_type ctxt (f,T) = | 
| 0 | 561 | let val (aTs,rT) = strip_type T; | 
| 562 | fun find_refl(r::rs) = | |
| 59582 | 563 | let val (Const(eq,eqT),_,_) = dest_red(Thm.concl_of r) | 
| 60789 | 564 | in if Sign.typ_instance (Proof_Context.theory_of ctxt) (rT, hd(binder_types eqT)) | 
| 15531 | 565 | then SOME(r,(eq,body_type eqT)) else find_refl rs | 
| 1459 | 566 | end | 
| 15531 | 567 | | find_refl([]) = NONE; | 
| 0 | 568 | in case find_refl refl_thms of | 
| 60789 | 569 | NONE => [] | SOME(refl) => [mk_cong ctxt (f,aTs,rT) refl] | 
| 0 | 570 | end; | 
| 571 | ||
| 60789 | 572 | fun mk_congs' ctxt f = | 
| 573 | let val T = case Sign.const_type (Proof_Context.theory_of ctxt) f of | |
| 15531 | 574 | NONE => error(f^" not declared") | SOME(T) => T; | 
| 16876 | 575 | val T' = Logic.incr_tvar 9 T; | 
| 60789 | 576 | in mk_cong_type ctxt (Const(f,T'),T') end; | 
| 0 | 577 | |
| 60789 | 578 | val mk_congs = maps o mk_congs'; | 
| 0 | 579 | |
| 60789 | 580 | fun mk_typed_congs ctxt = | 
| 22675 
acf10be7dcca
cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
 wenzelm parents: 
22596diff
changeset | 581 | let | 
| 
acf10be7dcca
cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
 wenzelm parents: 
22596diff
changeset | 582 | fun readfT(f,s) = | 
| 
acf10be7dcca
cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
 wenzelm parents: 
22596diff
changeset | 583 | let | 
| 60789 | 584 | val T = Logic.incr_tvar 9 (Syntax.read_typ ctxt s); | 
| 585 | val t = case Sign.const_type (Proof_Context.theory_of ctxt) f of | |
| 22675 
acf10be7dcca
cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
 wenzelm parents: 
22596diff
changeset | 586 | SOME(_) => Const(f,T) | NONE => Free(f,T) | 
| 
acf10be7dcca
cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
 wenzelm parents: 
22596diff
changeset | 587 | in (t,T) end | 
| 60789 | 588 | in maps (mk_cong_type ctxt o readfT) end; | 
| 0 | 589 | |
| 22675 
acf10be7dcca
cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
 wenzelm parents: 
22596diff
changeset | 590 | end; | 
| 
acf10be7dcca
cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
 wenzelm parents: 
22596diff
changeset | 591 | end; |