| author | haftmann | 
| Sun, 05 Jun 2005 14:41:23 +0200 | |
| changeset 16276 | 3a50bf1f04d0 | 
| parent 15794 | 5de27a5fc5ed | 
| child 16363 | c686a606dfba | 
| permissions | -rw-r--r-- | 
| 14620 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 1 | (* Title: HOL/Import/proof_kernel.ML | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 2 | ID: $Id$ | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 3 | Author: Sebastian Skalberg (TU Muenchen) | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 4 | *) | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 5 | |
| 14516 | 6 | signature ProofKernel = | 
| 7 | sig | |
| 8 | type hol_type | |
| 9 | type tag | |
| 10 | type term | |
| 11 | type thm | |
| 12 |     type ('a,'b) subst
 | |
| 13 | ||
| 14 | type proof_info | |
| 15 | datatype proof = Proof of proof_info * proof_content | |
| 16 | and proof_content | |
| 17 | = PRefl of term | |
| 18 | | PInstT of proof * (hol_type,hol_type) subst | |
| 19 | | PSubst of proof list * term * proof | |
| 20 | | PAbs of proof * term | |
| 21 | | PDisch of proof * term | |
| 22 | | PMp of proof * proof | |
| 23 | | PHyp of term | |
| 24 | | PAxm of string * term | |
| 25 | | PDef of string * string * term | |
| 26 | | PTmSpec of string * string list * proof | |
| 27 | | PTyDef of string * string * proof | |
| 28 | | PTyIntro of string * string * string * string * term * term * proof | |
| 29 | | POracle of tag * term list * term | |
| 30 | | PDisk | |
| 31 | | PSpec of proof * term | |
| 32 | | PInst of proof * (term,term) subst | |
| 33 | | PGen of proof * term | |
| 34 | | PGenAbs of proof * term option * term list | |
| 35 | | PImpAS of proof * proof | |
| 36 | | PSym of proof | |
| 37 | | PTrans of proof * proof | |
| 38 | | PComb of proof * proof | |
| 39 | | PEqMp of proof * proof | |
| 40 | | PEqImp of proof | |
| 41 | | PExists of proof * term * term | |
| 42 | | PChoose of term * proof * proof | |
| 43 | | PConj of proof * proof | |
| 44 | | PConjunct1 of proof | |
| 45 | | PConjunct2 of proof | |
| 46 | | PDisj1 of proof * term | |
| 47 | | PDisj2 of proof * term | |
| 48 | | PDisjCases of proof * proof * proof | |
| 49 | | PNotI of proof | |
| 50 | | PNotE of proof | |
| 51 | | PContr of proof * term | |
| 52 | ||
| 53 | exception PK of string * string | |
| 54 | ||
| 14620 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 55 | val get_proof_dir: string -> theory -> string option | 
| 14516 | 56 | val debug : bool ref | 
| 57 | val disk_info_of : proof -> (string * string) option | |
| 58 | val set_disk_info_of : proof -> string -> string -> unit | |
| 59 | val mk_proof : proof_content -> proof | |
| 60 | val content_of : proof -> proof_content | |
| 61 | val import_proof : string -> string -> theory -> (theory -> term) option * (theory -> proof) | |
| 62 | ||
| 63 | val rewrite_hol4_term: Term.term -> theory -> Thm.thm | |
| 64 | ||
| 65 | val type_of : term -> hol_type | |
| 66 | ||
| 67 | val get_thm : string -> string -> theory -> (theory * thm option) | |
| 68 | val get_def : string -> string -> term -> theory -> (theory * thm option) | |
| 69 | val get_axiom: string -> string -> theory -> (theory * thm option) | |
| 70 | ||
| 71 | val store_thm : string -> string -> thm -> theory -> theory * thm | |
| 72 | ||
| 73 | val to_isa_thm : thm -> (term * term) list * Thm.thm | |
| 74 | val to_isa_term: term -> Term.term | |
| 75 | ||
| 76 | val REFL : term -> theory -> theory * thm | |
| 77 | val ASSUME : term -> theory -> theory * thm | |
| 78 | val INST_TYPE : (hol_type,hol_type) subst -> thm -> theory -> theory * thm | |
| 79 | val INST : (term,term)subst -> thm -> theory -> theory * thm | |
| 80 | val EQ_MP : thm -> thm -> theory -> theory * thm | |
| 81 | val EQ_IMP_RULE : thm -> theory -> theory * thm | |
| 82 | val SUBST : thm list -> term -> thm -> theory -> theory * thm | |
| 83 | val DISJ_CASES : thm -> thm -> thm -> theory -> theory * thm | |
| 84 | val DISJ1: thm -> term -> theory -> theory * thm | |
| 85 | val DISJ2: term -> thm -> theory -> theory * thm | |
| 86 | val IMP_ANTISYM: thm -> thm -> theory -> theory * thm | |
| 87 | val SYM : thm -> theory -> theory * thm | |
| 88 | val MP : thm -> thm -> theory -> theory * thm | |
| 89 | val GEN : term -> thm -> theory -> theory * thm | |
| 90 | val CHOOSE : term -> thm -> thm -> theory -> theory * thm | |
| 91 | val EXISTS : term -> term -> thm -> theory -> theory * thm | |
| 92 | val ABS : term -> thm -> theory -> theory * thm | |
| 93 | val GEN_ABS : term option -> term list -> thm -> theory -> theory * thm | |
| 94 | val TRANS : thm -> thm -> theory -> theory * thm | |
| 95 | val CCONTR : term -> thm -> theory -> theory * thm | |
| 96 | val CONJ : thm -> thm -> theory -> theory * thm | |
| 97 | val CONJUNCT1: thm -> theory -> theory * thm | |
| 98 | val CONJUNCT2: thm -> theory -> theory * thm | |
| 99 | val NOT_INTRO: thm -> theory -> theory * thm | |
| 100 | val NOT_ELIM : thm -> theory -> theory * thm | |
| 101 | val SPEC : term -> thm -> theory -> theory * thm | |
| 102 | val COMB : thm -> thm -> theory -> theory * thm | |
| 103 | val DISCH: term -> thm -> theory -> theory * thm | |
| 104 | ||
| 105 | val type_introduction: string -> string -> string -> string -> string -> term * term -> thm -> theory -> theory * thm | |
| 106 | ||
| 107 | val new_definition : string -> string -> term -> theory -> theory * thm | |
| 108 | val new_specification : string -> string -> string list -> thm -> theory -> theory * thm | |
| 109 | val new_type_definition : string -> string -> string -> thm -> theory -> theory * thm | |
| 110 | val new_axiom : string -> term -> theory -> theory * thm | |
| 111 | ||
| 112 | end | |
| 113 | ||
| 114 | structure ProofKernel :> ProofKernel = | |
| 115 | struct | |
| 116 | type hol_type = Term.typ | |
| 117 | type term = Term.term | |
| 118 | datatype tag = Tag of string list | |
| 119 | type ('a,'b) subst = ('a * 'b) list
 | |
| 120 | datatype thm = HOLThm of (Term.term * Term.term) list * Thm.thm | |
| 121 | ||
| 122 | datatype proof_info | |
| 123 |   = Info of {disk_info: (string * string) option ref}
 | |
| 124 | ||
| 125 | datatype proof = Proof of proof_info * proof_content | |
| 126 | and proof_content | |
| 127 | = PRefl of term | |
| 128 | | PInstT of proof * (hol_type,hol_type) subst | |
| 129 | | PSubst of proof list * term * proof | |
| 130 | | PAbs of proof * term | |
| 131 | | PDisch of proof * term | |
| 132 | | PMp of proof * proof | |
| 133 | | PHyp of term | |
| 134 | | PAxm of string * term | |
| 135 | | PDef of string * string * term | |
| 136 | | PTmSpec of string * string list * proof | |
| 137 | | PTyDef of string * string * proof | |
| 138 | | PTyIntro of string * string * string * string * term * term * proof | |
| 139 | | POracle of tag * term list * term | |
| 140 | | PDisk | |
| 141 | | PSpec of proof * term | |
| 142 | | PInst of proof * (term,term) subst | |
| 143 | | PGen of proof * term | |
| 144 | | PGenAbs of proof * term option * term list | |
| 145 | | PImpAS of proof * proof | |
| 146 | | PSym of proof | |
| 147 | | PTrans of proof * proof | |
| 148 | | PComb of proof * proof | |
| 149 | | PEqMp of proof * proof | |
| 150 | | PEqImp of proof | |
| 151 | | PExists of proof * term * term | |
| 152 | | PChoose of term * proof * proof | |
| 153 | | PConj of proof * proof | |
| 154 | | PConjunct1 of proof | |
| 155 | | PConjunct2 of proof | |
| 156 | | PDisj1 of proof * term | |
| 157 | | PDisj2 of proof * term | |
| 158 | | PDisjCases of proof * proof * proof | |
| 159 | | PNotI of proof | |
| 160 | | PNotE of proof | |
| 161 | | PContr of proof * term | |
| 162 | ||
| 163 | exception PK of string * string | |
| 164 | fun ERR f mesg = PK (f,mesg) | |
| 165 | ||
| 166 | fun print_exn e = | |
| 167 | case e of | |
| 168 | 	PK (m,s) => (writeln ("PK (" ^ m ^ "): " ^ s); raise e)
 | |
| 169 | | _ => Goals.print_exn e | |
| 170 | ||
| 171 | (* Compatibility. *) | |
| 172 | ||
| 14685 | 173 | fun mk_syn thy c = | 
| 174 | if Syntax.is_identifier c andalso not (Syntax.is_keyword (Theory.syn_of thy) c) then NoSyn | |
| 175 | else Syntax.literal c | |
| 14516 | 176 | |
| 14673 | 177 | fun quotename c = | 
| 14685 | 178 | if Syntax.is_identifier c andalso not (OuterSyntax.is_keyword c) then c else quote c | 
| 14516 | 179 | |
| 180 | fun smart_string_of_cterm ct = | |
| 181 | let | |
| 182 | 	val {sign,t,T,...} = rep_cterm ct
 | |
| 183 | (* Hack to avoid parse errors with Trueprop *) | |
| 184 | val ct = (cterm_of sign (HOLogic.dest_Trueprop t) | |
| 185 | handle TERM _ => ct) | |
| 186 | fun match cu = t aconv (term_of cu) | |
| 187 | fun G 0 = I | |
| 188 | | G 1 = Library.setmp show_types true | |
| 189 | | G 2 = Library.setmp show_all_types true | |
| 190 | 	  | G _ = error ("ProofKernel.smart_string_of_cterm internal error: " ^ (G 2 string_of_cterm ct))
 | |
| 191 | fun F sh_br n = | |
| 192 | let | |
| 14980 | 193 | val str = Library.setmp show_brackets sh_br (G n string_of_cterm) ct | 
| 14516 | 194 | val cu = transform_error (read_cterm sign) (str,T) | 
| 195 | in | |
| 196 | if match cu | |
| 197 | then quote str | |
| 198 | else F false (n+1) | |
| 199 | end | |
| 200 | handle ERROR_MESSAGE mesg => | |
| 201 | if String.isPrefix "Ambiguous" mesg andalso | |
| 202 | not sh_br | |
| 203 | then F true n | |
| 204 | else F false (n+1) | |
| 205 | in | |
| 15647 | 206 | transform_error (Library.setmp print_mode [] (Library.setmp Syntax.ambiguity_is_error true (F false))) 0 | 
| 14516 | 207 | end | 
| 208 | handle ERROR_MESSAGE mesg => | |
| 209 | (writeln "Exception in smart_string_of_cterm!"; | |
| 210 | writeln mesg; | |
| 14980 | 211 | quote (string_of_cterm ct)) | 
| 14516 | 212 | |
| 213 | val smart_string_of_thm = smart_string_of_cterm o cprop_of | |
| 214 | ||
| 215 | fun prth th = writeln ((Library.setmp print_mode [] string_of_thm) th) | |
| 216 | fun prc ct = writeln ((Library.setmp print_mode [] string_of_cterm) ct) | |
| 217 | val prin = Library.setmp print_mode [] prin | |
| 218 | fun pth (HOLThm(ren,thm)) = | |
| 219 | let | |
| 220 | val _ = writeln "Renaming:" | |
| 221 | val _ = app (fn(v,w) => (prin v; writeln " -->"; prin w)) ren | |
| 222 | val _ = prth thm | |
| 223 | in | |
| 224 | () | |
| 225 | end | |
| 226 | ||
| 227 | fun disk_info_of (Proof(Info{disk_info,...},_)) = !disk_info
 | |
| 15531 | 228 | fun mk_proof p = Proof(Info{disk_info = ref NONE},p)
 | 
| 14516 | 229 | fun content_of (Proof(_,p)) = p | 
| 230 | ||
| 231 | fun set_disk_info_of (Proof(Info{disk_info,...},_)) thyname thmname =
 | |
| 15531 | 232 | disk_info := SOME(thyname,thmname) | 
| 14516 | 233 | |
| 234 | structure Lib = | |
| 235 | struct | |
| 236 | fun wrap b e s = String.concat[b,s,e] | |
| 237 | ||
| 238 | fun assoc x = | |
| 239 | let | |
| 240 | 	fun F [] = raise PK("Lib.assoc","Not found")
 | |
