| author | wenzelm | 
| Fri, 27 Jan 2006 19:03:19 +0100 | |
| changeset 18812 | a4554848b59e | 
| parent 17959 | 8db36a108213 | 
| child 19064 | bf19cc5a7899 | 
| permissions | -rw-r--r-- | 
| 14620 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 1 | (* Title: HOL/Import/replay.ML | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 2 | ID: $Id$ | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 3 | Author: Sebastian Skalberg (TU Muenchen) | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 4 | *) | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 5 | |
| 14516 | 6 | structure Replay = | 
| 7 | struct | |
| 8 | ||
| 9 | structure P = ProofKernel | |
| 10 | ||
| 11 | open ProofKernel | |
| 12 | ||
| 13 | exception REPLAY of string * string | |
| 14 | fun ERR f mesg = REPLAY (f,mesg) | |
| 15 | fun NY f = raise ERR f "NOT YET!" | |
| 16 | ||
| 17 | fun replay_proof int_thms thyname thmname prf thy = | |
| 18 | let | |
| 19 | fun rp (PRefl tm) thy = P.REFL tm thy | |
| 20 | | rp (PInstT(p,lambda)) thy = | |
| 21 | let | |
| 22 | val (thy',th) = rp' p thy | |
| 23 | in | |
| 24 | P.INST_TYPE lambda th thy' | |
| 25 | end | |
| 26 | | rp (PSubst(prfs,ctxt,prf)) thy = | |
| 27 | let | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 28 | val (thy',ths) = foldr (fn (p,(thy,ths)) => | 
| 14516 | 29 | let | 
| 30 | val (thy',th) = rp' p thy | |
| 31 | in | |
| 32 | (thy',th::ths) | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 33 | end) (thy,[]) prfs | 
| 14516 | 34 | val (thy'',th) = rp' prf thy' | 
| 35 | in | |
| 36 | P.SUBST ths ctxt th thy'' | |
| 37 | end | |
| 38 | | rp (PAbs(prf,v)) thy = | |
| 39 | let | |
| 40 | val (thy',th) = rp' prf thy | |
| 41 | in | |
| 42 | P.ABS v th thy' | |
| 43 | end | |
| 44 | | rp (PDisch(prf,tm)) thy = | |
| 45 | let | |
| 46 | val (thy',th) = rp' prf thy | |
| 47 | in | |
| 48 | P.DISCH tm th thy' | |
| 49 | end | |
| 50 | | rp (PMp(prf1,prf2)) thy = | |
| 51 | let | |
| 52 | val (thy1,th1) = rp' prf1 thy | |
| 53 | val (thy2,th2) = rp' prf2 thy1 | |
| 54 | in | |
| 55 | P.MP th1 th2 thy2 | |
| 56 | end | |
| 57 | | rp (PHyp tm) thy = P.ASSUME tm thy | |
| 58 | | rp (PDef(seg,name,rhs)) thy = | |
| 59 | (case P.get_def seg name rhs thy of | |
| 15531 | 60 | (thy',SOME res) => (thy',res) | 
| 61 | | (thy',NONE) => | |
| 14516 | 62 | if seg = thyname | 
| 63 | then P.new_definition seg name rhs thy' | |
| 17322 | 64 | 		 else raise ERR "replay_proof" ("Too late for term definition: "^seg^" != "^thyname))
 | 
| 14516 | 65 | | rp (POracle(tg,asl,c)) thy = (*P.mk_oracle_thm tg (map be_contract asl,c)*) NY "ORACLE" | 
| 66 | | rp (PSpec(prf,tm)) thy = | |
| 67 | let | |
| 68 | val (thy',th) = rp' prf thy | |
| 69 | in | |
| 70 | P.SPEC tm th thy' | |
| 71 | end | |
| 72 | | rp (PInst(prf,theta)) thy = | |
| 73 | let | |
| 74 | val (thy',th) = rp' prf thy | |
| 75 | in | |
| 76 | P.INST theta th thy' | |
| 77 | end | |
| 78 | | rp (PGen(prf,v)) thy = | |
| 79 | let | |
| 80 | val (thy',th) = rp' prf thy | |
| 17592 | 81 | val p = P.GEN v th thy' | 
| 14516 | 82 | in | 
| 17592 | 83 | p | 
| 14516 | 84 | end | 
| 85 | | rp (PGenAbs(prf,opt,vl)) thy = | |
| 86 | let | |
| 87 | val (thy',th) = rp' prf thy | |
| 88 | in | |
| 89 | P.GEN_ABS opt vl th thy' | |
| 90 | end | |
| 91 | | rp (PImpAS(prf1,prf2)) thy = | |
| 92 | let | |
| 93 | val (thy1,th1) = rp' prf1 thy | |
| 94 | val (thy2,th2) = rp' prf2 thy1 | |
| 95 | in | |
| 96 | P.IMP_ANTISYM th1 th2 thy2 | |
| 97 | end | |
| 98 | | rp (PSym prf) thy = | |
| 99 | let | |
| 100 | val (thy1,th) = rp' prf thy | |
| 101 | in | |
| 102 | P.SYM th thy1 | |
| 103 | end | |
| 104 | | rp (PTrans(prf1,prf2)) thy = | |
| 105 | let | |
| 106 | val (thy1,th1) = rp' prf1 thy | |
| 107 | val (thy2,th2) = rp' prf2 thy1 | |
| 108 | in | |
| 109 | P.TRANS th1 th2 thy2 | |
| 110 | end | |
| 111 | | rp (PComb(prf1,prf2)) thy = | |
| 112 | let | |
| 113 | val (thy1,th1) = rp' prf1 thy | |
| 114 | val (thy2,th2) = rp' prf2 thy1 | |
| 115 | in | |
| 116 | P.COMB th1 th2 thy2 | |
| 117 | end | |
| 118 | | rp (PEqMp(prf1,prf2)) thy = | |
| 119 | let | |
| 120 | val (thy1,th1) = rp' prf1 thy | |
| 121 | val (thy2,th2) = rp' prf2 thy1 | |
| 122 | in | |
| 123 | P.EQ_MP th1 th2 thy2 | |
| 124 | end | |
| 125 | | rp (PEqImp prf) thy = | |
| 126 | let | |
| 127 | val (thy',th) = rp' prf thy | |
| 128 | in | |
| 129 | P.EQ_IMP_RULE th thy' | |
| 130 | end | |
| 131 | | rp (PExists(prf,ex,wit)) thy = | |
| 132 | let | |
| 133 | val (thy',th) = rp' prf thy | |
| 134 | in | |
| 135 | P.EXISTS ex wit th thy' | |
| 136 | end | |
| 137 | | rp (PChoose(v,prf1,prf2)) thy = | |
| 138 | let | |
| 139 | val (thy1,th1) = rp' prf1 thy | |
| 140 | val (thy2,th2) = rp' prf2 thy1 | |
| 141 | in | |
| 142 | P.CHOOSE v th1 th2 thy2 | |
| 143 | end | |
| 144 | | rp (PConj(prf1,prf2)) thy = | |
| 145 | let | |
| 146 | val (thy1,th1) = rp' prf1 thy | |
| 147 | val (thy2,th2) = rp' prf2 thy1 | |
| 148 | in | |
| 149 | P.CONJ th1 th2 thy2 | |
| 150 | end | |
| 151 | | rp (PConjunct1 prf) thy = | |
| 152 | let | |
| 153 | val (thy',th) = rp' prf thy | |
| 154 | in | |
| 155 | P.CONJUNCT1 th thy' | |
| 156 | end | |
| 157 | | rp (PConjunct2 prf) thy = | |
| 158 | let | |
| 159 | val (thy',th) = rp' prf thy | |
| 160 | in | |
| 161 | P.CONJUNCT2 th thy' | |
| 162 | end | |
| 163 | | rp (PDisj1(prf,tm)) thy = | |
| 164 | let | |
| 165 | val (thy',th) = rp' prf thy | |
| 166 | in | |
| 167 | P.DISJ1 th tm thy' | |
| 168 | end | |
| 169 | | rp (PDisj2(prf,tm)) thy = | |
| 170 | let | |
| 171 | val (thy',th) = rp' prf thy | |
| 172 | in | |
| 173 | P.DISJ2 tm th thy' | |
| 174 | end | |
| 175 | | rp (PDisjCases(prf,prf1,prf2)) thy = | |
| 176 | let | |
| 177 | val (thy',th) = rp' prf thy | |
| 178 | val (thy1,th1) = rp' prf1 thy' | |
| 179 | val (thy2,th2) = rp' prf2 thy1 | |
| 180 | in | |
| 181 | P.DISJ_CASES th th1 th2 thy2 | |
| 182 | end | |
| 183 | | rp (PNotI prf) thy = | |
| 184 | let | |
| 185 | val (thy',th) = rp' prf thy | |
| 186 | in | |
| 187 | P.NOT_INTRO th thy' | |
| 188 | end | |
| 189 | | rp (PNotE prf) thy = | |
| 190 | let | |
| 191 | val (thy',th) = rp' prf thy | |
| 192 | in | |
| 193 | P.NOT_ELIM th thy' | |
| 194 | end | |
| 195 | | rp (PContr(prf,tm)) thy = | |
| 196 | let | |
| 197 | val (thy',th) = rp' prf thy | |
| 198 | in | |
| 199 | P.CCONTR tm th thy' | |
| 200 | end | |
| 201 | | rp (PTmSpec _) _ = raise ERR "rp" "Shouldn't reach here (PTmSpec)" | |
| 202 | | rp (PTyDef _) _ = raise ERR "rp" "Shouldn't reach here (PTyDef)" | |
| 203 | | rp (PTyIntro _) _ = raise ERR "rp" "Shouldn't reach here (PTyIntro)" | |
| 204 | | rp PDisk _ = raise ERR "rp" "Shouldn't reach here (PDisk)" | |
| 205 | | rp _ _ = raise ERR "rp" "What the hell is this? Which case did I forget?" | |
| 206 | and rp' p thy = | |
| 207 | let | |
| 208 | val pc = content_of p | |
| 209 | in | |
| 210 | case pc of | |
| 211 | PDisk => (case disk_info_of p of | |
| 15531 | 212 | SOME(thyname',thmname) => | 
| 14516 | 213 | (case Int.fromString thmname of | 
| 214 | SOME i => | |
| 215 | if thyname' = thyname | |
| 216 | then | |
| 217 | (case Array.sub(int_thms,i-1) of | |
| 15531 | 218 | NONE => | 
| 14516 | 219 | let | 
| 220 | val (thy',th) = rp' (snd (import_proof thyname' thmname thy) thy) thy | |
| 15531 | 221 | val _ = Array.update(int_thms,i-1,SOME th) | 
| 14516 | 222 | in | 
| 223 | (thy',th) | |
| 224 | end | |
| 15531 | 225 | | SOME th => (thy,th)) | 
| 14516 | 226 | 				       else raise ERR "replay_proof" ("Library " ^ thyname' ^ " should be built before " ^ thyname ^ " (" ^ thmname ^ ")")
 | 
| 227 | | NONE => | |
| 228 | (case P.get_thm thyname' thmname thy of | |
| 17594 | 229 | (thy',SOME res) => (thy',res) | 
| 15531 | 230 | | (thy',NONE) => | 
| 14516 | 231 | if thyname' = thyname | 
| 232 | then | |
| 233 | let | |
| 234 | 						    val _ = writeln ("Found no " ^ thmname ^ " theorem, replaying...")
 | |
| 235 | val (f_opt,prf) = import_proof thyname' thmname thy' | |
| 236 | val prf = prf thy' | |
| 237 | val (thy',th) = replay_proof int_thms thyname' thmname prf thy' | |
| 17322 | 238 |                                                     val _ = writeln ("Successfully finished replaying "^thmname^" !")
 | 
| 14516 | 239 | in | 
| 240 | case content_of prf of | |
| 241 | PTmSpec _ => (thy',th) | |
| 242 | | PTyDef _ => (thy',th) | |
| 243 | | PTyIntro _ => (thy',th) | |
| 244 | | _ => P.store_thm thyname' thmname th thy' | |
| 245 | end | |
| 246 | 					    else raise ERR "replay_proof" ("Library " ^ thyname' ^ " should be built before " ^ thyname ^ " (" ^ thmname ^ ")")))
 | |
| 15531 | 247 | | NONE => raise ERR "rp'.PDisk" "Not enough information") | 
| 14516 | 248 | | PAxm(name,c) => | 
| 249 | (case P.get_axiom thyname name thy of | |
| 15531 | 250 | (thy',SOME res) => (thy',res) | 
| 251 | | (thy',NONE) => P.new_axiom name c thy') | |
| 14516 | 252 | | PTmSpec(seg,names,prf') => | 
| 253 | let | |
| 254 | val (thy',th) = rp' prf' thy | |
| 255 | in | |
| 256 | P.new_specification seg thmname names th thy' | |
| 257 | end | |
| 258 | | PTyDef(seg,name,prf') => | |
| 259 | let | |
| 260 | val (thy',th) = rp' prf' thy | |
| 261 | in | |
| 262 | P.new_type_definition seg thmname name th thy' | |
| 263 | end | |
| 264 | | PTyIntro(seg,name,abs_name,rep_name,P,t,prf') => | |
| 265 | let | |
| 266 | val (thy',th) = rp' prf' thy | |
| 267 | in | |
| 268 | P.type_introduction seg thmname name abs_name rep_name (P,t) th thy' | |
| 269 | end | |
| 270 | | _ => rp pc thy | |
| 271 | end | |
| 272 | in | |
| 17959 | 273 | rp' prf thy handle e => (writeln "Exception in replay_proof"; OldGoals.print_exn e) | 
| 14516 | 274 | end | 
| 275 | ||
| 276 | fun setup_int_thms thyname thy = | |
| 277 | let | |
| 14620 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 278 | val fname = | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 279 | case P.get_proof_dir thyname thy of | 
| 15531 | 280 | 		SOME p => OS.Path.joinDirFile {dir=p,file=OS.Path.joinBaseExt{base = "facts",ext=SOME "lst"}}
 | 
| 281 | | NONE => error "Cannot find proof files" | |
| 14620 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14516diff
changeset | 282 | val is = TextIO.openIn fname | 
| 14516 | 283 | val (num_int_thms,facts) = | 
| 284 | let | |
| 285 | fun get_facts facts = | |
| 286 | case TextIO.inputLine is of | |
| 287 | "" => (case facts of | |
| 17440 
df77edc4f5d0
fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
 obua parents: 
17322diff
changeset | 288 | i::facts => (valOf (Int.fromString i),map P.protect_factname (rev facts)) | 
| 14516 | 289 | | _ => raise ERR "replay_thm" "Bad facts.lst file") | 
| 290 | | fact => get_facts ((String.substring(fact,0,String.size fact -1 ))::facts) | |
| 291 | in | |
| 292 | get_facts [] | |
| 293 | end | |
| 294 | val _ = TextIO.closeIn is | |
| 15531 | 295 | val int_thms = Array.array(num_int_thms,NONE:thm option) | 
| 14516 | 296 | in | 
| 297 | (int_thms,facts) | |
| 298 | end | |
| 299 | ||
| 300 | fun import_single_thm thyname int_thms thmname thy = | |
| 301 | let | |
| 302 | fun replay_fact (thmname,thy) = | |
| 303 | let | |
| 304 | val prf = mk_proof PDisk | |
| 305 | val _ = set_disk_info_of prf thyname thmname | |
| 17594 | 306 |                 val _ = writeln ("Replaying "^thmname^" ...")
 | 
| 17592 | 307 | val p = fst (replay_proof int_thms thyname thmname prf thy) | 
| 14516 | 308 | in | 
| 17592 | 309 | p | 
| 14516 | 310 | end | 
| 311 | in | |
| 312 | replay_fact (thmname,thy) | |
| 313 | end | |
| 314 | ||
| 315 | fun import_thms thyname int_thms thmnames thy = | |
| 316 | let | |
| 317 | fun replay_fact (thy,thmname) = | |
| 318 | let | |
| 17592 | 319 | val prf = mk_proof PDisk | 
| 17594 | 320 | val _ = set_disk_info_of prf thyname thmname | 
| 321 |                 val _ = writeln ("Replaying "^thmname^" ...")
 | |
| 17592 | 322 | val p = fst (replay_proof int_thms thyname thmname prf thy) | 
| 14516 | 323 | in | 
| 17592 | 324 | p | 
| 14516 | 325 | end | 
| 15570 | 326 | val res_thy = Library.foldl replay_fact (thy,thmnames) | 
| 14516 | 327 | in | 
| 328 | res_thy | |
| 329 | end | |
| 330 | ||
| 331 | fun import_thm thyname thmname thy = | |
| 332 | let | |
| 333 | val int_thms = fst (setup_int_thms thyname thy) | |
| 334 | fun replay_fact (thmname,thy) = | |
| 335 | let | |
| 17592 | 336 | val prf = mk_proof PDisk | 
| 17594 | 337 | val _ = set_disk_info_of prf thyname thmname | 
| 338 |                 val _ = writeln ("Replaying "^thmname^" ...")
 | |
| 17592 | 339 | val p = fst (replay_proof int_thms thyname thmname prf thy) | 
| 340 | in | |
| 341 | p | |
| 14516 | 342 | end | 
| 343 | in | |
| 344 | replay_fact (thmname,thy) | |
| 345 | end | |
| 346 | ||
| 17959 | 347 | end |