| author | wenzelm | 
| Mon, 29 Sep 2008 21:45:44 +0200 | |
| changeset 28417 | 74e580c6f2b5 | 
| parent 26928 | ca87aff1ad2d | 
| child 29265 | 5b4247055bd7 | 
| permissions | -rw-r--r-- | 
| 23150 | 1 | (* Title: HOL/Tools/TFL/rules.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Konrad Slind, Cambridge University Computer Laboratory | |
| 4 | Copyright 1997 University of Cambridge | |
| 5 | ||
| 6 | Emulation of HOL inference rules for TFL | |
| 7 | *) | |
| 8 | ||
| 9 | signature RULES = | |
| 10 | sig | |
| 11 | val dest_thm: thm -> term list * term | |
| 12 | ||
| 13 | (* Inference rules *) | |
| 14 | val REFL: cterm -> thm | |
| 15 | val ASSUME: cterm -> thm | |
| 16 | val MP: thm -> thm -> thm | |
| 17 | val MATCH_MP: thm -> thm -> thm | |
| 18 | val CONJUNCT1: thm -> thm | |
| 19 | val CONJUNCT2: thm -> thm | |
| 20 | val CONJUNCTS: thm -> thm list | |
| 21 | val DISCH: cterm -> thm -> thm | |
| 22 | val UNDISCH: thm -> thm | |
| 23 | val SPEC: cterm -> thm -> thm | |
| 24 | val ISPEC: cterm -> thm -> thm | |
| 25 | val ISPECL: cterm list -> thm -> thm | |
| 26 | val GEN: cterm -> thm -> thm | |
| 27 | val GENL: cterm list -> thm -> thm | |
| 28 | val LIST_CONJ: thm list -> thm | |
| 29 | ||
| 30 | val SYM: thm -> thm | |
| 31 | val DISCH_ALL: thm -> thm | |
| 32 | val FILTER_DISCH_ALL: (term -> bool) -> thm -> thm | |
| 33 | val SPEC_ALL: thm -> thm | |
| 34 | val GEN_ALL: thm -> thm | |
| 35 | val IMP_TRANS: thm -> thm -> thm | |
| 36 | val PROVE_HYP: thm -> thm -> thm | |
| 37 | ||
| 38 | val CHOOSE: cterm * thm -> thm -> thm | |
| 39 | val EXISTS: cterm * cterm -> thm -> thm | |
| 40 | val EXISTL: cterm list -> thm -> thm | |
| 41 | val IT_EXISTS: (cterm*cterm) list -> thm -> thm | |
| 42 | ||
| 43 | val EVEN_ORS: thm list -> thm list | |
| 44 | val DISJ_CASESL: thm -> thm list -> thm | |
| 45 | ||
| 46 | val list_beta_conv: cterm -> cterm list -> thm | |
| 47 | val SUBS: thm list -> thm -> thm | |
| 48 | val simpl_conv: simpset -> thm list -> cterm -> thm | |
| 49 | ||
| 50 | val rbeta: thm -> thm | |
| 51 | (* For debugging my isabelle solver in the conditional rewriter *) | |
| 52 | val term_ref: term list ref | |
| 53 | val thm_ref: thm list ref | |
| 54 | val ss_ref: simpset list ref | |
| 55 | val tracing: bool ref | |
| 56 | val CONTEXT_REWRITE_RULE: term * term list * thm * thm list | |
| 57 | -> thm -> thm * term list | |
| 58 | val RIGHT_ASSOC: thm -> thm | |
| 59 | ||
| 60 | val prove: bool -> cterm * tactic -> thm | |
| 61 | end; | |
| 62 | ||
| 63 | structure Rules: RULES = | |
| 64 | struct | |
| 65 | ||
| 66 | structure S = USyntax; | |
| 67 | structure U = Utils; | |
| 68 | structure D = Dcterm; | |
| 69 | ||
| 70 | ||
| 71 | fun RULES_ERR func mesg = U.ERR {module = "Rules", func = func, mesg = mesg};
 | |