| 241 | | F ((x',y)::rest) = if x = x' | |
| 242 | then y | |
| 243 | else F rest | |
| 244 | in | |
| 245 | F | |
| 246 | end | |
| 247 | fun i mem L = | |
| 248 | let fun itr [] = false | |
| 249 | | itr (a::rst) = i=a orelse itr rst | |
| 250 | in itr L end; | |
| 251 | ||
| 252 | fun insert i L = if i mem L then L else i::L | |
| 253 | ||
| 254 | fun mk_set [] = [] | |
| 255 | | mk_set (a::rst) = insert a (mk_set rst) | |
| 256 | ||
| 257 | fun [] union S = S | |
| 258 | | S union [] = S | |
| 259 | | (a::rst) union S2 = rst union (insert a S2) | |
| 260 | ||
| 261 | fun implode_subst [] = [] | |
| 262 | | implode_subst (x::r::rest) = ((x,r)::(implode_subst rest)) | |
| 263 | | implode_subst _ = raise ERR "implode_subst" "malformed substitution list" | |
| 264 | ||
| 265 | fun apboth f (x,y) = (f x,f y) | |
| 266 | end | |
| 267 | open Lib | |
| 268 | ||
| 269 | structure Tag = | |
| 270 | struct | |
| 271 | val empty_tag = Tag [] | |
| 272 | fun read name = Tag [name] | |
| 273 | fun merge (Tag tag1) (Tag tag2) = Tag (Lib.union(tag1,tag2)) | |
| 274 | end | |
| 275 | ||
| 276 | (* Acutal code. *) | |
| 277 | ||
| 278 | fun get_segment thyname l = (Lib.assoc "s" l | |
| 279 | handle PK _ => thyname) | |
| 14518 
c3019a66180f
Added a number of explicit type casts and delayed evaluations (all seemingly
 skalberg parents: 
14516diff
changeset | 280 | val get_name : (string * string) list -> string = Lib.assoc "n" | 
| 14516 | 281 | |
| 282 | local | |
| 283 | open LazyScan | |
| 284 | infix 7 |-- --| | |
| 285 | infix 5 :-- -- ^^ | |
| 286 | infix 3 >> | |
| 287 | infix 0 || | |
| 288 | in | |
| 289 | exception XML of string | |
| 290 | ||
| 291 | datatype xml = Elem of string * (string * string) list * xml list | |
| 292 | datatype XMLtype = XMLty of xml | FullType of hol_type | |
| 293 | datatype XMLterm = XMLtm of xml | FullTerm of term | |
| 294 | ||
| 295 | fun pair x y = (x,y) | |
| 296 | ||
| 297 | fun scan_id toks = | |
| 298 | let | |
| 299 | val (x,toks2) = one Char.isAlpha toks | |
| 300 | val (xs,toks3) = any Char.isAlphaNum toks2 | |
| 301 | in | |
| 302 | (String.implode (x::xs),toks3) | |
| 303 | end | |
| 304 | ||
| 305 | fun scan_string str c = | |
| 306 | let | |
| 307 | fun F [] toks = (c,toks) | |
| 308 | | F (c::cs) toks = | |
| 309 | case LazySeq.getItem toks of | |
| 15531 | 310 | SOME(c',toks') => | 
| 14516 | 311 | if c = c' | 
| 312 | then F cs toks' | |
| 313 | else raise SyntaxError | |
| 15531 | 314 | | NONE => raise SyntaxError | 
| 14516 | 315 | in | 
| 316 | F (String.explode str) | |
| 317 | end | |
| 318 | ||
| 319 | local | |
| 320 | val scan_entity = | |
| 321 | (scan_string "amp;" #"&") | |
| 322 | || scan_string "quot;" #"\"" | |
| 323 | || scan_string "gt;" #">" | |
| 324 | || scan_string "lt;" #"<" | |
| 14620 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 325 | || scan_string "apos;" #"'" | 
| 14516 | 326 | in | 
| 327 | fun scan_nonquote toks = | |
| 328 | case LazySeq.getItem toks of | |
| 15531 | 329 | SOME (c,toks') => | 
| 14516 | 330 | (case c of | 
| 331 | #"\"" => raise SyntaxError | |
| 332 | | #"&" => scan_entity toks' | |
| 333 | | c => (c,toks')) | |
| 15531 | 334 | | NONE => raise SyntaxError | 
| 14516 | 335 | end | 
| 336 | ||
| 337 | val scan_string = $$ #"\"" |-- repeat scan_nonquote --| $$ #"\"" >> | |
| 338 | String.implode | |
| 339 | ||
| 340 | val scan_attribute = scan_id -- $$ #"=" |-- scan_string | |
| 341 | ||
| 342 | val scan_start_of_tag = $$ #"<" |-- scan_id -- | |
| 343 | repeat ($$ #" " |-- scan_attribute) | |
| 344 | ||
| 14518 
c3019a66180f
Added a number of explicit type casts and delayed evaluations (all seemingly
 skalberg parents: 
14516diff
changeset | 345 | (* The evaluation delay introduced through the 'toks' argument is needed | 
| 
c3019a66180f
Added a number of explicit type casts and delayed evaluations (all seemingly
 skalberg parents: 
14516diff
changeset | 346 | for the sake of the SML/NJ (110.9.1) compiler. Either that or an explicit | 
| 
c3019a66180f
Added a number of explicit type casts and delayed evaluations (all seemingly
 skalberg parents: 
14516diff
changeset | 347 | type :-( *) | 
| 
c3019a66180f
Added a number of explicit type casts and delayed evaluations (all seemingly
 skalberg parents: 
14516diff
changeset | 348 | fun scan_end_of_tag toks = ($$ #"/" |-- $$ #">" |-- succeed []) toks | 
| 
c3019a66180f
Added a number of explicit type casts and delayed evaluations (all seemingly
 skalberg parents: 
14516diff
changeset | 349 | |
| 14516 | 350 | val scan_end_tag = $$ #"<" |-- $$ #"/" |-- scan_id --| $$ #">" | 
| 351 | ||
| 352 | fun scan_children id = $$ #">" |-- repeat scan_tag -- scan_end_tag >> | |
| 353 | (fn (chldr,id') => if id = id' | |
| 354 | then chldr | |
| 355 | else raise XML "Tag mismatch") | |
| 356 | and scan_tag toks = | |
| 357 | let | |
| 358 | val ((id,atts),toks2) = scan_start_of_tag toks | |
| 359 | val (chldr,toks3) = (scan_children id || scan_end_of_tag) toks2 | |
| 360 | in | |
| 361 | (Elem (id,atts,chldr),toks3) | |
| 362 | end | |
| 363 | end | |
| 364 | ||
| 365 | val type_of = Term.type_of | |
| 366 | ||
| 367 | val boolT = Type("bool",[])
 | |
| 368 | val propT = Type("prop",[])
 | |
| 369 | ||
| 370 | fun mk_defeq name rhs thy = | |
| 371 | let | |
| 372 | val ty = type_of rhs | |
| 373 | in | |
| 374 | Logic.mk_equals (Const(Sign.intern_const (sign_of thy) name,ty),rhs) | |
| 375 | end | |
| 376 | ||
| 377 | fun mk_teq name rhs thy = | |
| 378 | let | |
| 379 | val ty = type_of rhs | |
| 380 | in | |
| 381 | HOLogic.mk_eq (Const(Sign.intern_const (sign_of thy) name,ty),rhs) | |
| 382 | end | |
| 383 | ||
| 384 | fun intern_const_name thyname const thy = | |
| 385 | case get_hol4_const_mapping thyname const thy of | |
| 15531 | 386 | SOME (_,cname,_) => cname | 
| 387 | | NONE => (case get_hol4_const_renaming thyname const thy of | |
| 388 | SOME cname => Sign.intern_const (sign_of thy) (thyname ^ "." ^ cname) | |
| 389 | | NONE => Sign.intern_const (sign_of thy) (thyname ^ "." ^ const)) | |
| 14516 | 390 | |
| 391 | fun intern_type_name thyname const thy = | |
| 392 | case get_hol4_type_mapping thyname const thy of | |
| 15531 | 393 | SOME (_,cname) => cname | 
| 394 | | NONE => Sign.intern_const (sign_of thy) (thyname ^ "." ^ const) | |
| 14516 | 395 | |
| 396 | fun mk_vartype name = TFree(name,["HOL.type"]) | |
| 397 | fun mk_thy_type thy Thy Tyop Args = Type(intern_type_name Thy Tyop thy,Args) | |
| 398 | ||
| 399 | val mk_var = Free | |
| 400 | ||
| 401 | fun dom_rng (Type("fun",[dom,rng])) = (dom,rng)
 | |
| 402 | | dom_rng _ = raise ERR "dom_rng" "Not a functional type" | |
| 403 | ||
| 404 | fun mk_thy_const thy Thy Name Ty = Const(intern_const_name Thy Name thy,Ty) | |
| 405 | ||
| 406 | local | |
| 407 | fun get_type sg thyname name = | |
| 408 | case Sign.const_type sg name of | |
| 15531 | 409 | SOME ty => ty | 
| 410 | | NONE => raise ERR "get_type" (name ^ ": No such constant") | |
| 14516 | 411 | in | 
| 412 | fun prim_mk_const thy Thy Name = | |
| 413 | let | |
| 414 | val name = intern_const_name Thy Name thy | |
| 415 | val cmaps = HOL4ConstMaps.get thy | |
| 416 | in | |
| 417 | case StringPair.lookup(cmaps,(Thy,Name)) of | |
| 15531 | 418 | SOME(_,_,SOME ty) => Const(name,ty) | 
| 14516 | 419 | | _ => Const(name,get_type (sign_of thy) Thy name) | 
| 420 | end | |
| 421 | end | |
| 422 | ||
| 423 | fun mk_comb(f,a) = f $ a | |
| 424 | fun mk_abs(x,a) = Term.lambda x a | |
| 425 | ||
| 426 | (* Needed for HOL Light *) | |
| 427 | fun protect_tyvarname s = | |
| 428 | let | |
| 429 | fun no_quest s = | |
| 430 | if Char.contains s #"?" | |
| 431 | then String.translate (fn #"?" => "q_" | c => Char.toString c) s | |
| 432 | else s | |
| 433 | fun beg_prime s = | |
| 434 | if String.isPrefix "'" s | |
| 435 | then s | |
| 436 | else "'" ^ s | |
| 437 | in | |
| 438 | s |> no_quest |> beg_prime | |
| 439 | end | |
| 440 | fun protect_varname s = | |
| 441 | let | |
| 442 | fun no_beg_underscore s = | |
| 443 | if String.isPrefix "_" s | |
| 444 | then "dummy" ^ s | |
| 445 | else s | |
| 446 | in | |
| 447 | s |> no_beg_underscore | |
| 448 | end | |
| 449 | ||
| 450 | structure TypeNet = | |
| 451 | struct | |
| 452 | fun get_type_from_index thy thyname types is = | |
| 453 | case Int.fromString is of | |
| 454 | SOME i => (case Array.sub(types,i) of | |
| 455 | FullType ty => ty | |
| 456 | | XMLty xty => | |
| 457 | let | |
| 458 | val ty = get_type_from_xml thy thyname types xty | |
| 459 | val _ = Array.update(types,i,FullType ty) | |
| 460 | in | |
| 461 | ty | |
| 462 | end) | |
| 463 | | NONE => raise ERR "get_type_from_index" "Bad index" | |
| 464 | and get_type_from_xml thy thyname types = | |
| 465 | let | |
| 466 | 	fun gtfx (Elem("tyi",[("i",iS)],[])) =
 | |
| 467 | get_type_from_index thy thyname types iS | |
| 468 | 	  | gtfx (Elem("tyc",atts,[])) =
 | |
| 469 | mk_thy_type thy | |
| 470 | (get_segment thyname atts) | |
| 471 | (get_name atts) | |
| 472 | [] | |
| 473 | 	  | gtfx (Elem("tyv",[("n",s)],[])) = mk_vartype (protect_tyvarname s)
 | |
| 474 | 	  | gtfx (Elem("tya",[],(Elem("tyc",atts,[]))::tys)) =
 | |
| 475 | mk_thy_type thy | |
| 476 | (get_segment thyname atts) | |
| 477 | (get_name atts) | |
| 478 | (map gtfx tys) | |
| 479 | | gtfx _ = raise ERR "get_type" "Bad type" | |
| 480 | in | |
| 481 | gtfx | |
| 482 | end | |
| 483 | ||
| 484 | fun input_types thyname (Elem("tylist",[("i",i)],xtys)) =
 | |
| 485 | let | |
| 486 | 	val types = Array.array(valOf (Int.fromString i),XMLty (Elem("",[],[])))
 | |
| 487 | fun IT _ [] = () | |
| 488 | | IT n (xty::xtys) = | |
| 489 | (Array.update(types,n,XMLty xty); | |
| 490 | IT (n+1) xtys) | |
| 491 | val _ = IT 0 xtys | |
| 492 | in | |
| 493 | types | |
| 494 | end | |
| 495 | | input_types _ _ = raise ERR "input_types" "Bad type list" | |
| 496 | end | |
| 497 | ||
| 498 | structure TermNet = | |
| 499 | struct | |
| 500 | fun get_term_from_index thy thyname types terms is = | |
| 501 | case Int.fromString is of | |
| 502 | SOME i => (case Array.sub(terms,i) of | |
| 503 | FullTerm tm => tm | |
| 504 | | XMLtm xtm => | |
| 505 | let | |
| 506 | val tm = get_term_from_xml thy thyname types terms xtm | |
| 507 | val _ = Array.update(terms,i,FullTerm tm) | |
| 508 | in | |
| 509 | tm | |
| 510 | end) | |
| 511 | | NONE => raise ERR "get_term_from_index" "Bad index" | |
| 512 | and get_term_from_xml thy thyname types terms = | |
| 513 | let | |
| 15531 | 514 | fun get_type [] = NONE | 
| 515 | | get_type [ty] = SOME (TypeNet.get_type_from_xml thy thyname types ty) | |
| 14516 | 516 | | get_type _ = raise ERR "get_term" "Bad type" | 
| 517 | ||
| 518 | 	fun gtfx (Elem("tmv",[("n",name),("t",tyi)],[])) =
 | |
| 519 | mk_var(protect_varname name,TypeNet.get_type_from_index thy thyname types tyi) | |
| 520 | 	  | gtfx (Elem("tmc",atts,[])) =
 | |
| 521 | let | |
| 522 | val segment = get_segment thyname atts | |
| 523 | val name = get_name atts | |
| 524 | in | |
| 525 | mk_thy_const thy segment name (TypeNet.get_type_from_index thy thyname types (Lib.assoc "t" atts)) | |
| 526 | handle PK _ => prim_mk_const thy segment name | |
| 527 | end | |
| 528 | 	  | gtfx (Elem("tma",[("f",tmf),("a",tma)],[])) =
 | |
| 529 | let | |
| 530 | val f = get_term_from_index thy thyname types terms tmf | |
| 531 | val a = get_term_from_index thy thyname types terms tma | |
| 532 | in | |
| 533 | mk_comb(f,a) | |
| 534 | end | |
| 535 | 	  | gtfx (Elem("tml",[("x",tmx),("a",tma)],[])) =
 | |
| 536 | let | |
| 537 | val x = get_term_from_index thy thyname types terms tmx | |
| 538 | val a = get_term_from_index thy thyname types terms tma | |
| 539 | in | |
| 540 | mk_abs(x,a) | |
| 541 | end | |
| 542 | 	  | gtfx (Elem("tmi",[("i",iS)],[])) =
 | |
| 543 | get_term_from_index thy thyname types terms iS | |
| 544 | | gtfx (Elem(tag,_,_)) = | |
| 545 | 	    raise ERR "get_term" ("Not a term: "^tag)
 | |
| 546 | in | |
| 547 | gtfx | |
| 548 | end | |
| 549 | ||
| 550 | fun input_terms thyname types (Elem("tmlist",[("i",i)],xtms)) =
 | |
| 551 | let | |
| 552 | 	val terms = Array.array(valOf(Int.fromString i),XMLtm (Elem("",[],[])))
 | |
| 553 | ||
| 554 | fun IT _ [] = () | |
| 555 | | IT n (xtm::xtms) = | |
| 556 | (Array.update(terms,n,XMLtm xtm); | |
| 557 | IT (n+1) xtms) | |
| 558 | val _ = IT 0 xtms | |
| 559 | in | |
| 560 | terms | |
| 561 | end | |
| 562 | | input_terms _ _ _ = raise ERR "input_terms" "Bad term list" | |
| 563 | end | |
| 564 | ||
| 565 | fun get_proof_dir (thyname:string) thy = | |
| 566 | let | |
| 567 | val import_segment = | |
| 568 | case get_segment2 thyname thy of | |
| 15531 | 569 | SOME seg => seg | 
| 570 | | NONE => get_import_segment thy | |
| 14620 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 571 | 	val defpath = [OS.Path.joinDirFile {dir=getenv "ISABELLE_HOME_USER",file="proofs"}]
 | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 572 | val path = space_explode ":" (getenv "PROOF_DIRS") @ defpath | 
| 15531 | 573 | fun find [] = NONE | 
| 14620 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 574 | | find (p::ps) = | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 575 | (let | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 576 | 		 val dir = OS.Path.joinDirFile {dir = p,file=import_segment}
 | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 577 | in | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 578 | if OS.FileSys.isDir dir | 
| 15531 | 579 | then SOME dir | 
| 14620 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 580 | else find ps | 
| 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 581 | end) handle OS.SysErr _ => find ps | 
| 14516 | 582 | in | 
| 15570 | 583 | 	Option.map (fn p => OS.Path.joinDirFile {dir = p, file = thyname}) (find path)
 | 
| 14516 | 584 | end | 
| 585 | ||
| 586 | fun proof_file_name thyname thmname thy = | |
| 587 | let | |
| 14620 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 588 | val path = case get_proof_dir thyname thy of | 
| 15531 | 589 | SOME p => p | 
| 590 | | NONE => error "Cannot find proof files" | |
| 14516 | 591 | val _ = OS.FileSys.mkDir path handle OS.SysErr _ => () | 
| 592 | in | |
| 14620 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 skalberg parents: 
14518diff
changeset | 593 | 	OS.Path.joinDirFile {dir = path, file = OS.Path.joinBaseExt {base = thmname, ext = SOME "prf"}}
 | 
| 14516 | 594 | end | 
| 595 | ||
| 596 | fun xml_to_proof thyname types terms prf thy = | |
| 597 | let | |
| 598 | val xml_to_hol_type = TypeNet.get_type_from_xml thy thyname types | |
| 599 | val xml_to_term = TermNet.get_term_from_xml thy thyname types terms | |
| 600 | ||
| 601 | fun index_to_term is = | |
| 602 | TermNet.get_term_from_index thy thyname types terms is | |
| 603 | ||
| 604 | 	fun x2p (Elem("prefl",[("i",is)],[])) = mk_proof (PRefl (index_to_term is))
 | |
| 605 | 	  | x2p (Elem("pinstt",[],p::lambda)) =
 | |
| 606 | let | |
| 607 | val p = x2p p | |
| 608 | val lambda = implode_subst (map xml_to_hol_type lambda) | |
| 609 | in | |
| 610 | mk_proof (PInstT(p,lambda)) | |
| 611 | end | |
| 612 | 	  | x2p (Elem("psubst",[("i",is)],prf::prfs)) =
 | |
| 613 | let | |
| 614 | val tm = index_to_term is | |
| 615 | val prf = x2p prf | |
| 616 | val prfs = map x2p prfs | |
| 617 | in | |
| 618 | mk_proof (PSubst(prfs,tm,prf)) | |
| 619 | end | |
| 620 | 	  | x2p (Elem("pabs",[("i",is)],[prf])) =
 | |
| 621 | let | |
| 622 | val p = x2p prf | |
| 623 | val t = index_to_term is | |
| 624 | in | |
| 625 | mk_proof (PAbs (p,t)) | |
| 626 | end | |
| 627 | 	  | x2p (Elem("pdisch",[("i",is)],[prf])) =
 | |
| 628 | let | |
| 629 | val p = x2p prf | |
| 630 | val t = index_to_term is | |
| 631 | in | |
| 632 | mk_proof (PDisch (p,t)) | |
| 633 | end | |
| 634 | 	  | x2p (Elem("pmp",[],[prf1,prf2])) =
 | |
| 635 | let | |
| 636 | val p1 = x2p prf1 | |
| 637 | val p2 = x2p prf2 | |
| 638 | in | |
| 639 | mk_proof (PMp(p1,p2)) | |
| 640 | end | |
| 641 | 	  | x2p (Elem("phyp",[("i",is)],[])) = mk_proof (PHyp (index_to_term is))
 | |
| 642 | 	  | x2p (Elem("paxiom",[("n",n),("i",is)],[])) =
 | |
| 643 | mk_proof (PAxm(n,index_to_term is)) | |
| 644 | 	  | x2p (Elem("pfact",atts,[])) =
 | |
| 645 | let | |
| 646 | val thyname = get_segment thyname atts | |
| 647 | val thmname = get_name atts | |
| 648 | val p = mk_proof PDisk | |
| 649 | val _ = set_disk_info_of p thyname thmname | |
| 650 | in | |
| 651 | p | |
| 652 | end | |
| 653 | 	  | x2p (Elem("pdef",[("s",seg),("n",name),("i",is)],[])) =
 | |
| 654 | mk_proof (PDef(seg,name,index_to_term is)) | |
| 655 | 	  | x2p (Elem("ptmspec",[("s",seg)],p::names)) =
 | |
| 656 | let | |
| 657 | 		val names = map (fn Elem("name",[("n",name)],[]) => name
 | |
| 658 | | _ => raise ERR "x2p" "Bad proof (ptmspec)") names | |
| 659 | in | |
| 660 | mk_proof (PTmSpec(seg,names,x2p p)) | |
| 661 | end | |
| 662 | 	  | x2p (Elem("ptyintro",[("s",seg),("n",name),("a",abs_name),("r",rep_name)],[xP,xt,p])) =
 | |
| 663 | let | |
| 664 | val P = xml_to_term xP | |
| 665 | val t = xml_to_term xt | |
| 666 | in | |
| 667 | mk_proof (PTyIntro(seg,name,abs_name,rep_name,P,t,x2p p)) | |
| 668 | end | |
| 669 | 	  | x2p (Elem("ptydef",[("s",seg),("n",name)],[p])) =
 | |
| 670 | mk_proof (PTyDef(seg,name,x2p p)) | |
| 671 | 	  | x2p (xml as Elem("poracle",[],chldr)) =
 | |
| 672 | let | |
| 673 | 		val (oracles,terms) = Library.partition (fn (Elem("oracle",_,_)) => true | _ => false) chldr
 | |
| 674 | 		val ors = map (fn (Elem("oracle",[("n",name)],[])) => name | xml => raise ERR "x2p" "bad oracle") oracles
 | |
| 675 | val (c,asl) = case terms of | |
| 676 | [] => raise ERR "x2p" "Bad oracle description" | |
| 677 | | (hd::tl) => (hd,tl) | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 678 | val tg = foldr (fn (oracle,tg) => Tag.merge (Tag.read oracle) tg) Tag.empty_tag ors | 
| 14516 | 679 | in | 
| 680 | mk_proof (POracle(tg,map xml_to_term asl,xml_to_term c)) | |
| 681 | end | |
| 682 | 	  | x2p (Elem("pspec",[("i",is)],[prf])) =
 | |
| 683 | let | |
| 684 | val p = x2p prf | |
| 685 | val tm = index_to_term is | |
| 686 | in | |
| 687 | mk_proof (PSpec(p,tm)) | |
| 688 | end | |
| 689 | 	  | x2p (Elem("pinst",[],p::theta)) =
 | |
| 690 | let | |
| 691 | val p = x2p p | |
| 692 | val theta = implode_subst (map xml_to_term theta) | |
| 693 | in | |
| 694 | mk_proof (PInst(p,theta)) | |
| 695 | end | |
| 696 | 	  | x2p (Elem("pgen",[("i",is)],[prf])) =
 | |
| 697 | let | |
| 698 | val p = x2p prf | |
| 699 | val tm = index_to_term is | |
| 700 | in | |
| 701 | mk_proof (PGen(p,tm)) | |
| 702 | end | |
| 703 | 	  | x2p (Elem("pgenabs",[],prf::tms)) =
 | |
| 704 | let | |
| 705 | val p = x2p prf | |
| 706 | val tml = map xml_to_term tms | |
| 707 | in | |
| 15531 | 708 | mk_proof (PGenAbs(p,NONE,tml)) | 
| 14516 | 709 | end | 
| 710 | 	  | x2p (Elem("pgenabs",[("i",is)],prf::tms)) =
 | |
| 711 | let | |
| 712 | val p = x2p prf | |
| 713 | val tml = map xml_to_term tms | |
| 714 | in | |
| 15531 | 715 | mk_proof (PGenAbs(p,SOME (index_to_term is),tml)) | 
| 14516 | 716 | end | 
| 717 | 	  | x2p (Elem("pimpas",[],[prf1,prf2])) =
 | |
| 718 | let | |
| 719 | val p1 = x2p prf1 | |
| 720 | val p2 = x2p prf2 | |
| 721 | in | |
| 722 | mk_proof (PImpAS(p1,p2)) | |
| 723 | end | |
| 724 | 	  | x2p (Elem("psym",[],[prf])) =
 | |
| 725 | let | |
| 726 | val p = x2p prf | |
| 727 | in | |
| 728 | mk_proof (PSym p) | |
| 729 | end | |
| 730 | 	  | x2p (Elem("ptrans",[],[prf1,prf2])) =
 | |
| 731 | let | |
| 732 | val p1 = x2p prf1 | |
| 733 | val p2 = x2p prf2 | |
| 734 | in | |
| 735 | mk_proof (PTrans(p1,p2)) | |
| 736 | end | |
| 737 | 	  | x2p (Elem("pcomb",[],[prf1,prf2])) =
 | |
| 738 | let | |
| 739 | val p1 = x2p prf1 | |
| 740 | val p2 = x2p prf2 | |
| 741 | in | |
| 742 | mk_proof (PComb(p1,p2)) | |
| 743 | end | |
| 744 | 	  | x2p (Elem("peqmp",[],[prf1,prf2])) =
 | |
| 745 | let | |
| 746 | val p1 = x2p prf1 | |
| 747 | val p2 = x2p prf2 | |
| 748 | in | |
| 749 | mk_proof (PEqMp(p1,p2)) | |
| 750 | end | |
| 751 | 	  | x2p (Elem("peqimp",[],[prf])) =
 | |
| 752 | let | |
| 753 | val p = x2p prf | |
| 754 | in | |
| 755 | mk_proof (PEqImp p) | |
| 756 | end | |
| 757 | 	  | x2p (Elem("pexists",[("e",ise),("w",isw)],[prf])) =
 | |
| 758 | let | |
| 759 | val p = x2p prf | |
| 760 | val ex = index_to_term ise | |
| 761 | val w = index_to_term isw | |
| 762 | in | |
| 763 | mk_proof (PExists(p,ex,w)) | |
| 764 | end | |
| 765 | 	  | x2p (Elem("pchoose",[("i",is)],[prf1,prf2])) =
 | |
| 766 | let | |
| 767 | val v = index_to_term is | |
| 768 | val p1 = x2p prf1 | |
| 769 | val p2 = x2p prf2 | |
| 770 | in | |
| 771 | mk_proof (PChoose(v,p1,p2)) | |
| 772 | end | |
| 773 | 	  | x2p (Elem("pconj",[],[prf1,prf2])) =
 | |
| 774 | let | |
| 775 | val p1 = x2p prf1 | |
| 776 | val p2 = x2p prf2 | |
| 777 | in | |
| 778 | mk_proof (PConj(p1,p2)) | |
| 779 | end | |
| 780 | 	  | x2p (Elem("pconjunct1",[],[prf])) =
 | |
| 781 | let | |
| 782 | val p = x2p prf | |
| 783 | in | |
| 784 | mk_proof (PConjunct1 p) | |
| 785 | end | |
| 786 | 	  | x2p (Elem("pconjunct2",[],[prf])) =
 | |
| 787 | let | |
| 788 | val p = x2p prf | |
| 789 | in | |
| 790 | mk_proof (PConjunct2 p) | |
| 791 | end | |
| 792 | 	  | x2p (Elem("pdisj1",[("i",is)],[prf])) =
 | |
| 793 | let | |
| 794 | val p = x2p prf | |
| 795 | val t = index_to_term is | |
| 796 | in | |
| 797 | mk_proof (PDisj1 (p,t)) | |
| 798 | end | |
| 799 | 	  | x2p (Elem("pdisj2",[("i",is)],[prf])) =
 | |
| 800 | let | |
| 801 | val p = x2p prf | |
| 802 | val t = index_to_term is | |
| 803 | in | |
| 804 | mk_proof (PDisj2 (p,t)) | |
| 805 | end | |
| 806 | 	  | x2p (Elem("pdisjcases",[],[prf1,prf2,prf3])) =
 | |
| 807 | let | |
| 808 | val p1 = x2p prf1 | |
| 809 | val p2 = x2p prf2 | |
| 810 | val p3 = x2p prf3 | |
| 811 | in | |
| 812 | mk_proof (PDisjCases(p1,p2,p3)) | |
| 813 | end | |
| 814 | 	  | x2p (Elem("pnoti",[],[prf])) =
 | |
| 815 | let | |
| 816 | val p = x2p prf | |
| 817 | in | |
| 818 | mk_proof (PNotI p) | |
| 819 | end | |
| 820 | 	  | x2p (Elem("pnote",[],[prf])) =
 | |
| 821 | let | |
| 822 | val p = x2p prf | |
| 823 | in | |
| 824 | mk_proof (PNotE p) | |
| 825 | end | |
| 826 | 	  | x2p (Elem("pcontr",[("i",is)],[prf])) =
 | |
| 827 | let | |
| 828 | val p = x2p prf | |
| 829 | val t = index_to_term is | |
| 830 | in | |
| 831 | mk_proof (PContr (p,t)) | |
| 832 | end | |
| 833 | | x2p xml = raise ERR "x2p" "Bad proof" | |
| 834 | in | |
| 835 | x2p prf | |
| 836 | end | |
| 837 | ||
| 838 | fun import_proof thyname thmname thy = | |
| 839 | let | |
| 840 | val is = TextIO.openIn(proof_file_name thyname thmname thy) | |
| 841 | val (proof_xml,_) = scan_tag (LazySeq.of_instream is) | |
| 842 | val _ = TextIO.closeIn is | |
| 843 | in | |
| 844 | case proof_xml of | |
| 845 | 	    Elem("proof",[],xtypes::xterms::prf::rest) =>
 | |
| 846 | let | |
| 847 | val types = TypeNet.input_types thyname xtypes | |
| 848 | val terms = TermNet.input_terms thyname types xterms | |
| 849 | in | |
| 850 | (case rest of | |
| 15531 | 851 | [] => NONE | 
| 852 | | [xtm] => SOME (fn thy => TermNet.get_term_from_xml thy thyname types terms xtm) | |
| 14516 | 853 | | _ => raise ERR "import_proof" "Bad argument list", | 
| 854 | xml_to_proof thyname types terms prf) | |
| 855 | end | |
| 856 | | _ => raise ERR "import_proof" "Bad proof" | |
| 857 | end | |
| 858 | ||
| 859 | fun uniq_compose m th i st = | |
| 860 | let | |
| 861 | val res = bicompose false (false,th,m) i st | |
| 862 | in | |
| 863 | case Seq.pull res of | |
| 15531 | 864 | SOME (th,rest) => (case Seq.pull rest of | 
| 865 | SOME _ => raise ERR "uniq_compose" "Not unique!" | |
| 866 | | NONE => th) | |
| 867 | | NONE => raise ERR "uniq_compose" "No result" | |
| 14516 | 868 | end | 
| 869 | ||
| 870 | val reflexivity_thm = thm "refl" | |
| 871 | val substitution_thm = thm "subst" | |
| 872 | val mp_thm = thm "mp" | |
| 873 | val imp_antisym_thm = thm "light_imp_as" | |
| 874 | val disch_thm = thm "impI" | |
| 875 | val ccontr_thm = thm "ccontr" | |
| 876 | ||
| 877 | val meta_eq_to_obj_eq_thm = thm "meta_eq_to_obj_eq" | |
| 878 | ||
| 879 | val gen_thm = thm "HOLallI" | |
| 880 | val choose_thm = thm "exE" | |
| 881 | val exists_thm = thm "exI" | |
| 882 | val conj_thm = thm "conjI" | |
| 883 | val conjunct1_thm = thm "conjunct1" | |
| 884 | val conjunct2_thm = thm "conjunct2" | |
| 885 | val spec_thm = thm "spec" | |
| 886 | val disj_cases_thm = thm "disjE" | |
| 887 | val disj1_thm = thm "disjI1" | |
| 888 | val disj2_thm = thm "disjI2" | |
| 889 | ||
| 890 | local | |
| 891 | val th = thm "not_def" | |
| 892 | val sg = sign_of_thm th | |
| 893 |     val pp = reflexive (cterm_of sg (Const("Trueprop",boolT-->propT)))
 | |
| 894 | in | |
| 895 | val not_elim_thm = combination pp th | |
| 896 | end | |
| 897 | ||
| 898 | val not_intro_thm = symmetric not_elim_thm | |
| 899 | val abs_thm = thm "ext" | |
| 900 | val trans_thm = thm "trans" | |
| 901 | val symmetry_thm = thm "sym" | |
| 902 | val transitivity_thm = thm "trans" | |
| 903 | val eqmp_thm = thm "iffD1" | |
| 904 | val eqimp_thm = thm "HOL4Setup.eq_imp" | |
| 905 | val comb_thm = thm "cong" | |
| 906 | ||
| 907 | (* Beta-eta normalizes a theorem (only the conclusion, not the * | |
| 908 | hypotheses!) *) | |
| 909 | ||
| 910 | fun beta_eta_thm th = | |
| 911 | let | |
| 912 | val th1 = Thm.equal_elim (Thm.beta_conversion true (cprop_of th)) th | |
| 913 | val th2 = Thm.equal_elim (Thm.eta_conversion (cprop_of th1)) th1 | |
| 914 | in | |
| 915 | th2 | |
| 916 | end | |
| 917 | ||
| 918 | fun implies_elim_all th = | |
| 15570 | 919 | Library.foldl (fn (th,p) => implies_elim th (assume p)) (th,cprems_of th) | 
| 14516 | 920 | |
| 921 | fun norm_hyps th = | |
| 922 | th |> beta_eta_thm | |
| 923 | |> implies_elim_all | |
| 924 | |> implies_intr_hyps | |
| 925 | ||
| 926 | fun mk_GEN v th sg = | |
| 927 | let | |
| 928 | val c = HOLogic.dest_Trueprop (concl_of th) | |
| 929 | val cv = cterm_of sg v | |
| 930 | val lc = Term.lambda v c | |
| 931 | val clc = Thm.cterm_of sg lc | |
| 932 | val cvty = ctyp_of_term cv | |
| 933 | val th1 = implies_elim_all th | |
| 934 | val th2 = beta_eta_thm (forall_intr cv th1) | |
| 15531 | 935 | val th3 = th2 COMP (beta_eta_thm (Drule.instantiate' [SOME cvty] [SOME clc] gen_thm)) | 
| 14516 | 936 | val c = prop_of th3 | 
| 937 | val vname = fst(dest_Free v) | |
| 938 | val (cold,cnew) = case c of | |
| 939 | 			      tpc $ (Const("All",allT) $ Abs(oldname,ty,body)) =>
 | |
| 940 | (Abs(oldname,dummyT,Bound 0),Abs(vname,dummyT,Bound 0)) | |
| 941 | 			    | tpc $ (Const("All",allT) $ rest) => (tpc,tpc)
 | |
| 942 | | _ => raise ERR "mk_GEN" "Unknown conclusion" | |
| 943 | val th4 = Thm.rename_boundvars cold cnew th3 | |
| 944 | val res = implies_intr_hyps th4 | |
| 945 | in | |
| 946 | res | |
| 947 | end | |
| 948 | ||
| 949 | (* rotate left k places, leaving the first j and last l premises alone | |
| 950 | *) | |
| 951 | ||
| 952 | fun permute_prems j k 0 th = Thm.permute_prems j k th | |
| 953 | | permute_prems j k l th = | |
| 954 | th |> Thm.permute_prems 0 (~l) | |
| 955 | |> Thm.permute_prems (j+l) k | |
| 956 | |> Thm.permute_prems 0 l | |
| 957 | ||
| 958 | fun rearrange sg tm th = | |
| 959 | let | |
| 960 | val tm' = Pattern.beta_eta_contract tm | |
| 961 | fun find [] n = permute_prems 0 1 0 (implies_intr (Thm.cterm_of sg tm) th) | |
| 962 | | find (p::ps) n = if tm' aconv (Pattern.beta_eta_contract p) | |
| 963 | then permute_prems n 1 0 th | |
| 964 | else find ps (n+1) | |
| 965 | in | |
| 966 | find (prems_of th) 0 | |
| 967 | end | |
| 968 | ||
| 969 | fun zip (x::xs) (y::ys) = (x,y)::(zip xs ys) | |
| 970 | | zip [] [] = [] | |
| 971 | | zip _ _ = raise ERR "zip" "arguments not of same length" | |
| 972 | ||
| 973 | fun mk_INST dom rng th = | |
| 974 | th |> forall_intr_list dom | |
| 975 | |> forall_elim_list rng | |
| 976 | ||
| 977 | fun apply_tyinst_typ tyinst = | |
| 978 | let | |
| 979 | fun G (ty as TFree _) = | |
| 980 | (case try (Lib.assoc ty) tyinst of | |
| 15531 | 981 | SOME ty' => ty' | 
| 982 | | NONE => ty) | |
| 14516 | 983 | | G (Type(tyname,tys)) = Type(tyname,map G tys) | 
| 984 | | G (TVar _) = raise ERR "apply_tyinst_typ" "Scematic variable found" | |
| 985 | in | |
| 986 | G | |
| 987 | end | |
| 988 | ||
| 989 | fun apply_tyinst_term tyinst = | |
| 990 | let | |
| 991 | val G = apply_tyinst_typ tyinst | |
| 992 | fun F (tm as Bound _) = tm | |
| 993 | | F (tm as Free(vname,ty)) = Free(vname,G ty) | |
| 994 | | F (tm as Const(vname,ty)) = Const(vname,G ty) | |
| 995 | | F (tm1 $ tm2) = (F tm1) $ (F tm2) | |
| 996 | | F (Abs(vname,ty,body)) = Abs(vname,G ty,F body) | |
| 997 | | F (Var _) = raise ERR "apply_tyinst_term" "Schematic variable found" | |
| 998 | in | |
| 999 | F | |
| 1000 | end | |
| 1001 | ||
| 1002 | fun apply_inst_term tminst = | |
| 1003 | let | |
| 1004 | fun F (tm as Bound _) = tm | |
| 1005 | | F (tm as Free _) = | |
| 1006 | (case try (Lib.assoc tm) tminst of | |
| 15531 | 1007 | SOME tm' => tm' | 
| 1008 | | NONE => tm) | |
| 14516 | 1009 | | F (tm as Const _) = tm | 
| 1010 | | F (tm1 $ tm2) = (F tm1) $ (F tm2) | |
| 1011 | | F (Abs(vname,ty,body)) = Abs(vname,ty,F body) | |
| 1012 | | F (Var _) = raise ERR "apply_inst_term" "Schematic variable found" | |
| 1013 | in | |
| 1014 | F | |
| 1015 | end | |
| 1016 | ||
| 1017 | val collect_vars = | |
| 1018 | let | |
| 1019 | fun F vars (Bound _) = vars | |
| 1020 | | F vars (tm as Free _) = | |
| 1021 | if tm mem vars | |
| 1022 | then vars | |
| 1023 | else (tm::vars) | |
| 1024 | | F vars (Const _) = vars | |
| 1025 | | F vars (tm1 $ tm2) = F (F vars tm1) tm2 | |
| 1026 | | F vars (Abs(_,_,body)) = F vars body | |
| 1027 | | F vars (Var _) = raise ERR "collect_vars" "Schematic variable found" | |
| 1028 | in | |
| 1029 | F [] | |
| 1030 | end | |
| 1031 | ||
| 1032 | (* Code for disambiguating variablenames (wrt. types) *) | |
| 1033 | ||
| 1034 | val disamb_info_empty = {vars=[],rens=[]}
 | |
| 1035 | ||
| 1036 | fun rens_of {vars,rens} = rens
 | |
| 1037 | ||
| 1038 | fun name_of_var (Free(vname,_)) = vname | |
| 1039 | | name_of_var _ = raise ERR "name_of_var" "Not a variable" | |
| 1040 | ||
| 1041 | fun disamb_helper {vars,rens} tm =
 | |
| 1042 | let | |
| 1043 | val varstm = collect_vars tm | |
| 1044 | fun process (v as Free(vname,ty),(vars,rens,inst)) = | |
| 1045 | if v mem vars | |
| 1046 | then (vars,rens,inst) | |
| 1047 | else (case try (Lib.assoc v) rens of | |
| 15531 | 1048 | SOME vnew => (vars,rens,(v,vnew)::inst) | 
| 1049 | | NONE => if exists (fn Free(vname',_) => vname = vname' | _ => raise ERR "disamb_helper" "Bad varlist") vars | |
| 14516 | 1050 | then | 
| 1051 | let | |
| 1052 | val tmnames = map name_of_var varstm | |
| 1053 | val varnames = map name_of_var vars | |
| 1054 | val (dom,rng) = ListPair.unzip rens | |
| 1055 | val rensnames = (map name_of_var dom) @ (map name_of_var rng) | |
| 1056 | val instnames = map name_of_var (snd (ListPair.unzip inst)) | |
| 1057 | val allnames = tmnames @ varnames @ rensnames @ instnames | |
| 1058 | val vnewname = Term.variant allnames vname | |
| 1059 | val vnew = Free(vnewname,ty) | |
| 1060 | in | |
| 1061 | (vars,(v,vnew)::rens,(v,vnew)::inst) | |
| 1062 | end | |
| 1063 | else (v::vars,rens,inst)) | |
| 1064 | | process _ = raise ERR "disamb_helper" "Internal error" | |
| 1065 | ||
| 1066 | val (vars',rens',inst) = | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1067 | foldr process (vars,rens,[]) varstm | 
| 14516 | 1068 | in | 
| 1069 | 	({vars=vars',rens=rens'},inst)
 | |
| 1070 | end | |
| 1071 | ||
| 1072 | fun disamb_term_from info tm = | |
| 1073 | let | |
| 1074 | val (info',inst) = disamb_helper info tm | |
| 1075 | in | |
| 1076 | (info',apply_inst_term inst tm) | |
| 1077 | end | |
| 1078 | ||
| 1079 | fun swap (x,y) = (y,x) | |
| 1080 | ||
| 1081 | fun has_ren (HOLThm([],_)) = false | |
| 1082 | | has_ren _ = true | |
| 1083 | ||
| 1084 | fun prinfo {vars,rens} = (writeln "Vars:";
 | |
| 1085 | app prin vars; | |
| 1086 | writeln "Renaming:"; | |
| 1087 | app (fn(x,y)=>(prin x; writeln " -->"; prin y)) rens) | |
| 1088 | ||
| 1089 | fun disamb_thm_from info (hth as HOLThm(rens,thm)) = | |
| 1090 | let | |
| 1091 | val inv_rens = map swap rens | |
| 1092 | val orig_thm = apply_inst_term inv_rens (prop_of thm) | |
| 1093 | val (info',inst) = disamb_helper info orig_thm | |
| 1094 | val rens' = map (apsnd (apply_inst_term inst)) inv_rens | |
| 1095 | val (dom,rng) = ListPair.unzip (rens' @ inst) | |
| 1096 | val sg = sign_of_thm thm | |
| 1097 | val thm' = mk_INST (map (cterm_of sg) dom) (map (cterm_of sg) rng) thm | |
| 1098 | in | |
| 1099 | (info',thm') | |
| 1100 | end | |
| 1101 | ||
| 1102 | fun disamb_terms_from info tms = | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1103 | foldr (fn (tm,(info,tms)) => | 
| 14516 | 1104 | let | 
| 1105 | val (info',tm') = disamb_term_from info tm | |
| 1106 | in | |
| 1107 | (info',tm'::tms) | |
| 1108 | end) | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1109 | (info,[]) tms | 
| 14516 | 1110 | |
| 1111 | fun disamb_thms_from info hthms = | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1112 | foldr (fn (hthm,(info,thms)) => | 
| 14516 | 1113 | let | 
| 1114 | val (info',tm') = disamb_thm_from info hthm | |
| 1115 | in | |
| 1116 | (info',tm'::thms) | |
| 1117 | end) | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1118 | (info,[]) hthms | 
| 14516 | 1119 | |
| 1120 | fun disamb_term tm = disamb_term_from disamb_info_empty tm | |
| 1121 | fun disamb_terms tms = disamb_terms_from disamb_info_empty tms | |
| 1122 | fun disamb_thm thm = disamb_thm_from disamb_info_empty thm | |
| 1123 | fun disamb_thms thms = disamb_thms_from disamb_info_empty thms | |
| 1124 | ||
| 1125 | fun norm_hthm sg (hth as HOLThm([],_)) = hth | |
| 1126 | | norm_hthm sg (hth as HOLThm(rens,th)) = | |
| 1127 | let | |
| 1128 | val vars = collect_vars (prop_of th) | |
| 1129 | val (rens',inst,_) = | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1130 | foldr (fn((ren as (vold as Free(vname,_),vnew)), | 
| 14516 | 1131 | (rens,inst,vars)) => | 
| 1132 | (case Library.find_first (fn Free(v,_) => v = vname | _ => false) vars of | |
| 15531 | 1133 | SOME v' => if v' = vold | 
| 14516 | 1134 | then (rens,(vnew,vold)::inst,vold::vars) | 
| 1135 | else (ren::rens,(vold,vnew)::inst,vnew::vars) | |
| 15531 | 1136 | | NONE => (rens,(vnew,vold)::inst,vold::vars)) | 
| 14516 | 1137 | | _ => raise ERR "norm_hthm" "Internal error") | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1138 | ([],[],vars) rens | 
| 14516 | 1139 | val (dom,rng) = ListPair.unzip inst | 
| 1140 | val th' = mk_INST (map (cterm_of sg) dom) (map (cterm_of sg) rng) th | |
| 1141 | val nvars = collect_vars (prop_of th') | |
| 15570 | 1142 | val rens' = List.filter (fn(_,v) => v mem nvars) rens | 
| 14516 | 1143 | val res = HOLThm(rens',th') | 
| 1144 | in | |
| 1145 | res | |
| 1146 | end | |
| 1147 | ||
| 1148 | (* End of disambiguating code *) | |
| 1149 | ||
| 1150 | val debug = ref false | |
| 1151 | ||
| 1152 | fun if_debug f x = if !debug then f x else () | |
| 1153 | val message = if_debug writeln | |
| 1154 | ||
| 1155 | val conjE_helper = Thm.permute_prems 0 1 conjE | |
| 1156 | ||
| 1157 | fun get_hol4_thm thyname thmname thy = | |
| 1158 | case get_hol4_theorem thyname thmname thy of | |
| 15531 | 1159 | SOME hth => SOME (HOLThm hth) | 
| 1160 | | NONE => | |
| 14516 | 1161 | let | 
| 1162 | val pending = HOL4Pending.get thy | |
| 1163 | in | |
| 1164 | case StringPair.lookup (pending,(thyname,thmname)) of | |
| 15531 | 1165 | SOME hth => SOME (HOLThm hth) | 
| 1166 | | NONE => NONE | |
| 14516 | 1167 | end | 
| 1168 | ||
| 1169 | fun non_trivial_term_consts tm = | |
| 15570 | 1170 | List.filter (fn c => not (c = "Trueprop" orelse | 
| 14516 | 1171 | c = "All" orelse | 
| 1172 | c = "op -->" orelse | |
| 1173 | c = "op &" orelse | |
| 1174 | c = "op =")) (Term.term_consts tm) | |
| 1175 | ||
| 1176 | fun match_consts t (* th *) = | |
| 1177 | let | |
| 1178 | fun add_consts (Const (c, _), cs) = | |
| 1179 | (case c of | |
| 1180 | "op =" => "==" ins_string cs | |
| 1181 | | "op -->" => "==>" ins_string cs | |
| 1182 | | "All" => cs | |
| 1183 | | "all" => cs | |
| 1184 | | "op &" => cs | |
| 1185 | | "Trueprop" => cs | |
| 1186 | | _ => c ins_string cs) | |
| 1187 | | add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs)) | |
| 1188 | | add_consts (Abs (_, _, t), cs) = add_consts (t, cs) | |
| 1189 | | add_consts (_, cs) = cs | |
| 1190 | val t_consts = add_consts(t,[]) | |
| 1191 | in | |
| 1192 | fn th => eq_set(t_consts,add_consts(prop_of th,[])) | |
| 1193 | end | |
| 1194 | ||
| 1195 | fun split_name str = | |
| 1196 | let | |
| 1197 | val sub = Substring.all str | |
| 1198 | val (f,idx) = apsnd Substring.string (Substring.splitr Char.isDigit sub) | |
| 1199 | val (newstr,u) = apboth Substring.string (Substring.splitr (fn c => c = #"_") f) | |
| 1200 | in | |
| 1201 | if not (idx = "") andalso u = "_" | |
| 15531 | 1202 | then SOME (newstr,valOf(Int.fromString idx)) | 
| 1203 | else NONE | |
| 14516 | 1204 | end | 
| 15531 | 1205 | handle _ => NONE | 
| 14516 | 1206 | |
| 1207 | fun rewrite_hol4_term t thy = | |
| 1208 | let | |
| 1209 | val sg = sign_of thy | |
| 1210 | ||
| 1211 | val hol4rews1 = map (transfer_sg sg) (HOL4Rewrites.get thy) | |
| 1212 | val hol4ss = empty_ss setmksimps single addsimps hol4rews1 | |
| 1213 | in | |
| 1214 | transfer_sg sg (Simplifier.full_rewrite hol4ss (cterm_of sg t)) | |
| 1215 | end | |
| 1216 | ||
| 1217 | ||
| 1218 | fun get_isabelle_thm thyname thmname hol4conc thy = | |
| 1219 | let | |
| 15647 | 1220 | val _ = message "get_isabelle_thm called..." | 
| 14516 | 1221 | val sg = sign_of thy | 
| 1222 | ||
| 1223 | val (info,hol4conc') = disamb_term hol4conc | |
| 1224 | val i2h_conc = symmetric (rewrite_hol4_term (HOLogic.mk_Trueprop hol4conc') thy) | |
| 1225 | val isaconc = | |
| 1226 | case concl_of i2h_conc of | |
| 1227 | 		Const("==",_) $ lhs $ _ => lhs
 | |
| 1228 | | _ => error "get_isabelle_thm" "Bad rewrite rule" | |
| 1229 | val _ = (message "Original conclusion:"; | |
| 1230 | if_debug prin hol4conc'; | |
| 1231 | message "Modified conclusion:"; | |
| 1232 | if_debug prin isaconc) | |
| 1233 | ||
| 1234 | fun mk_res th = HOLThm(rens_of info,equal_elim i2h_conc th) | |
| 1235 | in | |
| 1236 | case get_hol4_mapping thyname thmname thy of | |
| 15531 | 1237 | SOME (SOME thmname) => | 
| 14516 | 1238 | let | 
| 1239 | 		val _ = message ("Looking for " ^ thmname)
 | |
| 15531 | 1240 | val th1 = (SOME (transform_error (PureThy.get_thm thy) (thmname, NONE)) | 
| 14516 | 1241 | handle ERROR_MESSAGE _ => | 
| 1242 | (case split_name thmname of | |
| 15570 | 1243 | SOME (listname,idx) => (SOME (List.nth(PureThy.get_thms thy (listname, NONE),idx-1)) | 
| 15531 | 1244 | handle _ => NONE) | 
| 1245 | | NONE => NONE)) | |
| 14516 | 1246 | in | 
| 1247 | case th1 of | |
| 15531 | 1248 | SOME th2 => | 
| 14516 | 1249 | (case Shuffler.set_prop thy isaconc [(thmname,th2)] of | 
| 15531 | 1250 | SOME (_,th) => (message "YES";(thy, SOME (mk_res th))) | 
| 1251 | | NONE => (message "NO2";error "get_isabelle_thm" "Bad mapping")) | |
| 1252 | | NONE => (message "NO1";error "get_isabelle_thm" "Bad mapping") | |
| 14516 | 1253 | end | 
| 15531 | 1254 | 	  | SOME NONE => error ("Trying to access ignored theorem " ^ thmname)
 | 
| 1255 | | NONE => | |
| 14516 | 1256 | let | 
| 1257 | val _ = (message "Looking for conclusion:"; | |
| 1258 | if_debug prin isaconc) | |
| 1259 | val cs = non_trivial_term_consts isaconc | |
| 1260 | val _ = (message "Looking for consts:"; | |
| 15647 | 1261 | message (commas cs)) | 
| 14516 | 1262 | val pot_thms = Shuffler.find_potential thy isaconc | 
| 1263 | val _ = message ((Int.toString (length pot_thms)) ^ " potential theorems") | |
| 1264 | in | |
| 1265 | case Shuffler.set_prop thy isaconc pot_thms of | |
| 15531 | 1266 | SOME (isaname,th) => | 
| 14516 | 1267 | let | 
| 1268 | val hth as HOLThm args = mk_res th | |
| 1269 | val thy' = thy |> add_hol4_theorem thyname thmname args | |
| 1270 | |> add_hol4_mapping thyname thmname isaname | |
| 1271 | in | |
| 15531 | 1272 | (thy',SOME hth) | 
| 14516 | 1273 | end | 
| 15531 | 1274 | | NONE => (thy,NONE) | 
| 14516 | 1275 | end | 
| 1276 | end | |
| 15647 | 1277 | handle e => (message "Exception in get_isabelle_thm"; if_debug print_exn e handle _ => (); (thy,NONE)) | 
| 14516 | 1278 | |
| 1279 | fun get_thm thyname thmname thy = | |
| 1280 | case get_hol4_thm thyname thmname thy of | |
| 15531 | 1281 | SOME hth => (thy,SOME hth) | 
| 1282 | | NONE => ((case fst (import_proof thyname thmname thy) of | |
| 1283 | SOME f => get_isabelle_thm thyname thmname (f thy) thy | |
| 15647 | 1284 | | NONE => (message "No conclusion"; (thy,NONE))) | 
| 1285 | handle e as IO.Io _ => (message "IO exception"; (thy,NONE)) | |
| 1286 | | e as PK _ => (message "PK exception"; (thy,NONE))) | |
| 14516 | 1287 | |
| 1288 | fun rename_const thyname thy name = | |
| 1289 | case get_hol4_const_renaming thyname name thy of | |
| 15531 | 1290 | SOME cname => cname | 
| 1291 | | NONE => name | |
| 14516 | 1292 | |
| 1293 | fun get_def thyname constname rhs thy = | |
| 1294 | let | |
| 1295 | val constname = rename_const thyname thy constname | |
| 1296 | val (thmname,thy') = get_defname thyname constname thy | |
| 1297 | 	val _ = message ("Looking for definition " ^ thyname ^ "." ^ thmname)
 | |
| 1298 | in | |
| 1299 | get_isabelle_thm thyname thmname (mk_teq (thyname ^ "." ^ constname) rhs thy') thy' | |
| 1300 | end | |
| 1301 | ||
| 1302 | fun get_axiom thyname axname thy = | |
| 1303 | case get_thm thyname axname thy of | |
| 15531 | 1304 | arg as (_,SOME _) => arg | 
| 14516 | 1305 |       | _ => raise ERR "get_axiom" ("Trying to retrieve axiom (" ^ axname ^ ")")
 | 
| 1306 | ||
| 1307 | fun intern_store_thm gen_output thyname thmname hth thy = | |
| 1308 | let | |
| 1309 | val sg = sign_of thy | |
| 1310 | val (hth' as HOLThm (args as (_,th))) = norm_hthm sg hth | |
| 1311 | 	val _ = if has_ren hth' then warning ("Theorem " ^ thmname ^ " needs variable-disambiguating")
 | |
| 1312 | else () | |
| 1313 | val rew = rewrite_hol4_term (concl_of th) thy | |
| 1314 | val th = equal_elim rew th | |
| 1315 | val thy' = add_hol4_pending thyname thmname args thy | |
| 1316 | val thy2 = if gen_output | |
| 1317 | 		   then add_dump ("lemma " ^ thmname ^ ": " ^ (smart_string_of_thm th) ^ "\n  by (import " ^ thyname ^ " " ^ thmname ^ ")") thy'
 | |
| 1318 | else thy' | |
| 1319 | in | |
| 1320 | (thy2,hth') | |
| 1321 | end | |
| 1322 | ||
| 1323 | val store_thm = intern_store_thm true | |
| 1324 | ||
| 1325 | fun mk_REFL ctm = | |
| 1326 | let | |
| 1327 | val cty = Thm.ctyp_of_term ctm | |
| 1328 | in | |
| 15531 | 1329 | Drule.instantiate' [SOME cty] [SOME ctm] reflexivity_thm | 
| 14516 | 1330 | end | 
| 1331 | ||
| 1332 | fun REFL tm thy = | |
| 1333 | let | |
| 1334 | val _ = message "REFL:" | |
| 1335 | val (info,tm') = disamb_term tm | |
| 1336 | val sg = sign_of thy | |
| 1337 | val ctm = Thm.cterm_of sg tm' | |
| 1338 | val res = HOLThm(rens_of info,mk_REFL ctm) | |
| 1339 | val _ = if_debug pth res | |
| 1340 | in | |
| 1341 | (thy,res) | |
| 1342 | end | |
| 1343 | ||
| 1344 | fun ASSUME tm thy = | |
| 1345 | let | |
| 1346 | val _ = message "ASSUME:" | |
| 1347 | val (info,tm') = disamb_term tm | |
| 1348 | val sg = sign_of thy | |
| 1349 | val ctm = Thm.cterm_of sg (HOLogic.mk_Trueprop tm') | |
| 1350 | val th = Thm.trivial ctm | |
| 1351 | val res = HOLThm(rens_of info,th) | |
| 1352 | val _ = if_debug pth res | |
| 1353 | in | |
| 1354 | (thy,res) | |
| 1355 | end | |
| 1356 | ||
| 1357 | fun INST_TYPE lambda (hth as HOLThm(rens,th)) thy = | |
| 1358 | let | |
| 1359 | val _ = message "INST_TYPE:" | |
| 1360 | val _ = if_debug pth hth | |
| 1361 | val sg = sign_of thy | |
| 1362 | val tys_before = add_term_tfrees (prop_of th,[]) | |
| 1363 | val th1 = varifyT th | |
| 1364 | val tys_after = add_term_tvars (prop_of th1,[]) | |
| 15794 
5de27a5fc5ed
Adapted to new interface of instantiation and unification / matching functions.
 berghofe parents: 
15647diff
changeset | 1365 | val tyinst = map (fn (bef, iS) => | 
| 14516 | 1366 | (case try (Lib.assoc (TFree bef)) lambda of | 
| 15794 
5de27a5fc5ed
Adapted to new interface of instantiation and unification / matching functions.
 berghofe parents: 
15647diff
changeset | 1367 | SOME ty => (ctyp_of sg (TVar iS), ctyp_of sg ty) | 
| 
5de27a5fc5ed
Adapted to new interface of instantiation and unification / matching functions.
 berghofe parents: 
15647diff
changeset | 1368 | | NONE => (ctyp_of sg (TVar iS), ctyp_of sg (TFree bef)) | 
| 14516 | 1369 | )) | 
| 1370 | (zip tys_before tys_after) | |
| 1371 | val res = Drule.instantiate (tyinst,[]) th1 | |
| 1372 | val appty = apboth (apply_tyinst_term lambda) | |
| 1373 | val hth = HOLThm(map appty rens,res) | |
| 1374 | val res = norm_hthm sg hth | |
| 1375 | val _ = message "RESULT:" | |
| 1376 | val _ = if_debug pth res | |
| 1377 | in | |
| 1378 | (thy,res) | |
| 1379 | end | |
| 1380 | ||
| 1381 | fun INST sigma hth thy = | |
| 1382 | let | |
| 1383 | val _ = message "INST:" | |
| 1384 | val _ = if_debug (app (fn (x,y) => (prin x; prin y))) sigma | |
| 1385 | val _ = if_debug pth hth | |
| 1386 | val sg = sign_of thy | |
| 1387 | val (sdom,srng) = ListPair.unzip sigma | |
| 1388 | val (info,th) = disamb_thm hth | |
| 1389 | val (info',srng') = disamb_terms_from info srng | |
| 1390 | val rens = rens_of info' | |
| 1391 | val sdom' = map (apply_inst_term rens) sdom | |
| 1392 | val th1 = mk_INST (map (cterm_of sg) sdom') (map (cterm_of sg) srng') th | |
| 1393 | val res = HOLThm(rens,th1) | |
| 1394 | val _ = message "RESULT:" | |
| 1395 | val _ = if_debug pth res | |
| 1396 | in | |
| 1397 | (thy,res) | |
| 1398 | end | |
| 1399 | ||
| 1400 | fun EQ_IMP_RULE (hth as HOLThm(rens,th)) thy = | |
| 1401 | let | |
| 1402 | val _ = message "EQ_IMP_RULE:" | |
| 1403 | val _ = if_debug pth hth | |
| 1404 | val res = HOLThm(rens,th RS eqimp_thm) | |
| 1405 | val _ = message "RESULT:" | |
| 1406 | val _ = if_debug pth res | |
| 1407 | in | |
| 1408 | (thy,res) | |
| 1409 | end | |
| 1410 | ||
| 1411 | fun mk_EQ_MP th1 th2 = [beta_eta_thm th1,beta_eta_thm th2] MRS eqmp_thm | |
| 1412 | ||
| 1413 | fun EQ_MP hth1 hth2 thy = | |
| 1414 | let | |
| 1415 | val _ = message "EQ_MP:" | |
| 1416 | val _ = if_debug pth hth1 | |
| 1417 | val _ = if_debug pth hth2 | |
| 1418 | val (info,[th1,th2]) = disamb_thms [hth1,hth2] | |
| 1419 | val res = HOLThm(rens_of info,mk_EQ_MP th1 th2) | |
| 1420 | val _ = message "RESULT:" | |
| 1421 | val _ = if_debug pth res | |
| 1422 | in | |
| 1423 | (thy,res) | |
| 1424 | end | |
| 1425 | ||
| 1426 | fun mk_COMB th1 th2 sg = | |
| 1427 | let | |
| 1428 | val (f,g) = case concl_of th1 of | |
| 1429 | 			_ $ (Const("op =",_) $ f $ g) => (f,g)
 | |
| 1430 | | _ => raise ERR "mk_COMB" "First theorem not an equality" | |
| 1431 | val (x,y) = case concl_of th2 of | |
| 1432 | 			_ $ (Const("op =",_) $ x $ y) => (x,y)
 | |
| 1433 | | _ => raise ERR "mk_COMB" "Second theorem not an equality" | |
| 1434 | val fty = type_of f | |
| 1435 | val (fd,fr) = dom_rng fty | |
| 1436 | val comb_thm' = Drule.instantiate' | |
| 15531 | 1437 | [SOME (ctyp_of sg fd),SOME (ctyp_of sg fr)] | 
| 1438 | [SOME (cterm_of sg f),SOME (cterm_of sg g), | |
| 1439 | SOME (cterm_of sg x),SOME (cterm_of sg y)] comb_thm | |
| 14516 | 1440 | in | 
| 1441 | [th1,th2] MRS comb_thm' | |
| 1442 | end | |
| 1443 | ||
| 1444 | fun SUBST rews ctxt hth thy = | |
| 1445 | let | |
| 1446 | val _ = message "SUBST:" | |
| 1447 | val _ = if_debug (app pth) rews | |
| 1448 | val _ = if_debug prin ctxt | |
| 1449 | val _ = if_debug pth hth | |
| 1450 | val (info,th) = disamb_thm hth | |
| 1451 | val (info1,ctxt') = disamb_term_from info ctxt | |
| 1452 | val (info2,rews') = disamb_thms_from info1 rews | |
| 1453 | ||
| 1454 | val sg = sign_of thy | |
| 1455 | val cctxt = cterm_of sg ctxt' | |
| 1456 | fun subst th [] = th | |
| 1457 | | subst th (rew::rews) = subst (mk_COMB th rew sg) rews | |
| 1458 | val res = HOLThm(rens_of info2,mk_EQ_MP (subst (mk_REFL cctxt) rews') th) | |
| 1459 | val _ = message "RESULT:" | |
| 1460 | val _ = if_debug pth res | |
| 1461 | in | |
| 1462 | (thy,res) | |
| 1463 | end | |
| 1464 | ||
| 1465 | fun DISJ_CASES hth hth1 hth2 thy = | |
| 1466 | let | |
| 1467 | val _ = message "DISJ_CASES:" | |
| 1468 | val _ = if_debug (app pth) [hth,hth1,hth2] | |
| 1469 | val (info,th) = disamb_thm hth | |
| 1470 | val (info1,th1) = disamb_thm_from info hth1 | |
| 1471 | val (info2,th2) = disamb_thm_from info1 hth2 | |
| 1472 | val sg = sign_of thy | |
| 1473 | val th1 = norm_hyps th1 | |
| 1474 | val th2 = norm_hyps th2 | |
| 1475 | val (l,r) = case concl_of th of | |
| 1476 | 			_ $ (Const("op |",_) $ l $ r) => (l,r)
 | |
| 1477 | | _ => raise ERR "DISJ_CASES" "Conclusion not a disjunction" | |
| 1478 | val th1' = rearrange sg (HOLogic.mk_Trueprop l) th1 | |
| 1479 | val th2' = rearrange sg (HOLogic.mk_Trueprop r) th2 | |
| 1480 | val res1 = th RS disj_cases_thm | |
| 1481 | val res2 = uniq_compose ((nprems_of th1')-1) th1' ((nprems_of th)+1) res1 | |
| 1482 | val res3 = uniq_compose ((nprems_of th2')-1) th2' (nprems_of res2) res2 | |
| 1483 | val res = HOLThm(rens_of info2,res3) | |
| 1484 | val _ = message "RESULT:" | |
| 1485 | val _ = if_debug pth res | |
| 1486 | in | |
| 1487 | (thy,res) | |
| 1488 | end | |
| 1489 | ||
| 1490 | fun DISJ1 hth tm thy = | |
| 1491 | let | |
| 1492 | val _ = message "DISJ1:" | |
| 1493 | val _ = if_debug pth hth | |
| 1494 | val _ = if_debug prin tm | |
| 1495 | val (info,th) = disamb_thm hth | |
| 1496 | val (info',tm') = disamb_term_from info tm | |
| 1497 | val sg = sign_of thy | |
| 1498 | val ct = Thm.cterm_of sg tm' | |
| 15531 | 1499 | val disj1_thm' = Drule.instantiate' [] [NONE,SOME ct] disj1_thm | 
| 14516 | 1500 | val res = HOLThm(rens_of info',th RS disj1_thm') | 
| 1501 | val _ = message "RESULT:" | |
| 1502 | val _ = if_debug pth res | |
| 1503 | in | |
| 1504 | (thy,res) | |
| 1505 | end | |
| 1506 | ||
| 1507 | fun DISJ2 tm hth thy = | |
| 1508 | let | |
| 1509 | val _ = message "DISJ1:" | |
| 1510 | val _ = if_debug prin tm | |
| 1511 | val _ = if_debug pth hth | |
| 1512 | val (info,th) = disamb_thm hth | |
| 1513 | val (info',tm') = disamb_term_from info tm | |
| 1514 | val sg = sign_of thy | |
| 1515 | val ct = Thm.cterm_of sg tm' | |
| 15531 | 1516 | val disj2_thm' = Drule.instantiate' [] [NONE,SOME ct] disj2_thm | 
| 14516 | 1517 | val res = HOLThm(rens_of info',th RS disj2_thm') | 
| 1518 | val _ = message "RESULT:" | |
| 1519 | val _ = if_debug pth res | |
| 1520 | in | |
| 1521 | (thy,res) | |
| 1522 | end | |
| 1523 | ||
| 1524 | fun IMP_ANTISYM hth1 hth2 thy = | |
| 1525 | let | |
| 1526 | val _ = message "IMP_ANTISYM:" | |
| 1527 | val _ = if_debug pth hth1 | |
| 1528 | val _ = if_debug pth hth2 | |
| 1529 | val (info,[th1,th2]) = disamb_thms [hth1,hth2] | |
| 1530 | val th = [beta_eta_thm th1,beta_eta_thm th2] MRS imp_antisym_thm | |
| 1531 | val res = HOLThm(rens_of info,th) | |
| 1532 | val _ = message "RESULT:" | |
| 1533 | val _ = if_debug pth res | |
| 1534 | in | |
| 1535 | (thy,res) | |
| 1536 | end | |
| 1537 | ||
| 1538 | fun SYM (hth as HOLThm(rens,th)) thy = | |
| 1539 | let | |
| 1540 | val _ = message "SYM:" | |
| 1541 | val _ = if_debug pth hth | |
| 1542 | val th = th RS symmetry_thm | |
| 1543 | val res = HOLThm(rens,th) | |
| 1544 | val _ = message "RESULT:" | |
| 1545 | val _ = if_debug pth res | |
| 1546 | in | |
| 1547 | (thy,res) | |
| 1548 | end | |
| 1549 | ||
| 1550 | fun MP hth1 hth2 thy = | |
| 1551 | let | |
| 1552 | val _ = message "MP:" | |
| 1553 | val _ = if_debug pth hth1 | |
| 1554 | val _ = if_debug pth hth2 | |
| 1555 | val (info,[th1,th2]) = disamb_thms [hth1,hth2] | |
| 1556 | val th = [beta_eta_thm th1,beta_eta_thm th2] MRS mp_thm | |
| 1557 | val res = HOLThm(rens_of info,th) | |
| 1558 | val _ = message "RESULT:" | |
| 1559 | val _ = if_debug pth res | |
| 1560 | in | |
| 1561 | (thy,res) | |
| 1562 | end | |
| 1563 | ||
| 1564 | fun CONJ hth1 hth2 thy = | |
| 1565 | let | |
| 1566 | val _ = message "CONJ:" | |
| 1567 | val _ = if_debug pth hth1 | |
| 1568 | val _ = if_debug pth hth2 | |
| 1569 | val (info,[th1,th2]) = disamb_thms [hth1,hth2] | |
| 1570 | val th = [th1,th2] MRS conj_thm | |
| 1571 | val res = HOLThm(rens_of info,th) | |
| 1572 | val _ = message "RESULT:" | |
| 1573 | val _ = if_debug pth res | |
| 1574 | in | |
| 1575 | (thy,res) | |
| 1576 | end | |
| 1577 | ||
| 1578 | fun CONJUNCT1 (hth as HOLThm(rens,th)) thy = | |
| 1579 | let | |
| 1580 | val _ = message "CONJUNCT1:" | |
| 1581 | val _ = if_debug pth hth | |
| 1582 | val res = HOLThm(rens,th RS conjunct1_thm) | |
| 1583 | val _ = message "RESULT:" | |
| 1584 | val _ = if_debug pth res | |
| 1585 | in | |
| 1586 | (thy,res) | |
| 1587 | end | |
| 1588 | ||
| 1589 | fun CONJUNCT2 (hth as HOLThm(rens,th)) thy = | |
| 1590 | let | |
| 1591 | val _ = message "CONJUNCT1:" | |
| 1592 | val _ = if_debug pth hth | |
| 1593 | val res = HOLThm(rens,th RS conjunct2_thm) | |
| 1594 | val _ = message "RESULT:" | |
| 1595 | val _ = if_debug pth res | |
| 1596 | in | |
| 1597 | (thy,res) | |
| 1598 | end | |
| 1599 | ||
| 1600 | fun EXISTS ex wit hth thy = | |
| 1601 | let | |
| 1602 | val _ = message "EXISTS:" | |
| 1603 | val _ = if_debug prin ex | |
| 1604 | val _ = if_debug prin wit | |
| 1605 | val _ = if_debug pth hth | |
| 1606 | val (info,th) = disamb_thm hth | |
| 1607 | val (info',[ex',wit']) = disamb_terms_from info [ex,wit] | |
| 1608 | val sg = sign_of thy | |
| 1609 | val cwit = cterm_of sg wit' | |
| 1610 | val cty = ctyp_of_term cwit | |
| 1611 | val a = case ex' of | |
| 1612 | 		    (Const("Ex",_) $ a) => a
 | |
| 1613 | | _ => raise ERR "EXISTS" "Argument not existential" | |
| 1614 | val ca = cterm_of sg a | |
| 15531 | 1615 | val exists_thm' = beta_eta_thm (Drule.instantiate' [SOME cty] [SOME ca,SOME cwit] exists_thm) | 
| 14516 | 1616 | val th1 = beta_eta_thm th | 
| 1617 | val th2 = implies_elim_all th1 | |
| 1618 | val th3 = th2 COMP exists_thm' | |
| 1619 | val th = implies_intr_hyps th3 | |
| 1620 | val res = HOLThm(rens_of info',th) | |
| 1621 | val _ = message "RESULT:" | |
| 1622 | val _ = if_debug pth res | |
| 1623 | in | |
| 1624 | (thy,res) | |
| 1625 | end | |
| 1626 | ||
| 1627 | fun CHOOSE v hth1 hth2 thy = | |
| 1628 | let | |
| 1629 | val _ = message "CHOOSE:" | |
| 1630 | val _ = if_debug prin v | |
| 1631 | val _ = if_debug pth hth1 | |
| 1632 | val _ = if_debug pth hth2 | |
| 1633 | val (info,[th1,th2]) = disamb_thms [hth1,hth2] | |
| 1634 | val (info',v') = disamb_term_from info v | |
| 1635 | fun strip 0 _ th = th | |
| 1636 | | strip n (p::ps) th = | |
| 1637 | strip (n-1) ps (implies_elim th (assume p)) | |
| 1638 | | strip _ _ _ = raise ERR "CHOOSE" "strip error" | |
| 1639 | val sg = sign_of thy | |
| 1640 | val cv = cterm_of sg v' | |
| 1641 | val th2 = norm_hyps th2 | |
| 1642 | val cvty = ctyp_of_term cv | |
| 14518 
c3019a66180f
Added a number of explicit type casts and delayed evaluations (all seemingly
 skalberg parents: 
14516diff
changeset | 1643 | val c = HOLogic.dest_Trueprop (concl_of th2) | 
| 14516 | 1644 | val cc = cterm_of sg c | 
| 1645 | val a = case concl_of th1 of | |
| 1646 | 		    _ $ (Const("Ex",_) $ a) => a
 | |
| 1647 | | _ => raise ERR "CHOOSE" "Conclusion not existential" | |
| 1648 | val ca = cterm_of (sign_of_thm th1) a | |
| 15531 | 1649 | val choose_thm' = beta_eta_thm (Drule.instantiate' [SOME cvty] [SOME ca,SOME cc] choose_thm) | 
| 14516 | 1650 | val th21 = rearrange sg (HOLogic.mk_Trueprop (a $ v')) th2 | 
| 1651 | val th22 = strip ((nprems_of th21)-1) (cprems_of th21) th21 | |
| 1652 | val th23 = beta_eta_thm (forall_intr cv th22) | |
| 1653 | val th11 = implies_elim_all (beta_eta_thm th1) | |
| 1654 | val th' = th23 COMP (th11 COMP choose_thm') | |
| 1655 | val th = implies_intr_hyps th' | |
| 1656 | val res = HOLThm(rens_of info',th) | |
| 1657 | val _ = message "RESULT:" | |
| 1658 | val _ = if_debug pth res | |
| 1659 | in | |
| 1660 | (thy,res) | |
| 1661 | end | |
| 1662 | ||
| 1663 | fun GEN v hth thy = | |
| 1664 | let | |
| 1665 | val _ = message "GEN:" | |
| 1666 | val _ = if_debug prin v | |
| 1667 | val _ = if_debug pth hth | |
| 1668 | val (info,th) = disamb_thm hth | |
| 1669 | val (info',v') = disamb_term_from info v | |
| 1670 | val res = HOLThm(rens_of info',mk_GEN v' th (sign_of thy)) | |
| 1671 | val _ = message "RESULT:" | |
| 1672 | val _ = if_debug pth res | |
| 1673 | in | |
| 1674 | (thy,res) | |
| 1675 | end | |
| 1676 | ||
| 1677 | fun SPEC tm hth thy = | |
| 1678 | let | |
| 1679 | val _ = message "SPEC:" | |
| 1680 | val _ = if_debug prin tm | |
| 1681 | val _ = if_debug pth hth | |
| 1682 | val (info,th) = disamb_thm hth | |
| 1683 | val (info',tm') = disamb_term_from info tm | |
| 1684 | val sg = sign_of thy | |
| 1685 | val ctm = Thm.cterm_of sg tm' | |
| 1686 | val cty = Thm.ctyp_of_term ctm | |
| 15531 | 1687 | val spec' = Drule.instantiate' [SOME cty] [NONE,SOME ctm] spec_thm | 
| 14516 | 1688 | val th = th RS spec' | 
| 1689 | val res = HOLThm(rens_of info',th) | |
| 1690 | val _ = message "RESULT:" | |
| 1691 | val _ = if_debug pth res | |
| 1692 | in | |
| 1693 | (thy,res) | |
| 1694 | end | |
| 1695 | ||
| 1696 | fun COMB hth1 hth2 thy = | |
| 1697 | let | |
| 1698 | val _ = message "COMB:" | |
| 1699 | val _ = if_debug pth hth1 | |
| 1700 | val _ = if_debug pth hth2 | |
| 1701 | val (info,[th1,th2]) = disamb_thms [hth1,hth2] | |
| 1702 | val sg = sign_of thy | |
| 1703 | val res = HOLThm(rens_of info,mk_COMB th1 th2 sg) | |
| 1704 | val _ = message "RESULT:" | |
| 1705 | val _ = if_debug pth res | |
| 1706 | in | |
| 1707 | (thy,res) | |
| 1708 | end | |
| 1709 | ||
| 1710 | fun TRANS hth1 hth2 thy = | |
| 1711 | let | |
| 1712 | val _ = message "TRANS:" | |
| 1713 | val _ = if_debug pth hth1 | |
| 1714 | val _ = if_debug pth hth2 | |
| 1715 | val (info,[th1,th2]) = disamb_thms [hth1,hth2] | |
| 1716 | val th = [th1,th2] MRS trans_thm | |
| 1717 | val res = HOLThm(rens_of info,th) | |
| 1718 | val _ = message "RESULT:" | |
| 1719 | val _ = if_debug pth res | |
| 1720 | in | |
| 1721 | (thy,res) | |
| 1722 | end | |
| 1723 | ||
| 1724 | ||
| 1725 | fun CCONTR tm hth thy = | |
| 1726 | let | |
| 1727 | val _ = message "SPEC:" | |
| 1728 | val _ = if_debug prin tm | |
| 1729 | val _ = if_debug pth hth | |
| 1730 | val (info,th) = disamb_thm hth | |
| 1731 | val (info',tm') = disamb_term_from info tm | |
| 1732 | val th = norm_hyps th | |
| 1733 | val sg = sign_of thy | |
| 1734 | val ct = cterm_of sg tm' | |
| 1735 | 	val th1 = rearrange sg (HOLogic.mk_Trueprop (Const("Not",boolT-->boolT) $ tm')) th
 | |
| 15531 | 1736 | val ccontr_thm' = Drule.instantiate' [] [SOME ct] ccontr_thm | 
| 14516 | 1737 | val res1 = uniq_compose ((nprems_of th1) - 1) th1 1 ccontr_thm' | 
| 1738 | val res = HOLThm(rens_of info',res1) | |
| 1739 | val _ = message "RESULT:" | |
| 1740 | val _ = if_debug pth res | |
| 1741 | in | |
| 1742 | (thy,res) | |
| 1743 | end | |
| 1744 | ||
| 1745 | fun mk_ABS v th sg = | |
| 1746 | let | |
| 1747 | val cv = cterm_of sg v | |
| 1748 | val th1 = implies_elim_all (beta_eta_thm th) | |
| 1749 | val (f,g) = case concl_of th1 of | |
| 1750 | 			_ $ (Const("op =",_) $ f $ g) => (Term.lambda v f,Term.lambda v g)
 | |
| 1751 | | _ => raise ERR "mk_ABS" "Bad conclusion" | |
| 1752 | val (fd,fr) = dom_rng (type_of f) | |
| 15531 | 1753 | val abs_thm' = Drule.instantiate' [SOME (ctyp_of sg fd), SOME (ctyp_of sg fr)] [SOME (cterm_of sg f), SOME (cterm_of sg g)] abs_thm | 
| 14516 | 1754 | val th2 = forall_intr cv th1 | 
| 1755 | val th3 = th2 COMP abs_thm' | |
| 1756 | val res = implies_intr_hyps th3 | |
| 1757 | in | |
| 1758 | res | |
| 1759 | end | |
| 1760 | ||
| 1761 | fun ABS v hth thy = | |
| 1762 | let | |
| 1763 | val _ = message "ABS:" | |
| 1764 | val _ = if_debug prin v | |
| 1765 | val _ = if_debug pth hth | |
| 1766 | val (info,th) = disamb_thm hth | |
| 1767 | val (info',v') = disamb_term_from info v | |
| 1768 | val sg = sign_of thy | |
| 1769 | val res = HOLThm(rens_of info',mk_ABS v' th sg) | |
| 1770 | val _ = message "RESULT:" | |
| 1771 | val _ = if_debug pth res | |
| 1772 | in | |
| 1773 | (thy,res) | |
| 1774 | end | |
| 1775 | ||
| 1776 | fun GEN_ABS copt vlist hth thy = | |
| 1777 | let | |
| 1778 | val _ = message "GEN_ABS:" | |
| 1779 | val _ = case copt of | |
| 15531 | 1780 | SOME c => if_debug prin c | 
| 1781 | | NONE => () | |
| 14516 | 1782 | val _ = if_debug (app prin) vlist | 
| 1783 | val _ = if_debug pth hth | |
| 1784 | val (info,th) = disamb_thm hth | |
| 1785 | val (info',vlist') = disamb_terms_from info vlist | |
| 1786 | val sg = sign_of thy | |
| 1787 | val th1 = | |
| 1788 | case copt of | |
| 15531 | 1789 | SOME (c as Const(cname,cty)) => | 
| 14516 | 1790 | let | 
| 1791 | fun inst_type ty1 ty2 (TVar _) = raise ERR "GEN_ABS" "Type variable found!" | |
| 1792 | | inst_type ty1 ty2 (ty as TFree _) = if ty1 = ty | |
| 1793 | then ty2 | |
| 1794 | else ty | |
| 1795 | | inst_type ty1 ty2 (ty as Type(name,tys)) = | |
| 1796 | Type(name,map (inst_type ty1 ty2) tys) | |
| 1797 | in | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1798 | foldr (fn (v,th) => | 
| 14516 | 1799 | let | 
| 1800 | val cdom = fst (dom_rng (fst (dom_rng cty))) | |
| 1801 | val vty = type_of v | |
| 1802 | val newcty = inst_type cdom vty cty | |
| 1803 | val cc = cterm_of sg (Const(cname,newcty)) | |
| 1804 | in | |
| 1805 | mk_COMB (mk_REFL cc) (mk_ABS v th sg) sg | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1806 | end) th vlist' | 
| 14516 | 1807 | end | 
| 15531 | 1808 | | SOME _ => raise ERR "GEN_ABS" "Bad constant" | 
| 1809 | | NONE => | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1810 | foldr (fn (v,th) => mk_ABS v th sg) th vlist' | 
| 14516 | 1811 | val res = HOLThm(rens_of info',th1) | 
| 1812 | val _ = message "RESULT:" | |
| 1813 | val _ = if_debug pth res | |
| 1814 | in | |
| 1815 | (thy,res) | |
| 1816 | end | |
| 1817 | ||
| 1818 | fun NOT_INTRO (hth as HOLThm(rens,th)) thy = | |
| 1819 | let | |
| 1820 | val _ = message "NOT_INTRO:" | |
| 1821 | val _ = if_debug pth hth | |
| 1822 | val sg = sign_of thy | |
| 1823 | val th1 = implies_elim_all (beta_eta_thm th) | |
| 1824 | val a = case concl_of th1 of | |
| 1825 | 		    _ $ (Const("op -->",_) $ a $ Const("False",_)) => a
 | |
| 1826 | | _ => raise ERR "NOT_INTRO" "Conclusion of bad form" | |
| 1827 | val ca = cterm_of sg a | |
| 15531 | 1828 | val th2 = equal_elim (Drule.instantiate' [] [SOME ca] not_intro_thm) th1 | 
| 14516 | 1829 | val res = HOLThm(rens,implies_intr_hyps th2) | 
| 1830 | val _ = message "RESULT:" | |
| 1831 | val _ = if_debug pth res | |
| 1832 | in | |
| 1833 | (thy,res) | |
| 1834 | end | |
| 1835 | ||
| 1836 | fun NOT_ELIM (hth as HOLThm(rens,th)) thy = | |
| 1837 | let | |
| 1838 | val _ = message "NOT_INTRO:" | |
| 1839 | val _ = if_debug pth hth | |
| 1840 | val sg = sign_of thy | |
| 1841 | val th1 = implies_elim_all (beta_eta_thm th) | |
| 1842 | val a = case concl_of th1 of | |
| 1843 | 		    _ $ (Const("Not",_) $ a) => a
 | |
| 1844 | | _ => raise ERR "NOT_ELIM" "Conclusion of bad form" | |
| 1845 | val ca = cterm_of sg a | |
| 15531 | 1846 | val th2 = equal_elim (Drule.instantiate' [] [SOME ca] not_elim_thm) th1 | 
| 14516 | 1847 | val res = HOLThm(rens,implies_intr_hyps th2) | 
| 1848 | val _ = message "RESULT:" | |
| 1849 | val _ = if_debug pth res | |
| 1850 | in | |
| 1851 | (thy,res) | |
| 1852 | end | |
| 1853 | ||
| 1854 | fun DISCH tm hth thy = | |
| 1855 | let | |
| 1856 | val _ = message "DISCH:" | |
| 1857 | val _ = if_debug prin tm | |
| 1858 | val _ = if_debug pth hth | |
| 1859 | val (info,th) = disamb_thm hth | |
| 1860 | val (info',tm') = disamb_term_from info tm | |
| 1861 | val prems = prems_of th | |
| 1862 | val sg = sign_of thy | |
| 1863 | val th1 = beta_eta_thm th | |
| 1864 | val th2 = implies_elim_all th1 | |
| 1865 | val th3 = implies_intr (cterm_of sg (HOLogic.mk_Trueprop tm')) th2 | |
| 1866 | val th4 = th3 COMP disch_thm | |
| 1867 | val res = HOLThm(rens_of info',implies_intr_hyps th4) | |
| 1868 | val _ = message "RESULT:" | |
| 1869 | val _ = if_debug pth res | |
| 1870 | in | |
| 1871 | (thy,res) | |
| 1872 | end | |
| 1873 | ||
| 1874 | val spaces = String.concat o separate " " | |
| 1875 | ||
| 1876 | fun new_definition thyname constname rhs thy = | |
| 1877 | let | |
| 1878 | val constname = rename_const thyname thy constname | |
| 14685 | 1879 | val sg = sign_of thy | 
| 15570 | 1880 | val redeclared = isSome (Sign.const_type sg (Sign.intern_const sg constname)); | 
| 14516 | 1881 | 	val _ = warning ("Introducing constant " ^ constname)
 | 
| 1882 | val (thmname,thy) = get_defname thyname constname thy | |
| 1883 | val (info,rhs') = disamb_term rhs | |
| 1884 | val ctype = type_of rhs' | |
| 14685 | 1885 | val csyn = mk_syn thy constname | 
| 14516 | 1886 | val thy1 = case HOL4DefThy.get thy of | 
| 1887 | Replaying _ => thy | |
| 1888 | | _ => Theory.add_consts_i [(constname,ctype,csyn)] thy | |
| 1889 | val eq = mk_defeq constname rhs' thy1 | |
| 1890 | val (thy2,thms) = PureThy.add_defs_i false [((thmname,eq),[])] thy1 | |
| 1891 | val def_thm = hd thms | |
| 1892 | val thm' = def_thm RS meta_eq_to_obj_eq_thm | |
| 1893 | val (thy',th) = (thy2, thm') | |
| 1894 | val fullcname = Sign.intern_const (sign_of thy') constname | |
| 1895 | val thy'' = add_hol4_const_mapping thyname constname true fullcname thy' | |
| 1896 | val (linfo,tm24) = disamb_term (mk_teq constname rhs' thy'') | |
| 1897 | val sg = sign_of thy'' | |
| 1898 | val rew = rewrite_hol4_term eq thy'' | |
| 1899 | val crhs = cterm_of sg (#2 (Logic.dest_equals (prop_of rew))) | |
| 14685 | 1900 | val thy22 = if (def_name constname) = thmname andalso not redeclared andalso csyn = NoSyn | 
| 14516 | 1901 | then | 
| 14980 | 1902 | 			add_dump ("constdefs\n  " ^ (quotename constname) ^ " :: \"" ^ string_of_ctyp (ctyp_of sg ctype) ^ "\" " ^ (Syntax.string_of_mixfix csyn) ^ "\n  " ^ (smart_string_of_cterm crhs)) thy''
 | 
| 14516 | 1903 | else | 
| 14980 | 1904 | 			add_dump ("consts\n  " ^ (quotename constname) ^ " :: \"" ^ string_of_ctyp (ctyp_of sg ctype) ^
 | 
| 14516 | 1905 | "\" " ^ (Syntax.string_of_mixfix csyn) ^ "\n\ndefs\n " ^ (quotename thmname) ^ ": " ^ (smart_string_of_cterm crhs)) | 
| 1906 | thy'' | |
| 1907 | ||
| 1908 | 	val hth = case Shuffler.set_prop thy22 (HOLogic.mk_Trueprop tm24) [("",th)] of
 | |
| 15531 | 1909 | SOME (_,res) => HOLThm(rens_of linfo,res) | 
| 1910 | | NONE => raise ERR "new_definition" "Bad conclusion" | |
| 14516 | 1911 | val fullname = Sign.full_name sg thmname | 
| 1912 | val thy22' = case opt_get_output_thy thy22 of | |
| 1913 | "" => add_hol4_mapping thyname thmname fullname thy22 | |
| 1914 | | output_thy => | |
| 1915 | let | |
| 1916 | val moved_thmname = output_thy ^ "." ^ thyname ^ "." ^ thmname | |
| 1917 | in | |
| 1918 | thy22 |> add_hol4_move fullname moved_thmname | |
| 1919 | |> add_hol4_mapping thyname thmname moved_thmname | |
| 1920 | end | |
| 1921 | val _ = message "new_definition:" | |
| 1922 | val _ = if_debug pth hth | |
| 1923 | in | |
| 1924 | (thy22',hth) | |
| 1925 | end | |
| 1926 | handle e => (message "exception in new_definition"; print_exn e) | |
| 1927 | ||
| 1928 | val commafy = String.concat o separate ", " | |
| 1929 | ||
| 1930 | local | |
| 1931 | val helper = thm "termspec_help" | |
| 1932 | in | |
| 1933 | fun new_specification thyname thmname names hth thy = | |
| 1934 | case HOL4DefThy.get thy of | |
| 1935 | Replaying _ => (thy,hth) | |
| 1936 | | _ => | |
| 1937 | let | |
| 1938 | val _ = message "NEW_SPEC:" | |
| 1939 | val _ = if_debug pth hth | |
| 1940 | val names = map (rename_const thyname thy) names | |
| 1941 | 	    val _ = warning ("Introducing constants " ^ (commafy names))
 | |
| 1942 | val (HOLThm(rens,th)) = norm_hthm (sign_of thy) hth | |
| 1943 | val thy1 = case HOL4DefThy.get thy of | |
| 1944 | Replaying _ => thy | |
| 1945 | | _ => | |
| 1946 | let | |
| 1947 | fun dest_eta_abs (Abs(x,xT,body)) = (x,xT,body) | |
| 1948 | | dest_eta_abs body = | |
| 1949 | let | |
| 1950 | val (dT,rT) = dom_rng (type_of body) | |
| 1951 | in | |
| 1952 | 				       ("x",dT,body $ Bound 0)
 | |
| 1953 | end | |
| 1954 | handle TYPE _ => raise ERR "new_specification" "not an abstraction type" | |
| 1955 | 			       fun dest_exists (Const("Ex",_) $ abody) =
 | |
| 1956 | dest_eta_abs abody | |
| 1957 | | dest_exists tm = | |
| 1958 | raise ERR "new_specification" "Bad existential formula" | |
| 1959 | ||
| 15570 | 1960 | val (consts,_) = Library.foldl (fn ((cs,ex),cname) => | 
| 14516 | 1961 | let | 
| 1962 | val (_,cT,p) = dest_exists ex | |
| 1963 | in | |
| 14685 | 1964 | ((cname,cT,mk_syn thy cname)::cs,p) | 
| 14516 | 1965 | end) (([],HOLogic.dest_Trueprop (concl_of th)),names) | 
| 1966 | val sg = sign_of thy | |
| 15570 | 1967 | val str = Library.foldl (fn (acc,(c,T,csyn)) => | 
| 14980 | 1968 | 						   acc ^ "\n  " ^ (quotename c) ^ " :: \"" ^ string_of_ctyp (ctyp_of sg T) ^ "\" " ^ (Syntax.string_of_mixfix csyn)) ("consts",consts)
 | 
| 14516 | 1969 | val thy' = add_dump str thy | 
| 1970 | in | |
| 1971 | Theory.add_consts_i consts thy' | |
| 1972 | end | |
| 1973 | ||
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1974 | val thy1 = foldr (fn(name,thy)=> | 
| 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1975 | snd (get_defname thyname name thy)) thy1 names | 
| 14516 | 1976 | fun new_name name = fst (get_defname thyname name thy1) | 
| 15531 | 1977 | val (thy',res) = SpecificationPackage.add_specification_i NONE | 
| 14516 | 1978 | (map (fn name => (new_name name,name,false)) names) | 
| 1979 | (thy1,th) | |
| 1980 | val res' = Drule.freeze_all res | |
| 1981 | val hth = HOLThm(rens,res') | |
| 1982 | val rew = rewrite_hol4_term (concl_of res') thy' | |
| 1983 | val th = equal_elim rew res' | |
| 1984 | fun handle_const (name,thy) = | |
| 1985 | let | |
| 1986 | val defname = def_name name | |
| 1987 | val (newname,thy') = get_defname thyname name thy | |
| 1988 | in | |
| 1989 | (if defname = newname | |
| 1990 | then quotename name | |
| 1991 | else (quotename newname) ^ ": " ^ (quotename name),thy') | |
| 1992 | end | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1993 | val (new_names,thy') = foldr (fn(name,(names,thy)) => | 
| 14516 | 1994 | let | 
| 1995 | val (name',thy') = handle_const (name,thy) | |
| 1996 | in | |
| 1997 | (name'::names,thy') | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 1998 | end) ([],thy') names | 
| 14516 | 1999 | 	    val thy'' = add_dump ("specification (" ^ (spaces new_names) ^ ") " ^ thmname ^ ": " ^ (smart_string_of_thm th) ^
 | 
| 2000 | "\n by (import " ^ thyname ^ " " ^ thmname ^ ")") | |
| 2001 | thy' | |
| 2002 | val _ = message "RESULT:" | |
| 2003 | val _ = if_debug pth hth | |
| 2004 | in | |
| 2005 | intern_store_thm false thyname thmname hth thy'' | |
| 2006 | end | |
| 2007 | handle e => (message "exception in new_specification"; print_exn e) | |
| 2008 | ||
| 2009 | end | |
| 2010 | ||
| 2011 | fun new_axiom name tm thy = raise ERR "new_axiom" ("Oh, no you don't! (" ^ name ^ ")")
 | |
| 2012 | ||
| 2013 | fun to_isa_thm (hth as HOLThm(_,th)) = | |
| 2014 | let | |
| 2015 | val (HOLThm args) = norm_hthm (sign_of_thm th) hth | |
| 2016 | in | |
| 2017 | apsnd strip_shyps args | |
| 2018 | end | |
| 2019 | ||
| 2020 | fun to_isa_term tm = tm | |
| 2021 | ||
| 2022 | local | |
| 2023 | val light_nonempty = thm "light_ex_imp_nonempty" | |
| 2024 | val ex_imp_nonempty = thm "ex_imp_nonempty" | |
| 2025 | val typedef_hol2hol4 = thm "typedef_hol2hol4" | |
| 2026 | val typedef_hol2hollight = thm "typedef_hol2hollight" | |
| 2027 | in | |
| 2028 | fun new_type_definition thyname thmname tycname hth thy = | |
| 2029 | case HOL4DefThy.get thy of | |
| 2030 | Replaying _ => (thy,hth) | |
| 2031 | | _ => | |
| 2032 | let | |
| 2033 | val _ = message "TYPE_DEF:" | |
| 2034 | val _ = if_debug pth hth | |
| 2035 | 	    val _ = warning ("Introducing type " ^ tycname)
 | |
| 2036 | val (HOLThm(rens,td_th)) = norm_hthm (sign_of thy) hth | |
| 2037 | val th2 = beta_eta_thm (td_th RS ex_imp_nonempty) | |
| 2038 | val c = case concl_of th2 of | |
| 2039 | 			_ $ (Const("Ex",_) $ Abs(_,_,Const("op :",_) $ _ $ c)) => c
 | |
| 2040 | | _ => raise ERR "new_type_definition" "Bad type definition theorem" | |
| 2041 | val tfrees = term_tfrees c | |
| 2042 | val tnames = map fst tfrees | |
| 14685 | 2043 | val tsyn = mk_syn thy tycname | 
| 14516 | 2044 | val typ = (tycname,tnames,tsyn) | 
| 15531 | 2045 | val (thy',typedef_info) = TypedefPackage.add_typedef_i false (SOME thmname) typ c NONE (rtac th2 1) thy | 
| 14516 | 2046 | |
| 2047 | val th3 = (#type_definition typedef_info) RS typedef_hol2hol4 | |
| 2048 | ||
| 2049 | val fulltyname = Sign.intern_tycon (sign_of thy') tycname | |
| 2050 | val thy'' = add_hol4_type_mapping thyname tycname true fulltyname thy' | |
| 2051 | ||
| 2052 | val sg = sign_of thy'' | |
| 2053 | val (hth' as HOLThm args) = norm_hthm sg (HOLThm(rens,th3)) | |
| 2054 | 	    val _ = if has_ren hth' then warning ("Theorem " ^ thmname ^ " needs variable-disambiguating")
 | |
| 2055 | else () | |
| 2056 | val thy4 = add_hol4_pending thyname thmname args thy'' | |
| 2057 | ||
| 2058 | val sg = sign_of thy4 | |
| 2059 | val rew = rewrite_hol4_term (concl_of td_th) thy4 | |
| 2060 | val th = equal_elim rew (transfer_sg sg td_th) | |
| 2061 | val c = case HOLogic.dest_Trueprop (prop_of th) of | |
| 2062 | 			  Const("Ex",exT) $ P =>
 | |
| 2063 | let | |
| 2064 | val PT = domain_type exT | |
| 2065 | in | |
| 2066 | 			      Const("Collect",PT-->HOLogic.mk_setT (domain_type PT)) $ P
 | |
| 2067 | end | |
| 2068 | | _ => error "Internal error in ProofKernel.new_typedefinition" | |
| 2069 | val tnames_string = if null tnames | |
| 2070 | then "" | |
| 2071 | 				else "(" ^ (commafy tnames) ^ ") "
 | |
| 2072 | val proc_prop = if null tnames | |
| 2073 | then smart_string_of_cterm | |
| 2074 | else Library.setmp show_all_types true smart_string_of_cterm | |
| 2075 | 	    val thy5 = add_dump ("typedef (open) " ^ tnames_string ^ (quotename tycname) ^ " = " ^ (proc_prop (cterm_of sg c)) ^ " " ^ (Syntax.string_of_mixfix tsyn) ^ "\n  by (rule typedef_helper,import " ^ thyname ^ " " ^ thmname ^ ")") thy4
 | |
| 2076 | 	    val thy6 = add_dump ("lemmas " ^ thmname ^ " = typedef_hol2hol4 [OF type_definition_" ^ tycname ^ "]")
 | |
| 2077 | thy5 | |
| 2078 | val _ = message "RESULT:" | |
| 2079 | val _ = if_debug pth hth' | |
| 2080 | in | |
| 2081 | (thy6,hth') | |
| 2082 | end | |
| 2083 | handle e => (message "exception in new_type_definition"; print_exn e) | |
| 2084 | ||
| 2085 | fun type_introduction thyname thmname tycname abs_name rep_name (P,t) hth thy = | |
| 2086 | case HOL4DefThy.get thy of | |
| 2087 | Replaying _ => (thy,hth) | |
| 2088 | | _ => | |
| 2089 | let | |
| 2090 | val _ = message "TYPE_INTRO:" | |
| 2091 | val _ = if_debug pth hth | |
| 2092 | 	    val _ = warning ("Introducing type " ^ tycname ^ " (with morphisms " ^ abs_name ^ " and " ^ rep_name ^ ")")
 | |
| 2093 | val (HOLThm(rens,td_th)) = norm_hthm (sign_of thy) hth | |
| 2094 | val sg = sign_of thy | |
| 2095 | val tT = type_of t | |
| 2096 | val light_nonempty' = | |
| 15531 | 2097 | Drule.instantiate' [SOME (ctyp_of sg tT)] | 
| 2098 | [SOME (cterm_of sg P), | |
| 2099 | SOME (cterm_of sg t)] light_nonempty | |
| 14516 | 2100 | val th2 = beta_eta_thm (td_th RS (beta_eta_thm light_nonempty')) | 
| 2101 | val c = case concl_of th2 of | |
| 2102 | 			_ $ (Const("Ex",_) $ Abs(_,_,Const("op :",_) $ _ $ c)) => c
 | |
| 2103 | | _ => raise ERR "type_introduction" "Bad type definition theorem" | |
| 2104 | val tfrees = term_tfrees c | |
| 2105 | val tnames = map fst tfrees | |
| 14685 | 2106 | val tsyn = mk_syn thy tycname | 
| 14516 | 2107 | val typ = (tycname,tnames,tsyn) | 
| 15531 | 2108 | val (thy',typedef_info) = TypedefPackage.add_typedef_i false (SOME thmname) typ c (SOME(rep_name,abs_name)) (rtac th2 1) thy | 
| 14516 | 2109 | |
| 2110 | val th3 = (#type_definition typedef_info) RS typedef_hol2hollight | |
| 2111 | ||
| 2112 | val th4 = Drule.freeze_all th3 | |
| 2113 | val fulltyname = Sign.intern_tycon (sign_of thy') tycname | |
| 2114 | val thy'' = add_hol4_type_mapping thyname tycname true fulltyname thy' | |
| 2115 | ||
| 2116 | val sg = sign_of thy'' | |
| 2117 | val (hth' as HOLThm args) = norm_hthm sg (HOLThm(rens,th4)) | |
| 2118 | val _ = if #maxidx (rep_thm th4) <> ~1 | |
| 2119 | then (Library.setmp show_types true pth hth' ; error "SCHEME!") | |
| 2120 | else () | |
| 2121 | 	    val _ = if has_ren hth' then warning ("Theorem " ^ thmname ^ " needs variable-disambiguating")
 | |
| 2122 | else () | |
| 2123 | val thy4 = add_hol4_pending thyname thmname args thy'' | |
| 2124 | ||
| 2125 | val sg = sign_of thy4 | |
| 2126 | val P' = #2 (Logic.dest_equals (concl_of (rewrite_hol4_term P thy4))) | |
| 2127 | val c = | |
| 2128 | let | |
| 2129 | val PT = type_of P' | |
| 2130 | in | |
| 2131 | 		    Const("Collect",PT-->HOLogic.mk_setT (domain_type PT)) $ P'
 | |
| 2132 | end | |
| 2133 | ||
| 2134 | val tnames_string = if null tnames | |
| 2135 | then "" | |
| 2136 | 				else "(" ^ (commafy tnames) ^ ") "
 | |
| 2137 | val proc_prop = if null tnames | |
| 2138 | then smart_string_of_cterm | |
| 2139 | else Library.setmp show_all_types true smart_string_of_cterm | |
| 2140 | 	    val thy5 = add_dump ("typedef (open) " ^ tnames_string ^ (quotename tycname) ^ " = " ^ (proc_prop (cterm_of sg c)) ^ " " ^ (Syntax.string_of_mixfix tsyn) ^ "\n  by (rule light_ex_imp_nonempty,import " ^ thyname ^ " " ^ thmname ^ ")") thy4
 | |
| 2141 | 	    val thy6 = add_dump ("lemmas " ^ thmname ^ " = typedef_hol2hollight [OF type_definition_" ^ tycname ^ "]")
 | |
| 2142 | thy5 | |
| 2143 | val _ = message "RESULT:" | |
| 2144 | val _ = if_debug pth hth' | |
| 2145 | in | |
| 2146 | (thy6,hth') | |
| 2147 | end | |
| 2148 | handle e => (message "exception in type_introduction"; print_exn e) | |
| 2149 | end | |
| 2150 | ||
| 2151 | end |