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