| 72 | ||
| 73 | ||
| 74 | fun cconcl thm = D.drop_prop (#prop (Thm.crep_thm thm)); | |
| 75 | fun chyps thm = map D.drop_prop (#hyps (Thm.crep_thm thm)); | |
| 76 | ||
| 77 | fun dest_thm thm = | |
| 78 |   let val {prop,hyps,...} = Thm.rep_thm thm
 | |
| 79 | in (map HOLogic.dest_Trueprop hyps, HOLogic.dest_Trueprop prop) end | |
| 80 | handle TERM _ => raise RULES_ERR "dest_thm" "missing Trueprop"; | |
| 81 | ||
| 82 | ||
| 83 | (* Inference rules *) | |
| 84 | ||
| 85 | (*--------------------------------------------------------------------------- | |
| 86 | * Equality (one step) | |
| 87 | *---------------------------------------------------------------------------*) | |
| 88 | ||
| 89 | fun REFL tm = Thm.reflexive tm RS meta_eq_to_obj_eq | |
| 90 | handle THM (msg, _, _) => raise RULES_ERR "REFL" msg; | |
| 91 | ||
| 92 | fun SYM thm = thm RS sym | |
| 93 | handle THM (msg, _, _) => raise RULES_ERR "SYM" msg; | |
| 94 | ||
| 95 | fun ALPHA thm ctm1 = | |
| 96 | let | |
| 97 | val ctm2 = Thm.cprop_of thm; | |
| 98 | val ctm2_eq = Thm.reflexive ctm2; | |
| 99 | val ctm1_eq = Thm.reflexive ctm1; | |
| 100 | in Thm.equal_elim (Thm.transitive ctm2_eq ctm1_eq) thm end | |
| 101 | handle THM (msg, _, _) => raise RULES_ERR "ALPHA" msg; | |
| 102 | ||
| 103 | fun rbeta th = | |
| 104 | (case D.strip_comb (cconcl th) of | |
| 105 | (_, [l, r]) => Thm.transitive th (Thm.beta_conversion false r) | |
| 106 | | _ => raise RULES_ERR "rbeta" ""); | |
| 107 | ||
| 108 | ||
| 109 | (*---------------------------------------------------------------------------- | |
| 110 | * Implication and the assumption list | |
| 111 | * | |
| 112 | * Assumptions get stuck on the meta-language assumption list. Implications | |
| 113 | * are in the object language, so discharging an assumption "A" from theorem | |
| 114 | * "B" results in something that looks like "A --> B". | |
| 115 | *---------------------------------------------------------------------------*) | |
| 116 | ||
| 117 | fun ASSUME ctm = Thm.assume (D.mk_prop ctm); | |
| 118 | ||
| 119 | ||
| 120 | (*--------------------------------------------------------------------------- | |
| 121 | * Implication in TFL is -->. Meta-language implication (==>) is only used | |
| 122 | * in the implementation of some of the inference rules below. | |
| 123 | *---------------------------------------------------------------------------*) | |
| 124 | fun MP th1 th2 = th2 RS (th1 RS mp) | |
| 125 | handle THM (msg, _, _) => raise RULES_ERR "MP" msg; | |
| 126 | ||
| 127 | (*forces the first argument to be a proposition if necessary*) | |
| 128 | fun DISCH tm thm = Thm.implies_intr (D.mk_prop tm) thm COMP impI | |
| 129 | handle THM (msg, _, _) => raise RULES_ERR "DISCH" msg; | |
| 130 | ||
| 131 | fun DISCH_ALL thm = fold_rev DISCH (#hyps (Thm.crep_thm thm)) thm; | |
| 132 | ||
| 133 | ||
| 134 | fun FILTER_DISCH_ALL P thm = | |
| 135 | let fun check tm = P (#t (Thm.rep_cterm tm)) | |
| 136 | in foldr (fn (tm,th) => if check tm then DISCH tm th else th) | |
| 137 | thm (chyps thm) | |
| 138 | end; | |
| 139 | ||
| 140 | (* freezeT expensive! *) | |
| 141 | fun UNDISCH thm = | |
| 142 | let val tm = D.mk_prop (#1 (D.dest_imp (cconcl (Thm.freezeT thm)))) | |
| 143 | in Thm.implies_elim (thm RS mp) (ASSUME tm) end | |
| 144 | handle U.ERR _ => raise RULES_ERR "UNDISCH" "" | |
| 145 | | THM _ => raise RULES_ERR "UNDISCH" ""; | |
| 146 | ||
| 147 | fun PROVE_HYP ath bth = MP (DISCH (cconcl ath) bth) ath; | |
| 148 | ||
| 149 | fun IMP_TRANS th1 th2 = th2 RS (th1 RS Thms.imp_trans) | |
| 150 | handle THM (msg, _, _) => raise RULES_ERR "IMP_TRANS" msg; | |
| 151 | ||
| 152 | ||
| 153 | (*---------------------------------------------------------------------------- | |
| 154 | * Conjunction | |
| 155 | *---------------------------------------------------------------------------*) | |
| 156 | ||
| 157 | fun CONJUNCT1 thm = thm RS conjunct1 | |
| 158 | handle THM (msg, _, _) => raise RULES_ERR "CONJUNCT1" msg; | |
| 159 | ||
| 160 | fun CONJUNCT2 thm = thm RS conjunct2 | |
| 161 | handle THM (msg, _, _) => raise RULES_ERR "CONJUNCT2" msg; | |
| 162 | ||
| 163 | fun CONJUNCTS th = CONJUNCTS (CONJUNCT1 th) @ CONJUNCTS (CONJUNCT2 th) handle U.ERR _ => [th]; | |
| 164 | ||
| 165 | fun LIST_CONJ [] = raise RULES_ERR "LIST_CONJ" "empty list" | |
| 166 | | LIST_CONJ [th] = th | |
| 167 | | LIST_CONJ (th :: rst) = MP (MP (conjI COMP (impI RS impI)) th) (LIST_CONJ rst) | |
| 168 | handle THM (msg, _, _) => raise RULES_ERR "LIST_CONJ" msg; | |
| 169 | ||
| 170 | ||
| 171 | (*---------------------------------------------------------------------------- | |
| 172 | * Disjunction | |
| 173 | *---------------------------------------------------------------------------*) | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 174 | local val thy = Thm.theory_of_thm disjI1 | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 175 | val prop = Thm.prop_of disjI1 | 
| 23150 | 176 | val [P,Q] = term_vars prop | 
| 177 | val disj1 = Thm.forall_intr (Thm.cterm_of thy Q) disjI1 | |
| 178 | in | |
| 179 | fun DISJ1 thm tm = thm RS (forall_elim (D.drop_prop tm) disj1) | |
| 180 | handle THM (msg, _, _) => raise RULES_ERR "DISJ1" msg; | |
| 181 | end; | |
| 182 | ||
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 183 | local val thy = Thm.theory_of_thm disjI2 | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 184 | val prop = Thm.prop_of disjI2 | 
| 23150 | 185 | val [P,Q] = term_vars prop | 
| 186 | val disj2 = Thm.forall_intr (Thm.cterm_of thy P) disjI2 | |
| 187 | in | |
| 188 | fun DISJ2 tm thm = thm RS (forall_elim (D.drop_prop tm) disj2) | |
| 189 | handle THM (msg, _, _) => raise RULES_ERR "DISJ2" msg; | |
| 190 | end; | |
| 191 | ||
| 192 | ||
| 193 | (*---------------------------------------------------------------------------- | |
| 194 | * | |
| 195 | * A1 |- M1, ..., An |- Mn | |
| 196 | * --------------------------------------------------- | |
| 197 | * [A1 |- M1 \/ ... \/ Mn, ..., An |- M1 \/ ... \/ Mn] | |
| 198 | * | |
| 199 | *---------------------------------------------------------------------------*) | |
| 200 | ||
| 201 | ||
| 202 | fun EVEN_ORS thms = | |
| 203 | let fun blue ldisjs [] _ = [] | |
| 204 | | blue ldisjs (th::rst) rdisjs = | |
| 205 | let val tail = tl rdisjs | |
| 206 | val rdisj_tl = D.list_mk_disj tail | |
| 207 | in fold_rev DISJ2 ldisjs (DISJ1 th rdisj_tl) | |
| 208 | :: blue (ldisjs @ [cconcl th]) rst tail | |
| 209 | end handle U.ERR _ => [fold_rev DISJ2 ldisjs th] | |
| 210 | in blue [] thms (map cconcl thms) end; | |
| 211 | ||
| 212 | ||
| 213 | (*---------------------------------------------------------------------------- | |
| 214 | * | |
| 215 | * A |- P \/ Q B,P |- R C,Q |- R | |
| 216 | * --------------------------------------------------- | |
| 217 | * A U B U C |- R | |
| 218 | * | |
| 219 | *---------------------------------------------------------------------------*) | |
| 220 | ||
| 221 | fun DISJ_CASES th1 th2 th3 = | |
| 222 | let | |
| 223 | val c = D.drop_prop (cconcl th1); | |
| 224 | val (disj1, disj2) = D.dest_disj c; | |
| 225 | val th2' = DISCH disj1 th2; | |
| 226 | val th3' = DISCH disj2 th3; | |
| 227 | in | |
| 228 | th3' RS (th2' RS (th1 RS Thms.tfl_disjE)) | |
| 229 | handle THM (msg, _, _) => raise RULES_ERR "DISJ_CASES" msg | |
| 230 | end; | |
| 231 | ||
| 232 | ||
| 233 | (*----------------------------------------------------------------------------- | |
| 234 | * | |
| 235 | * |- A1 \/ ... \/ An [A1 |- M, ..., An |- M] | |
| 236 | * --------------------------------------------------- | |
| 237 | * |- M | |
| 238 | * | |
| 239 | * Note. The list of theorems may be all jumbled up, so we have to | |
| 240 | * first organize it to align with the first argument (the disjunctive | |
| 241 | * theorem). | |
| 242 | *---------------------------------------------------------------------------*) | |
| 243 | ||
| 244 | fun organize eq = (* a bit slow - analogous to insertion sort *) | |
| 245 | let fun extract a alist = | |
| 246 | let fun ex (_,[]) = raise RULES_ERR "organize" "not a permutation.1" | |
| 247 | | ex(left,h::t) = if (eq h a) then (h,rev left@t) else ex(h::left,t) | |
| 248 | in ex ([],alist) | |
| 249 | end | |
| 250 | fun place [] [] = [] | |
| 251 | | place (a::rst) alist = | |
| 252 | let val (item,next) = extract a alist | |
| 253 | in item::place rst next | |
| 254 | end | |
| 255 | | place _ _ = raise RULES_ERR "organize" "not a permutation.2" | |
| 256 | in place | |
| 257 | end; | |
| 258 | (* freezeT expensive! *) | |
| 259 | fun DISJ_CASESL disjth thl = | |
| 260 | let val c = cconcl disjth | |
| 261 | fun eq th atm = exists (fn t => HOLogic.dest_Trueprop t | |
| 262 | aconv term_of atm) | |
| 263 | (#hyps(rep_thm th)) | |
| 264 | val tml = D.strip_disj c | |
| 265 | fun DL th [] = raise RULES_ERR "DISJ_CASESL" "no cases" | |
| 266 | | DL th [th1] = PROVE_HYP th th1 | |
| 267 | | DL th [th1,th2] = DISJ_CASES th th1 th2 | |
| 268 | | DL th (th1::rst) = | |
| 269 | let val tm = #2(D.dest_disj(D.drop_prop(cconcl th))) | |
| 270 | in DISJ_CASES th th1 (DL (ASSUME tm) rst) end | |
| 271 | in DL (Thm.freezeT disjth) (organize eq tml thl) | |
| 272 | end; | |
| 273 | ||
| 274 | ||
| 275 | (*---------------------------------------------------------------------------- | |
| 276 | * Universals | |
| 277 | *---------------------------------------------------------------------------*) | |
| 278 | local (* this is fragile *) | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 279 | val thy = Thm.theory_of_thm spec | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 280 | val prop = Thm.prop_of spec | 
| 23150 | 281 | val x = hd (tl (term_vars prop)) | 
| 282 | val cTV = ctyp_of thy (type_of x) | |
| 283 | val gspec = forall_intr (cterm_of thy x) spec | |
| 284 | in | |
| 285 | fun SPEC tm thm = | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 286 | let val gspec' = instantiate ([(cTV, Thm.ctyp_of_term tm)], []) gspec | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 287 | in thm RS (forall_elim tm gspec') end | 
| 23150 | 288 | end; | 
| 289 | ||
| 290 | fun SPEC_ALL thm = fold SPEC (#1(D.strip_forall(cconcl thm))) thm; | |
| 291 | ||
| 292 | val ISPEC = SPEC | |
| 293 | val ISPECL = fold ISPEC; | |
| 294 | ||
| 295 | (* Not optimized! Too complicated. *) | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 296 | local val thy = Thm.theory_of_thm allI | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 297 | val prop = Thm.prop_of allI | 
| 23150 | 298 | val [P] = add_term_vars (prop, []) | 
| 299 | fun cty_theta s = map (fn (i, (S, ty)) => (ctyp_of s (TVar (i, S)), ctyp_of s ty)) | |
| 300 | fun ctm_theta s = map (fn (i, (_, tm2)) => | |
| 301 | let val ctm2 = cterm_of s tm2 | |
| 302 | in (cterm_of s (Var(i,#T(rep_cterm ctm2))), ctm2) | |
| 303 | end) | |
| 304 | fun certify s (ty_theta,tm_theta) = | |
| 305 | (cty_theta s (Vartab.dest ty_theta), | |
| 306 | ctm_theta s (Vartab.dest tm_theta)) | |
| 307 | in | |
| 308 | fun GEN v th = | |
| 309 | let val gth = forall_intr v th | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 310 | val thy = Thm.theory_of_thm gth | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 311 |        val Const("all",_)$Abs(x,ty,rst) = Thm.prop_of gth
 | 
| 23150 | 312 | val P' = Abs(x,ty, HOLogic.dest_Trueprop rst) (* get rid of trueprop *) | 
| 313 | val theta = Pattern.match thy (P,P') (Vartab.empty, Vartab.empty); | |
| 314 | val allI2 = instantiate (certify thy theta) allI | |
| 315 | val thm = Thm.implies_elim allI2 gth | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 316 | val tp $ (A $ Abs(_,_,M)) = Thm.prop_of thm | 
| 23150 | 317 | val prop' = tp $ (A $ Abs(x,ty,M)) | 
| 318 | in ALPHA thm (cterm_of thy prop') | |
| 319 | end | |
| 320 | end; | |
| 321 | ||
| 322 | val GENL = fold_rev GEN; | |
| 323 | ||
| 324 | fun GEN_ALL thm = | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 325 | let val thy = Thm.theory_of_thm thm | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 326 | val prop = Thm.prop_of thm | 
| 23150 | 327 | val tycheck = cterm_of thy | 
| 328 | val vlist = map tycheck (add_term_vars (prop, [])) | |
| 329 | in GENL vlist thm | |
| 330 | end; | |
| 331 | ||
| 332 | ||
| 333 | fun MATCH_MP th1 th2 = | |
| 334 | if (D.is_forall (D.drop_prop(cconcl th1))) | |
| 335 | then MATCH_MP (th1 RS spec) th2 | |
| 336 | else MP th1 th2; | |
| 337 | ||
| 338 | ||
| 339 | (*---------------------------------------------------------------------------- | |
| 340 | * Existentials | |
| 341 | *---------------------------------------------------------------------------*) | |
| 342 | ||
| 343 | ||
| 344 | ||
| 345 | (*--------------------------------------------------------------------------- | |
| 346 | * Existential elimination | |
| 347 | * | |
| 348 | * A1 |- ?x.t[x] , A2, "t[v]" |- t' | |
| 349 | * ------------------------------------ (variable v occurs nowhere) | |
| 350 | * A1 u A2 |- t' | |
| 351 | * | |
| 352 | *---------------------------------------------------------------------------*) | |
| 353 | ||
| 354 | fun CHOOSE (fvar, exth) fact = | |
| 355 | let | |
| 356 | val lam = #2 (D.dest_comb (D.drop_prop (cconcl exth))) | |
| 357 | val redex = D.capply lam fvar | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 358 | val thy = Thm.theory_of_cterm redex | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 359 | val t$u = Thm.term_of redex | 
| 23150 | 360 | val residue = Thm.cterm_of thy (Term.betapply (t, u)) | 
| 361 | in | |
| 362 | GEN fvar (DISCH residue fact) RS (exth RS Thms.choose_thm) | |
| 363 | handle THM (msg, _, _) => raise RULES_ERR "CHOOSE" msg | |
| 364 | end; | |
| 365 | ||
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 366 | local val thy = Thm.theory_of_thm exI | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 367 | val prop = Thm.prop_of exI | 
| 23150 | 368 | val [P,x] = term_vars prop | 
| 369 | in | |
| 370 | fun EXISTS (template,witness) thm = | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 371 | let val thy = Thm.theory_of_thm thm | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 372 | val prop = Thm.prop_of thm | 
| 23150 | 373 | val P' = cterm_of thy P | 
| 374 | val x' = cterm_of thy x | |
| 375 | val abstr = #2 (D.dest_comb template) | |
| 376 | in | |
| 377 | thm RS (cterm_instantiate[(P',abstr), (x',witness)] exI) | |
| 378 | handle THM (msg, _, _) => raise RULES_ERR "EXISTS" msg | |
| 379 | end | |
| 380 | end; | |
| 381 | ||
| 382 | (*---------------------------------------------------------------------------- | |
| 383 | * | |
| 384 | * A |- M | |
| 385 | * ------------------- [v_1,...,v_n] | |
| 386 | * A |- ?v1...v_n. M | |
| 387 | * | |
| 388 | *---------------------------------------------------------------------------*) | |
| 389 | ||
| 390 | fun EXISTL vlist th = | |
| 391 | fold_rev (fn v => fn thm => EXISTS(D.mk_exists(v,cconcl thm), v) thm) | |
| 392 | vlist th; | |
| 393 | ||
| 394 | ||
| 395 | (*---------------------------------------------------------------------------- | |
| 396 | * | |
| 397 | * A |- M[x_1,...,x_n] | |
| 398 | * ---------------------------- [(x |-> y)_1,...,(x |-> y)_n] | |
| 399 | * A |- ?y_1...y_n. M | |
| 400 | * | |
| 401 | *---------------------------------------------------------------------------*) | |
| 402 | (* Could be improved, but needs "subst_free" for certified terms *) | |
| 403 | ||
| 404 | fun IT_EXISTS blist th = | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 405 | let val thy = Thm.theory_of_thm th | 
| 23150 | 406 | val tych = cterm_of thy | 
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 407 | val blist' = map (pairself Thm.term_of) blist | 
| 23150 | 408 |        fun ex v M  = cterm_of thy (S.mk_exists{Bvar=v,Body = M})
 | 
| 409 | ||
| 410 | in | |
| 411 | fold_rev (fn (b as (r1,r2)) => fn thm => | |
| 412 | EXISTS(ex r2 (subst_free [b] | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 413 | (HOLogic.dest_Trueprop(Thm.prop_of thm))), tych r1) | 
| 23150 | 414 | thm) | 
| 415 | blist' th | |
| 416 | end; | |
| 417 | ||
| 418 | (*--------------------------------------------------------------------------- | |
| 419 | * Faster version, that fails for some as yet unknown reason | |
| 420 | * fun IT_EXISTS blist th = | |
| 421 |  *    let val {thy,...} = rep_thm th
 | |
| 422 | * val tych = cterm_of thy | |
| 423 | * fun detype (x,y) = ((#t o rep_cterm) x, (#t o rep_cterm) y) | |
| 424 | * in | |
| 425 | * fold (fn (b as (r1,r2), thm) => | |
| 426 | * EXISTS(D.mk_exists(r2, tych(subst_free[detype b](#t(rep_cterm(cconcl thm))))), | |
| 427 | * r1) thm) blist th | |
| 428 | * end; | |
| 429 | *---------------------------------------------------------------------------*) | |
| 430 | ||
| 431 | (*---------------------------------------------------------------------------- | |
| 432 | * Rewriting | |
| 433 | *---------------------------------------------------------------------------*) | |
| 434 | ||
| 435 | fun SUBS thl = | |
| 436 | rewrite_rule (map (fn th => th RS eq_reflection handle THM _ => th) thl); | |
| 437 | ||
| 438 | val rew_conv = MetaSimplifier.rewrite_cterm (true, false, false) (K (K NONE)); | |
| 439 | ||
| 440 | fun simpl_conv ss thl ctm = | |
| 441 | rew_conv (ss addsimps thl) ctm RS meta_eq_to_obj_eq; | |
| 442 | ||
| 443 | ||
| 444 | val RIGHT_ASSOC = rewrite_rule [Thms.disj_assoc]; | |
| 445 | ||
| 446 | ||
| 447 | ||
| 448 | (*--------------------------------------------------------------------------- | |
| 449 | * TERMINATION CONDITION EXTRACTION | |
| 450 | *---------------------------------------------------------------------------*) | |
| 451 | ||
| 452 | ||
| 453 | (* Object language quantifier, i.e., "!" *) | |
| 454 | fun Forall v M = S.mk_forall{Bvar=v, Body=M};
 | |
| 455 | ||
| 456 | ||
| 457 | (* Fragile: it's a cong if it is not "R y x ==> cut f R x y = f y" *) | |
| 458 | fun is_cong thm = | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 459 | case (Thm.prop_of thm) | 
| 23150 | 460 |      of (Const("==>",_)$(Const("Trueprop",_)$ _) $
 | 
| 26750 | 461 |          (Const("==",_) $ (Const (@{const_name "Wellfounded.cut"},_) $ f $ R $ a $ x) $ _)) => false
 | 
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 462 | | _ => true; | 
| 23150 | 463 | |
| 464 | ||
| 465 | fun dest_equal(Const ("==",_) $
 | |
| 466 |                (Const ("Trueprop",_) $ lhs)
 | |
| 467 |                $ (Const ("Trueprop",_) $ rhs)) = {lhs=lhs, rhs=rhs}
 | |
| 468 |   | dest_equal(Const ("==",_) $ lhs $ rhs)  = {lhs=lhs, rhs=rhs}
 | |
| 469 | | dest_equal tm = S.dest_eq tm; | |
| 470 | ||
| 471 | fun get_lhs tm = #lhs(dest_equal (HOLogic.dest_Trueprop tm)); | |
| 472 | ||
| 473 | fun dest_all used (Const("all",_) $ (a as Abs _)) = S.dest_abs used a
 | |
| 474 | | dest_all _ _ = raise RULES_ERR "dest_all" "not a !!"; | |
| 475 | ||
| 476 | val is_all = can (dest_all []); | |
| 477 | ||
| 478 | fun strip_all used fm = | |
| 479 | if (is_all fm) | |
| 480 |    then let val ({Bvar, Body}, used') = dest_all used fm
 | |
| 481 | val (bvs, core, used'') = strip_all used' Body | |
| 482 | in ((Bvar::bvs), core, used'') | |
| 483 | end | |
| 484 | else ([], fm, used); | |
| 485 | ||
| 486 | fun break_all(Const("all",_) $ Abs (_,_,body)) = body
 | |
| 487 | | break_all _ = raise RULES_ERR "break_all" "not a !!"; | |
| 488 | ||
| 489 | fun list_break_all(Const("all",_) $ Abs (s,ty,body)) =
 | |
| 490 | let val (L,core) = list_break_all body | |
| 491 | in ((s,ty)::L, core) | |
| 492 | end | |
| 493 | | list_break_all tm = ([],tm); | |
| 494 | ||
| 495 | (*--------------------------------------------------------------------------- | |
| 496 | * Rename a term of the form | |
| 497 | * | |
| 498 | * !!x1 ...xn. x1=M1 ==> ... ==> xn=Mn | |
| 499 | * ==> ((%v1...vn. Q) x1 ... xn = g x1 ... xn. | |
| 500 | * to one of | |
| 501 | * | |
| 502 | * !!v1 ... vn. v1=M1 ==> ... ==> vn=Mn | |
| 503 | * ==> ((%v1...vn. Q) v1 ... vn = g v1 ... vn. | |
| 504 | * | |
| 505 | * This prevents name problems in extraction, and helps the result to read | |
| 506 | * better. There is a problem with varstructs, since they can introduce more | |
| 507 | * than n variables, and some extra reasoning needs to be done. | |
| 508 | *---------------------------------------------------------------------------*) | |
| 509 | ||
| 510 | fun get ([],_,L) = rev L | |
| 511 | | get (ant::rst,n,L) = | |
| 512 | case (list_break_all ant) | |
| 513 | of ([],_) => get (rst, n+1,L) | |
| 514 | | (vlist,body) => | |
| 515 | let val eq = Logic.strip_imp_concl body | |
| 516 | val (f,args) = S.strip_comb (get_lhs eq) | |
| 517 | val (vstrl,_) = S.strip_abs f | |
| 518 | val names = | |
| 519 | Name.variant_list (add_term_names(body, [])) (map (#1 o dest_Free) vstrl) | |
| 520 | in get (rst, n+1, (names,n)::L) end | |
| 521 | handle TERM _ => get (rst, n+1, L) | |
| 522 | | U.ERR _ => get (rst, n+1, L); | |
| 523 | ||
| 524 | (* Note: rename_params_rule counts from 1, not 0 *) | |
| 525 | fun rename thm = | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 526 | let val thy = Thm.theory_of_thm thm | 
| 23150 | 527 | val tych = cterm_of thy | 
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 528 | val ants = Logic.strip_imp_prems (Thm.prop_of thm) | 
| 23150 | 529 | val news = get (ants,1,[]) | 
| 530 | in | |
| 531 | fold rename_params_rule news thm | |
| 532 | end; | |
| 533 | ||
| 534 | ||
| 535 | (*--------------------------------------------------------------------------- | |
| 536 | * Beta-conversion to the rhs of an equation (taken from hol90/drule.sml) | |
| 537 | *---------------------------------------------------------------------------*) | |
| 538 | ||
| 539 | fun list_beta_conv tm = | |
| 540 | let fun rbeta th = Thm.transitive th (beta_conversion false (#2(D.dest_eq(cconcl th)))) | |
| 541 | fun iter [] = Thm.reflexive tm | |
| 542 | | iter (v::rst) = rbeta (combination(iter rst) (Thm.reflexive v)) | |
| 543 | in iter end; | |
| 544 | ||
| 545 | ||
| 546 | (*--------------------------------------------------------------------------- | |
| 547 | * Trace information for the rewriter | |
| 548 | *---------------------------------------------------------------------------*) | |
| 549 | val term_ref = ref[] : term list ref | |
| 550 | val ss_ref = ref [] : simpset list ref; | |
| 551 | val thm_ref = ref [] : thm list ref; | |
| 552 | val tracing = ref false; | |
| 553 | ||
| 554 | fun say s = if !tracing then writeln s else (); | |
| 555 | ||
| 556 | fun print_thms s L = | |
| 26928 | 557 | say (cat_lines (s :: map Display.string_of_thm L)); | 
| 23150 | 558 | |
| 559 | fun print_cterms s L = | |
| 26928 | 560 | say (cat_lines (s :: map Display.string_of_cterm L)); | 
| 23150 | 561 | |
| 562 | ||
| 563 | (*--------------------------------------------------------------------------- | |
| 564 | * General abstraction handlers, should probably go in USyntax. | |
| 565 | *---------------------------------------------------------------------------*) | |
| 566 | fun mk_aabs (vstr, body) = | |
| 567 |   S.mk_abs {Bvar = vstr, Body = body}
 | |
| 568 |   handle U.ERR _ => S.mk_pabs {varstruct = vstr, body = body};
 | |
| 569 | ||
| 570 | fun list_mk_aabs (vstrl,tm) = | |
| 571 | fold_rev (fn vstr => fn tm => mk_aabs(vstr,tm)) vstrl tm; | |
| 572 | ||
| 573 | fun dest_aabs used tm = | |
| 574 |    let val ({Bvar,Body}, used') = S.dest_abs used tm
 | |
| 575 | in (Bvar, Body, used') end | |
| 576 | handle U.ERR _ => | |
| 577 |      let val {varstruct, body, used} = S.dest_pabs used tm
 | |
| 578 | in (varstruct, body, used) end; | |
| 579 | ||
| 580 | fun strip_aabs used tm = | |
| 581 | let val (vstr, body, used') = dest_aabs used tm | |
| 582 | val (bvs, core, used'') = strip_aabs used' body | |
| 583 | in (vstr::bvs, core, used'') end | |
| 584 | handle U.ERR _ => ([], tm, used); | |
| 585 | ||
| 586 | fun dest_combn tm 0 = (tm,[]) | |
| 587 | | dest_combn tm n = | |
| 588 |      let val {Rator,Rand} = S.dest_comb tm
 | |
| 589 | val (f,rands) = dest_combn Rator (n-1) | |
| 590 | in (f,Rand::rands) | |
| 591 | end; | |
| 592 | ||
| 593 | ||
| 594 | ||
| 595 | ||
| 596 | local fun dest_pair M = let val {fst,snd} = S.dest_pair M in (fst,snd) end
 | |
| 597 | fun mk_fst tm = | |
| 598 |           let val ty as Type("*", [fty,sty]) = type_of tm
 | |
| 599 |           in  Const ("fst", ty --> fty) $ tm  end
 | |
| 600 | fun mk_snd tm = | |
| 601 |           let val ty as Type("*", [fty,sty]) = type_of tm
 | |
| 602 |           in  Const ("snd", ty --> sty) $ tm  end
 | |
| 603 | in | |
| 604 | fun XFILL tych x vstruct = | |
| 605 | let fun traverse p xocc L = | |
| 606 | if (is_Free p) | |
| 607 | then tych xocc::L | |
| 608 | else let val (p1,p2) = dest_pair p | |
| 609 | in traverse p1 (mk_fst xocc) (traverse p2 (mk_snd xocc) L) | |
| 610 | end | |
| 611 | in | |
| 612 | traverse vstruct x [] | |
| 613 | end end; | |
| 614 | ||
| 615 | (*--------------------------------------------------------------------------- | |
| 616 | * Replace a free tuple (vstr) by a universally quantified variable (a). | |
| 617 | * Note that the notion of "freeness" for a tuple is different than for a | |
| 618 | * variable: if variables in the tuple also occur in any other place than | |
| 619 | * an occurrences of the tuple, they aren't "free" (which is thus probably | |
| 620 | * the wrong word to use). | |
| 621 | *---------------------------------------------------------------------------*) | |
| 622 | ||
| 623 | fun VSTRUCT_ELIM tych a vstr th = | |
| 624 | let val L = S.free_vars_lr vstr | |
| 625 | val bind1 = tych (HOLogic.mk_Trueprop (HOLogic.mk_eq(a,vstr))) | |
| 626 | val thm1 = implies_intr bind1 (SUBS [SYM(assume bind1)] th) | |
| 627 | val thm2 = forall_intr_list (map tych L) thm1 | |
| 628 | val thm3 = forall_elim_list (XFILL tych a vstr) thm2 | |
| 629 | in refl RS | |
| 630 | rewrite_rule [Thm.symmetric (surjective_pairing RS eq_reflection)] thm3 | |
| 631 | end; | |
| 632 | ||
| 633 | fun PGEN tych a vstr th = | |
| 634 | let val a1 = tych a | |
| 635 | val vstr1 = tych vstr | |
| 636 | in | |
| 637 | forall_intr a1 | |
| 638 | (if (is_Free vstr) | |
| 639 | then cterm_instantiate [(vstr1,a1)] th | |
| 640 | else VSTRUCT_ELIM tych a vstr th) | |
| 641 | end; | |
| 642 | ||
| 643 | ||
| 644 | (*--------------------------------------------------------------------------- | |
| 645 | * Takes apart a paired beta-redex, looking like "(\(x,y).N) vstr", into | |
| 646 | * | |
| 647 | * (([x,y],N),vstr) | |
| 648 | *---------------------------------------------------------------------------*) | |
| 649 | fun dest_pbeta_redex used M n = | |
| 650 | let val (f,args) = dest_combn M n | |
| 651 | val dummy = dest_aabs used f | |
| 652 | in (strip_aabs used f,args) | |
| 653 | end; | |
| 654 | ||
| 655 | fun pbeta_redex M n = can (U.C (dest_pbeta_redex []) n) M; | |
| 656 | ||
| 657 | fun dest_impl tm = | |
| 658 | let val ants = Logic.strip_imp_prems tm | |
| 659 | val eq = Logic.strip_imp_concl tm | |
| 660 | in (ants,get_lhs eq) | |
| 661 | end; | |
| 662 | ||
| 663 | fun restricted t = isSome (S.find_term | |
| 26750 | 664 |                             (fn (Const(@{const_name "Wellfounded.cut"},_)) =>true | _ => false)
 | 
| 23150 | 665 | t) | 
| 666 | ||
| 667 | fun CONTEXT_REWRITE_RULE (func, G, cut_lemma, congs) th = | |
| 668 | let val globals = func::G | |
| 669 | val ss0 = Simplifier.theory_context (Thm.theory_of_thm th) empty_ss | |
| 670 | val pbeta_reduce = simpl_conv ss0 [split_conv RS eq_reflection]; | |
| 671 | val tc_list = ref[]: term list ref | |
| 672 | val dummy = term_ref := [] | |
| 673 | val dummy = thm_ref := [] | |
| 674 | val dummy = ss_ref := [] | |
| 675 | val cut_lemma' = cut_lemma RS eq_reflection | |
| 676 | fun prover used ss thm = | |
| 677 | let fun cong_prover ss thm = | |
| 678 | let val dummy = say "cong_prover:" | |
| 679 | val cntxt = MetaSimplifier.prems_of_ss ss | |
| 680 | val dummy = print_thms "cntxt:" cntxt | |
| 681 | val dummy = say "cong rule:" | |
| 26928 | 682 | val dummy = say (Display.string_of_thm thm) | 
| 23150 | 683 | val dummy = thm_ref := (thm :: !thm_ref) | 
| 684 | val dummy = ss_ref := (ss :: !ss_ref) | |
| 685 | (* Unquantified eliminate *) | |
| 686 | fun uq_eliminate (thm,imp,thy) = | |
| 687 | let val tych = cterm_of thy | |
| 688 | val dummy = print_cterms "To eliminate:" [tych imp] | |
| 689 | val ants = map tych (Logic.strip_imp_prems imp) | |
| 690 | val eq = Logic.strip_imp_concl imp | |
| 691 | val lhs = tych(get_lhs eq) | |
| 692 | val ss' = MetaSimplifier.add_prems (map ASSUME ants) ss | |
| 693 | val lhs_eq_lhs1 = MetaSimplifier.rewrite_cterm (false,true,false) (prover used) ss' lhs | |
| 694 | handle U.ERR _ => Thm.reflexive lhs | |
| 695 | val dummy = print_thms "proven:" [lhs_eq_lhs1] | |
| 696 | val lhs_eq_lhs2 = implies_intr_list ants lhs_eq_lhs1 | |
| 697 | val lhs_eeq_lhs2 = lhs_eq_lhs2 RS meta_eq_to_obj_eq | |
| 698 | in | |
| 699 | lhs_eeq_lhs2 COMP thm | |
| 700 | end | |
| 701 | fun pq_eliminate (thm,thy,vlist,imp_body,lhs_eq) = | |
| 702 | let val ((vstrl, _, used'), args) = dest_pbeta_redex used lhs_eq (length vlist) | |
| 703 | val dummy = forall (op aconv) (ListPair.zip (vlist, args)) | |
| 704 | orelse error "assertion failed in CONTEXT_REWRITE_RULE" | |
| 705 | val imp_body1 = subst_free (ListPair.zip (args, vstrl)) | |
| 706 | imp_body | |
| 707 | val tych = cterm_of thy | |
| 708 | val ants1 = map tych (Logic.strip_imp_prems imp_body1) | |
| 709 | val eq1 = Logic.strip_imp_concl imp_body1 | |
| 710 | val Q = get_lhs eq1 | |
| 711 | val QeqQ1 = pbeta_reduce (tych Q) | |
| 712 | val Q1 = #2(D.dest_eq(cconcl QeqQ1)) | |
| 713 | val ss' = MetaSimplifier.add_prems (map ASSUME ants1) ss | |
| 714 | val Q1eeqQ2 = MetaSimplifier.rewrite_cterm (false,true,false) (prover used') ss' Q1 | |
| 715 | handle U.ERR _ => Thm.reflexive Q1 | |
| 716 | val Q2 = #2 (Logic.dest_equals (Thm.prop_of Q1eeqQ2)) | |
| 717 | val Q3 = tych(list_comb(list_mk_aabs(vstrl,Q2),vstrl)) | |
| 718 | val Q2eeqQ3 = Thm.symmetric(pbeta_reduce Q3 RS eq_reflection) | |
| 719 | val thA = Thm.transitive(QeqQ1 RS eq_reflection) Q1eeqQ2 | |
| 720 | val QeeqQ3 = Thm.transitive thA Q2eeqQ3 handle THM _ => | |
| 721 | ((Q2eeqQ3 RS meta_eq_to_obj_eq) | |
| 722 | RS ((thA RS meta_eq_to_obj_eq) RS trans)) | |
| 723 | RS eq_reflection | |
| 724 | val impth = implies_intr_list ants1 QeeqQ3 | |
| 725 | val impth1 = impth RS meta_eq_to_obj_eq | |
| 726 | (* Need to abstract *) | |
| 727 | val ant_th = U.itlist2 (PGEN tych) args vstrl impth1 | |
| 728 | in ant_th COMP thm | |
| 729 | end | |
| 730 | fun q_eliminate (thm,imp,thy) = | |
| 731 | let val (vlist, imp_body, used') = strip_all used imp | |
| 732 | val (ants,Q) = dest_impl imp_body | |
| 733 | in if (pbeta_redex Q) (length vlist) | |
| 734 | then pq_eliminate (thm,thy,vlist,imp_body,Q) | |
| 735 | else | |
| 736 | let val tych = cterm_of thy | |
| 737 | val ants1 = map tych ants | |
| 738 | val ss' = MetaSimplifier.add_prems (map ASSUME ants1) ss | |
| 739 | val Q_eeq_Q1 = MetaSimplifier.rewrite_cterm | |
| 740 | (false,true,false) (prover used') ss' (tych Q) | |
| 741 | handle U.ERR _ => Thm.reflexive (tych Q) | |
| 742 | val lhs_eeq_lhs2 = implies_intr_list ants1 Q_eeq_Q1 | |
| 743 | val lhs_eq_lhs2 = lhs_eeq_lhs2 RS meta_eq_to_obj_eq | |
| 744 | val ant_th = forall_intr_list(map tych vlist)lhs_eq_lhs2 | |
| 745 | in | |
| 746 | ant_th COMP thm | |
| 747 | end end | |
| 748 | ||
| 749 | fun eliminate thm = | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 750 | case Thm.prop_of thm | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 751 |                of Const("==>",_) $ imp $ _ =>
 | 
| 23150 | 752 | eliminate | 
| 753 | (if not(is_all imp) | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 754 | then uq_eliminate (thm, imp, Thm.theory_of_thm thm) | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 755 | else q_eliminate (thm, imp, Thm.theory_of_thm thm)) | 
| 23150 | 756 | (* Assume that the leading constant is ==, *) | 
| 757 | | _ => thm (* if it is not a ==> *) | |
| 758 | in SOME(eliminate (rename thm)) end | |
| 759 | handle U.ERR _ => NONE (* FIXME handle THM as well?? *) | |
| 760 | ||
| 761 | fun restrict_prover ss thm = | |
| 762 | let val dummy = say "restrict_prover:" | |
| 763 | val cntxt = rev(MetaSimplifier.prems_of_ss ss) | |
| 764 | val dummy = print_thms "cntxt:" cntxt | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 765 | val thy = Thm.theory_of_thm thm | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 766 |               val Const("==>",_) $ (Const("Trueprop",_) $ A) $ _ = Thm.prop_of thm
 | 
| 23150 | 767 | fun genl tm = let val vlist = subtract (op aconv) globals | 
| 768 | (add_term_frees(tm,[])) | |
| 769 | in fold_rev Forall vlist tm | |
| 770 | end | |
| 771 | (*-------------------------------------------------------------- | |
| 772 | * This actually isn't quite right, since it will think that | |
| 773 | * not-fully applied occs. of "f" in the context mean that the | |
| 774 | * current call is nested. The real solution is to pass in a | |
| 775 | * term "f v1..vn" which is a pattern that any full application | |
| 776 | * of "f" will match. | |
| 777 | *-------------------------------------------------------------*) | |
| 778 | val func_name = #1(dest_Const func) | |
| 779 | fun is_func (Const (name,_)) = (name = func_name) | |
| 780 | | is_func _ = false | |
| 781 | val rcontext = rev cntxt | |
| 782 | val cncl = HOLogic.dest_Trueprop o Thm.prop_of | |
| 783 | val antl = case rcontext of [] => [] | |
| 784 | | _ => [S.list_mk_conj(map cncl rcontext)] | |
| 785 | val TC = genl(S.list_mk_imp(antl, A)) | |
| 786 | val dummy = print_cterms "func:" [cterm_of thy func] | |
| 787 | val dummy = print_cterms "TC:" | |
| 788 | [cterm_of thy (HOLogic.mk_Trueprop TC)] | |
| 789 | val dummy = tc_list := (TC :: !tc_list) | |
| 790 | val nestedp = isSome (S.find_term is_func TC) | |
| 791 | val dummy = if nestedp then say "nested" else say "not_nested" | |
| 792 | val dummy = term_ref := ([func,TC]@(!term_ref)) | |
| 793 | val th' = if nestedp then raise RULES_ERR "solver" "nested function" | |
| 794 | else let val cTC = cterm_of thy | |
| 795 | (HOLogic.mk_Trueprop TC) | |
| 796 | in case rcontext of | |
| 797 | [] => SPEC_ALL(ASSUME cTC) | |
| 798 | | _ => MP (SPEC_ALL (ASSUME cTC)) | |
| 799 | (LIST_CONJ rcontext) | |
| 800 | end | |
| 801 | val th'' = th' RS thm | |
| 802 | in SOME (th'') | |
| 803 | end handle U.ERR _ => NONE (* FIXME handle THM as well?? *) | |
| 804 | in | |
| 805 | (if (is_cong thm) then cong_prover else restrict_prover) ss thm | |
| 806 | end | |
| 807 | val ctm = cprop_of th | |
| 808 | val names = add_term_names (term_of ctm, []) | |
| 809 | val th1 = MetaSimplifier.rewrite_cterm(false,true,false) | |
| 810 | (prover names) (ss0 addsimps [cut_lemma'] addeqcongs congs) ctm | |
| 811 | val th2 = equal_elim th1 th | |
| 812 | in | |
| 813 | (th2, List.filter (not o restricted) (!tc_list)) | |
| 814 | end; | |
| 815 | ||
| 816 | ||
| 817 | fun prove strict (ptm, tac) = | |
| 818 | let | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 819 | val thy = Thm.theory_of_cterm ptm; | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
23150diff
changeset | 820 | val t = Thm.term_of ptm; | 
| 23150 | 821 | val ctxt = ProofContext.init thy |> Variable.auto_fixes t; | 
| 822 | in | |
| 823 | if strict then Goal.prove ctxt [] [] t (K tac) | |
| 824 | else Goal.prove ctxt [] [] t (K tac) | |
| 825 | handle ERROR msg => (warning msg; raise RULES_ERR "prove" msg) | |
| 826 | end; | |
| 827 | ||
| 828 | end; |