| author | haftmann | 
| Fri, 14 Mar 2008 08:52:53 +0100 | |
| changeset 26268 | 80aaf4d034be | 
| parent 25573 | a0e695567236 | 
| child 26714 | 4773b832f1b1 | 
| permissions | -rw-r--r-- | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 1 | (* Title: Pure/variable.ML | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 2 | ID: $Id$ | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 3 | Author: Makarius | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 4 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 5 | Fixed type/term variables and polymorphic term abbreviations. | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 6 | *) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 7 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 8 | signature VARIABLE = | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 9 | sig | 
| 20303 | 10 | val is_body: Proof.context -> bool | 
| 11 | val set_body: bool -> Proof.context -> Proof.context | |
| 12 | val restore_body: Proof.context -> Proof.context -> Proof.context | |
| 13 | val names_of: Proof.context -> Name.context | |
| 14 | val fixes_of: Proof.context -> (string * string) list | |
| 15 | val binds_of: Proof.context -> (typ * term) Vartab.table | |
| 24765 | 16 | val maxidx_of: Proof.context -> int | 
| 20303 | 17 | val constraints_of: Proof.context -> typ Vartab.table * sort Vartab.table | 
| 18 | val is_declared: Proof.context -> string -> bool | |
| 19 | val is_fixed: Proof.context -> string -> bool | |
| 20 | val newly_fixed: Proof.context -> Proof.context -> string -> bool | |
| 21571 | 21 | val add_fixed: Proof.context -> term -> (string * typ) list -> (string * typ) list | 
| 20303 | 22 | val default_type: Proof.context -> string -> typ option | 
| 23 | val def_type: Proof.context -> bool -> indexname -> typ option | |
| 24 | val def_sort: Proof.context -> indexname -> sort option | |
| 25 | val declare_constraints: term -> Proof.context -> Proof.context | |
| 25573 | 26 | val declare_names: term -> Proof.context -> Proof.context | 
| 20303 | 27 | val declare_internal: term -> Proof.context -> Proof.context | 
| 28 | val declare_term: term -> Proof.context -> Proof.context | |
| 29 | val declare_prf: Proofterm.proof -> Proof.context -> Proof.context | |
| 30 | val declare_thm: thm -> Proof.context -> Proof.context | |
| 31 | val thm_context: thm -> Proof.context | |
| 32 | val variant_frees: Proof.context -> term list -> (string * 'a) list -> (string * 'a) list | |
| 33 | val add_binds: (indexname * term option) list -> Proof.context -> Proof.context | |
| 34 | val expand_binds: Proof.context -> term -> term | |
| 25325 | 35 | val lookup_const: Proof.context -> string -> string option | 
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 36 | val is_const: Proof.context -> string -> bool | 
| 25325 | 37 | val declare_const: string * string -> Proof.context -> Proof.context | 
| 20303 | 38 | val add_fixes: string list -> Proof.context -> string list * Proof.context | 
| 39 | val add_fixes_direct: string list -> Proof.context -> Proof.context | |
| 21369 | 40 | val auto_fixes: term -> Proof.context -> Proof.context | 
| 20797 
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
 wenzelm parents: 
20579diff
changeset | 41 | val variant_fixes: string list -> Proof.context -> string list * Proof.context | 
| 20303 | 42 | val invent_types: sort list -> Proof.context -> (string * sort) list * Proof.context | 
| 43 | val export_terms: Proof.context -> Proof.context -> term list -> term list | |
| 44 | val exportT_terms: Proof.context -> Proof.context -> term list -> term list | |
| 45 | val exportT: Proof.context -> Proof.context -> thm list -> thm list | |
| 46 | val export_prf: Proof.context -> Proof.context -> Proofterm.proof -> Proofterm.proof | |
| 47 | val export: Proof.context -> Proof.context -> thm list -> thm list | |
| 21522 | 48 | val export_morphism: Proof.context -> Proof.context -> morphism | 
| 20303 | 49 | val importT_inst: term list -> Proof.context -> ((indexname * sort) * typ) list * Proof.context | 
| 50 | val import_inst: bool -> term list -> Proof.context -> | |
| 51 | (((indexname * sort) * typ) list * ((indexname * typ) * term) list) * Proof.context | |
| 52 | val importT_terms: term list -> Proof.context -> term list * Proof.context | |
| 53 | val import_terms: bool -> term list -> Proof.context -> term list * Proof.context | |
| 22601 | 54 | val importT_thms: thm list -> Proof.context -> (ctyp list * thm list) * Proof.context | 
| 20303 | 55 | val import_prf: bool -> Proofterm.proof -> Proof.context -> Proofterm.proof * Proof.context | 
| 22568 
ed7aa5a350ef
renamed Variable.import to import_thms (avoid clash with Alice keywords);
 wenzelm parents: 
21799diff
changeset | 56 | val import_thms: bool -> thm list -> Proof.context -> | 
| 20303 | 57 | ((ctyp list * cterm list) * thm list) * Proof.context | 
| 21287 | 58 | val tradeT: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list | 
| 59 | val trade: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list | |
| 20303 | 60 | val focus: cterm -> Proof.context -> (cterm list * cterm) * Proof.context | 
| 61 | val focus_subgoal: int -> thm -> Proof.context -> (cterm list * cterm) * Proof.context | |
| 62 | val warn_extra_tfrees: Proof.context -> Proof.context -> unit | |
| 24694 | 63 | val polymorphic_types: Proof.context -> term list -> (indexname * sort) list * term list | 
| 20303 | 64 | val polymorphic: Proof.context -> term list -> term list | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 65 | end; | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 66 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 67 | structure Variable: VARIABLE = | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 68 | struct | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 69 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 70 | (** local context data **) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 71 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 72 | datatype data = Data of | 
| 20102 | 73 |  {is_body: bool,                        (*inner body mode*)
 | 
| 20162 | 74 | names: Name.context, (*type/term variable names*) | 
| 25325 | 75 | consts: string Symtab.table, (*consts within the local scope*) | 
| 20162 | 76 | fixes: (string * string) list, (*term fixes -- extern/intern*) | 
| 20102 | 77 | binds: (typ * term) Vartab.table, (*term bindings*) | 
| 20162 | 78 | type_occs: string list Symtab.table, (*type variables -- possibly within term variables*) | 
| 24765 | 79 | maxidx: int, (*maximum var index*) | 
| 20162 | 80 | constraints: | 
| 20102 | 81 | typ Vartab.table * (*type constraints*) | 
| 20162 | 82 | sort Vartab.table}; (*default sorts*) | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 83 | |
| 25325 | 84 | fun make_data (is_body, names, consts, fixes, binds, type_occs, maxidx, constraints) = | 
| 85 |   Data {is_body = is_body, names = names, consts = consts, fixes = fixes, binds = binds,
 | |
| 24765 | 86 | type_occs = type_occs, maxidx = maxidx, constraints = constraints}; | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 87 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 88 | structure Data = ProofDataFun | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 89 | ( | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 90 | type T = data; | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 91 | fun init thy = | 
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 92 | make_data (false, Name.context, Symtab.empty, [], Vartab.empty, Symtab.empty, | 
| 24765 | 93 | ~1, (Vartab.empty, Vartab.empty)); | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 94 | ); | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 95 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 96 | fun map_data f = | 
| 25325 | 97 |   Data.map (fn Data {is_body, names, consts, fixes, binds, type_occs, maxidx, constraints} =>
 | 
| 98 | make_data (f (is_body, names, consts, fixes, binds, type_occs, maxidx, constraints))); | |
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 99 | |
| 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 100 | fun map_names f = | 
| 25325 | 101 | map_data (fn (is_body, names, consts, fixes, binds, type_occs, maxidx, constraints) => | 
| 102 | (is_body, f names, consts, fixes, binds, type_occs, maxidx, constraints)); | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 103 | |
| 25325 | 104 | fun map_consts f = | 
| 105 | map_data (fn (is_body, names, consts, fixes, binds, type_occs, maxidx, constraints) => | |
| 106 | (is_body, names, f consts, fixes, binds, type_occs, maxidx, constraints)); | |
| 20162 | 107 | |
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 108 | fun map_fixes f = | 
| 25325 | 109 | map_data (fn (is_body, names, consts, fixes, binds, type_occs, maxidx, constraints) => | 
| 110 | (is_body, names, consts, f fixes, binds, type_occs, maxidx, constraints)); | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 111 | |
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 112 | fun map_binds f = | 
| 25325 | 113 | map_data (fn (is_body, names, consts, fixes, binds, type_occs, maxidx, constraints) => | 
| 114 | (is_body, names, consts, fixes, f binds, type_occs, maxidx, constraints)); | |
| 24765 | 115 | |
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 116 | fun map_type_occs f = | 
| 25325 | 117 | map_data (fn (is_body, names, consts, fixes, binds, type_occs, maxidx, constraints) => | 
| 118 | (is_body, names, consts, fixes, binds, f type_occs, maxidx, constraints)); | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 119 | |
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 120 | fun map_maxidx f = | 
| 25325 | 121 | map_data (fn (is_body, names, consts, fixes, binds, type_occs, maxidx, constraints) => | 
| 122 | (is_body, names, consts, fixes, binds, type_occs, f maxidx, constraints)); | |
| 20102 | 123 | |
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 124 | fun map_constraints f = | 
| 25325 | 125 | map_data (fn (is_body, names, consts, fixes, binds, type_occs, maxidx, constraints) => | 
| 126 | (is_body, names, consts, fixes, binds, type_occs, maxidx, f constraints)); | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 127 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 128 | fun rep_data ctxt = Data.get ctxt |> (fn Data args => args); | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 129 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 130 | val is_body = #is_body o rep_data; | 
| 24765 | 131 | |
| 25325 | 132 | fun set_body b = map_data (fn (_, names, consts, fixes, binds, type_occs, maxidx, constraints) => | 
| 133 | (b, names, consts, fixes, binds, type_occs, maxidx, constraints)); | |
| 24765 | 134 | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 135 | fun restore_body ctxt = set_body (is_body ctxt); | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 136 | |
| 20102 | 137 | val names_of = #names o rep_data; | 
| 138 | val fixes_of = #fixes o rep_data; | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 139 | val binds_of = #binds o rep_data; | 
| 20162 | 140 | val type_occs_of = #type_occs o rep_data; | 
| 24765 | 141 | val maxidx_of = #maxidx o rep_data; | 
| 20162 | 142 | val constraints_of = #constraints o rep_data; | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 143 | |
| 20162 | 144 | val is_declared = Name.is_declared o names_of; | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 145 | fun is_fixed ctxt x = exists (fn (_, y) => x = y) (fixes_of ctxt); | 
| 20149 | 146 | fun newly_fixed inner outer x = is_fixed inner x andalso not (is_fixed outer x); | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 147 | |
| 21571 | 148 | fun add_fixed ctxt = Term.fold_aterms | 
| 149 | (fn Free (x, T) => if is_fixed ctxt x then insert (op =) (x, T) else I | _ => I); | |
| 150 | ||
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 151 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 152 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 153 | (** declarations **) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 154 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 155 | (* default sorts and types *) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 156 | |
| 20162 | 157 | fun default_type ctxt x = Vartab.lookup (#1 (constraints_of ctxt)) (x, ~1); | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 158 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 159 | fun def_type ctxt pattern xi = | 
| 20162 | 160 |   let val {binds, constraints = (types, _), ...} = rep_data ctxt in
 | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 161 | (case Vartab.lookup types xi of | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 162 | NONE => | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 163 | if pattern then NONE | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 164 | else Vartab.lookup binds xi |> Option.map (TypeInfer.polymorphicT o #1) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 165 | | some => some) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 166 | end; | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 167 | |
| 20162 | 168 | val def_sort = Vartab.lookup o #2 o constraints_of; | 
| 169 | ||
| 170 | ||
| 171 | (* names *) | |
| 172 | ||
| 24765 | 173 | fun declare_type_names t = | 
| 174 | map_names (fold_types (fold_atyps (fn TFree (a, _) => Name.declare a | _ => I)) t) #> | |
| 175 | map_maxidx (fold_types Term.maxidx_typ t); | |
| 20162 | 176 | |
| 177 | fun declare_names t = | |
| 178 | declare_type_names t #> | |
| 24765 | 179 | map_names (fold_aterms (fn Free (x, _) => Name.declare x | _ => I) t) #> | 
| 180 | map_maxidx (Term.maxidx_term t); | |
| 20162 | 181 | |
| 182 | ||
| 183 | (* type occurrences *) | |
| 184 | ||
| 24719 
21d1cdab2d8c
declare_constraints: declare (fix) type variables within constraints, but not terms themselves;
 wenzelm parents: 
24694diff
changeset | 185 | fun decl_type_occsT T = fold_atyps (fn TFree (a, _) => Symtab.default (a, []) | _ => I) T; | 
| 
21d1cdab2d8c
declare_constraints: declare (fix) type variables within constraints, but not terms themselves;
 wenzelm parents: 
24694diff
changeset | 186 | |
| 22671 | 187 | val decl_type_occs = fold_term_types | 
| 20162 | 188 | (fn Free (x, _) => fold_atyps (fn TFree (a, _) => Symtab.insert_list (op =) (a, x) | _ => I) | 
| 24719 
21d1cdab2d8c
declare_constraints: declare (fix) type variables within constraints, but not terms themselves;
 wenzelm parents: 
24694diff
changeset | 189 | | _ => decl_type_occsT); | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 190 | |
| 24719 
21d1cdab2d8c
declare_constraints: declare (fix) type variables within constraints, but not terms themselves;
 wenzelm parents: 
24694diff
changeset | 191 | val declare_type_occsT = map_type_occs o fold_types decl_type_occsT; | 
| 22671 | 192 | val declare_type_occs = map_type_occs o decl_type_occs; | 
| 193 | ||
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 194 | |
| 20162 | 195 | (* constraints *) | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 196 | |
| 21355 | 197 | fun constrain_tvar (xi, S) = | 
| 198 | if S = dummyS then Vartab.delete_safe xi else Vartab.update (xi, S); | |
| 199 | ||
| 20162 | 200 | fun declare_constraints t = map_constraints (fn (types, sorts) => | 
| 201 | let | |
| 202 | val types' = fold_aterms | |
| 203 | (fn Free (x, T) => Vartab.update ((x, ~1), T) | |
| 204 | | Var v => Vartab.update v | |
| 205 | | _ => I) t types; | |
| 206 | val sorts' = fold_types (fold_atyps | |
| 21355 | 207 | (fn TFree (x, S) => constrain_tvar ((x, ~1), S) | 
| 208 | | TVar v => constrain_tvar v | |
| 20162 | 209 | | _ => I)) t sorts; | 
| 210 | in (types', sorts') end) | |
| 24719 
21d1cdab2d8c
declare_constraints: declare (fix) type variables within constraints, but not terms themselves;
 wenzelm parents: 
24694diff
changeset | 211 | #> declare_type_occsT t | 
| 22711 | 212 | #> declare_type_names t; | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 213 | |
| 20162 | 214 | |
| 215 | (* common declarations *) | |
| 216 | ||
| 217 | fun declare_internal t = | |
| 218 | declare_names t #> | |
| 219 | declare_type_occs t; | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 220 | |
| 20162 | 221 | fun declare_term t = | 
| 222 | declare_internal t #> | |
| 223 | declare_constraints t; | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 224 | |
| 20303 | 225 | val declare_prf = Proofterm.fold_proof_terms declare_internal (declare_internal o Logic.mk_type); | 
| 226 | ||
| 22691 | 227 | val declare_thm = Thm.fold_terms declare_internal; | 
| 21522 | 228 | fun thm_context th = declare_thm th (ProofContext.init (Thm.theory_of_thm th)); | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 229 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 230 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 231 | (* renaming term/type frees *) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 232 | |
| 20162 | 233 | fun variant_frees ctxt ts frees = | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 234 | let | 
| 20162 | 235 | val names = names_of (fold declare_names ts ctxt); | 
| 20084 | 236 | val xs = fst (Name.variants (map #1 frees) names); | 
| 237 | in xs ~~ map snd frees end; | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 238 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 239 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 240 | |
| 20303 | 241 | (** term bindings **) | 
| 242 | ||
| 243 | fun add_bind (xi, NONE) = map_binds (Vartab.delete_safe xi) | |
| 244 | | add_bind ((x, i), SOME t) = | |
| 245 | let | |
| 25051 | 246 | val u = Term.close_schematic_term t; | 
| 247 | val U = Term.fastype_of u; | |
| 248 | in declare_term u #> map_binds (Vartab.update ((x, i), (U, u))) end; | |
| 20303 | 249 | |
| 250 | val add_binds = fold add_bind; | |
| 251 | ||
| 252 | fun expand_binds ctxt = | |
| 253 | let | |
| 254 | val binds = binds_of ctxt; | |
| 21799 | 255 | val get = fn Var (xi, _) => Vartab.lookup binds xi | _ => NONE; | 
| 256 | in Envir.beta_norm o Envir.expand_term get end; | |
| 20303 | 257 | |
| 258 | ||
| 259 | ||
| 25325 | 260 | (** consts **) | 
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 261 | |
| 25325 | 262 | val lookup_const = Symtab.lookup o #consts o rep_data; | 
| 263 | val is_const = is_some oo lookup_const; | |
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 264 | |
| 25325 | 265 | val declare_fixed = map_consts o Symtab.delete_safe; | 
| 266 | val declare_const = map_consts o Symtab.update; | |
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 267 | |
| 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 268 | |
| 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 269 | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 270 | (** fixes **) | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 271 | |
| 20084 | 272 | local | 
| 273 | ||
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 274 | fun no_dups [] = () | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 275 |   | no_dups dups = error ("Duplicate fixed variable(s): " ^ commas_quote dups);
 | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 276 | |
| 20102 | 277 | fun new_fixes names' xs xs' = | 
| 278 | map_names (K names') #> | |
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 279 | fold declare_fixed xs #> | 
| 20102 | 280 | map_fixes (fn fixes => (rev (xs ~~ xs') @ fixes)) #> | 
| 20162 | 281 | fold (declare_constraints o Syntax.free) xs' #> | 
| 20084 | 282 | pair xs'; | 
| 283 | ||
| 284 | in | |
| 285 | ||
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 286 | fun add_fixes xs ctxt = | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 287 | let | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 288 | val _ = | 
| 20084 | 289 | (case filter (can Name.dest_skolem) xs of [] => () | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 290 |       | bads => error ("Illegal internal Skolem constant(s): " ^ commas_quote bads));
 | 
| 20084 | 291 | val _ = no_dups (duplicates (op =) xs); | 
| 292 | val (ys, zs) = split_list (fixes_of ctxt); | |
| 20102 | 293 | val names = names_of ctxt; | 
| 20084 | 294 | val (xs', names') = | 
| 20149 | 295 | if is_body ctxt then Name.variants xs names |>> map Name.skolem | 
| 20084 | 296 | else (no_dups (xs inter_string ys); no_dups (xs inter_string zs); | 
| 297 | (xs, fold Name.declare xs names)); | |
| 20102 | 298 | in ctxt |> new_fixes names' xs xs' end; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 299 | |
| 20797 
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
 wenzelm parents: 
20579diff
changeset | 300 | fun variant_fixes raw_xs ctxt = | 
| 20084 | 301 | let | 
| 20102 | 302 | val names = names_of ctxt; | 
| 20149 | 303 | val xs = map Name.clean raw_xs; | 
| 304 | val (xs', names') = Name.variants xs names |>> map Name.skolem; | |
| 20102 | 305 | in ctxt |> new_fixes names' xs xs' end; | 
| 20084 | 306 | |
| 307 | end; | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 308 | |
| 20251 | 309 | |
| 310 | fun add_fixes_direct xs ctxt = ctxt | |
| 311 | |> set_body false | |
| 312 | |> (snd o add_fixes xs) | |
| 313 | |> restore_body ctxt; | |
| 314 | ||
| 315 | fun fix_frees t ctxt = ctxt | |
| 316 | |> add_fixes_direct | |
| 317 | (rev (fold_aterms (fn Free (x, _) => | |
| 21369 | 318 | if is_fixed ctxt x then I else insert (op =) x | _ => I) t [])); | 
| 319 | ||
| 320 | fun auto_fixes t ctxt = | |
| 321 | (if is_body ctxt then ctxt else fix_frees t ctxt) | |
| 20251 | 322 | |> declare_term t; | 
| 323 | ||
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 324 | fun invent_types Ss ctxt = | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 325 | let | 
| 24848 | 326 | val tfrees = Name.invents (names_of ctxt) Name.aT (length Ss) ~~ Ss; | 
| 20162 | 327 | val ctxt' = fold (declare_constraints o Logic.mk_type o TFree) tfrees ctxt; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 328 | in (tfrees, ctxt') end; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 329 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 330 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 331 | |
| 22671 | 332 | (** export -- generalize type/term variables (beware of closure sizes) **) | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 333 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 334 | fun export_inst inner outer = | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 335 | let | 
| 20162 | 336 | val declared_outer = is_declared outer; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 337 | val fixes_inner = fixes_of inner; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 338 | val fixes_outer = fixes_of outer; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 339 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 340 | val gen_fixes = map #2 (Library.take (length fixes_inner - length fixes_outer, fixes_inner)); | 
| 20162 | 341 | val still_fixed = not o member (op =) gen_fixes; | 
| 22671 | 342 | |
| 343 | val type_occs_inner = type_occs_of inner; | |
| 344 | fun gen_fixesT ts = | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 345 | Symtab.fold (fn (a, xs) => | 
| 20162 | 346 | if declared_outer a orelse exists still_fixed xs | 
| 22671 | 347 | then I else cons a) (fold decl_type_occs ts type_occs_inner) []; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 348 | in (gen_fixesT, gen_fixes) end; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 349 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 350 | fun exportT_inst inner outer = #1 (export_inst inner outer); | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 351 | |
| 22671 | 352 | fun exportT_terms inner outer = | 
| 353 | let val mk_tfrees = exportT_inst inner outer in | |
| 354 | fn ts => ts |> map | |
| 355 | (TermSubst.generalize (mk_tfrees ts, []) | |
| 356 | (fold (Term.fold_types Term.maxidx_typ) ts ~1 + 1)) | |
| 357 | end; | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 358 | |
| 22671 | 359 | fun export_terms inner outer = | 
| 360 | let val (mk_tfrees, tfrees) = export_inst inner outer in | |
| 361 | fn ts => ts |> map | |
| 362 | (TermSubst.generalize (mk_tfrees ts, tfrees) | |
| 363 | (fold Term.maxidx_term ts ~1 + 1)) | |
| 364 | end; | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 365 | |
| 20303 | 366 | fun export_prf inner outer prf = | 
| 367 | let | |
| 22671 | 368 | val (mk_tfrees, frees) = export_inst (declare_prf prf inner) outer; | 
| 369 | val tfrees = mk_tfrees []; | |
| 20303 | 370 | val idx = Proofterm.maxidx_proof prf ~1 + 1; | 
| 22671 | 371 | val gen_term = TermSubst.generalize_option (tfrees, frees) idx; | 
| 372 | val gen_typ = TermSubst.generalizeT_option tfrees idx; | |
| 20303 | 373 | in Proofterm.map_proof_terms_option gen_term gen_typ prf end; | 
| 374 | ||
| 22671 | 375 | |
| 376 | fun gen_export (mk_tfrees, frees) ths = | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 377 | let | 
| 22671 | 378 | val tfrees = mk_tfrees (map Thm.full_prop_of ths); | 
| 379 | val maxidx = fold Thm.maxidx_thm ths ~1; | |
| 380 | in map (Thm.generalize (tfrees, frees) (maxidx + 1)) ths end; | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 381 | |
| 22671 | 382 | fun exportT inner outer = gen_export (exportT_inst inner outer, []); | 
| 383 | fun export inner outer = gen_export (export_inst inner outer); | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 384 | |
| 21522 | 385 | fun export_morphism inner outer = | 
| 386 | let | |
| 387 | val fact = export inner outer; | |
| 388 | val term = singleton (export_terms inner outer); | |
| 389 | val typ = Logic.type_map term; | |
| 390 |   in Morphism.morphism {name = I, var = I, typ = typ, term = term, fact = fact} end;
 | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 391 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 392 | |
| 24765 | 393 | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 394 | (** import -- fix schematic type/term variables **) | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 395 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 396 | fun importT_inst ts ctxt = | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 397 | let | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 398 | val tvars = rev (fold Term.add_tvars ts []); | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 399 | val (tfrees, ctxt') = invent_types (map #2 tvars) ctxt; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 400 | in (tvars ~~ map TFree tfrees, ctxt') end; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 401 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 402 | fun import_inst is_open ts ctxt = | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 403 | let | 
| 20198 | 404 | val ren = if is_open then I else Name.internal; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 405 | val (instT, ctxt') = importT_inst ts ctxt; | 
| 20509 | 406 | val vars = map (apsnd (TermSubst.instantiateT instT)) (rev (fold Term.add_vars ts [])); | 
| 20797 
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
 wenzelm parents: 
20579diff
changeset | 407 | val (xs, ctxt'') = variant_fixes (map (ren o #1 o #1) vars) ctxt'; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 408 | val inst = vars ~~ map Free (xs ~~ map #2 vars); | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 409 | in ((instT, inst), ctxt'') end; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 410 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 411 | fun importT_terms ts ctxt = | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 412 | let val (instT, ctxt') = importT_inst ts ctxt | 
| 20509 | 413 | in (map (TermSubst.instantiate (instT, [])) ts, ctxt') end; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 414 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 415 | fun import_terms is_open ts ctxt = | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 416 | let val (inst, ctxt') = import_inst is_open ts ctxt | 
| 20509 | 417 | in (map (TermSubst.instantiate inst) ts, ctxt') end; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 418 | |
| 22601 | 419 | fun importT_thms ths ctxt = | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 420 | let | 
| 21522 | 421 | val thy = ProofContext.theory_of ctxt; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 422 | val certT = Thm.ctyp_of thy; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 423 | val (instT, ctxt') = importT_inst (map Thm.full_prop_of ths) ctxt; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 424 | val instT' = map (fn (v, T) => (certT (TVar v), certT T)) instT; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 425 | val ths' = map (Thm.instantiate (instT', [])) ths; | 
| 20220 | 426 | in ((map #2 instT', ths'), ctxt') end; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 427 | |
| 20303 | 428 | fun import_prf is_open prf ctxt = | 
| 429 | let | |
| 430 | val ts = rev (Proofterm.fold_proof_terms cons (cons o Logic.mk_type) prf []); | |
| 431 | val (insts, ctxt') = import_inst is_open ts ctxt; | |
| 432 | in (Proofterm.instantiate insts prf, ctxt') end; | |
| 433 | ||
| 22568 
ed7aa5a350ef
renamed Variable.import to import_thms (avoid clash with Alice keywords);
 wenzelm parents: 
21799diff
changeset | 434 | fun import_thms is_open ths ctxt = | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 435 | let | 
| 21522 | 436 | val thy = ProofContext.theory_of ctxt; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 437 | val cert = Thm.cterm_of thy; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 438 | val certT = Thm.ctyp_of thy; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 439 | val ((instT, inst), ctxt') = import_inst is_open (map Thm.full_prop_of ths) ctxt; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 440 | val instT' = map (fn (v, T) => (certT (TVar v), certT T)) instT; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 441 | val inst' = map (fn (v, t) => (cert (Var v), cert t)) inst; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 442 | val ths' = map (Thm.instantiate (instT', inst')) ths; | 
| 20220 | 443 | in (((map #2 instT', map #2 inst'), ths'), ctxt') end; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 444 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 445 | |
| 19926 | 446 | (* import/export *) | 
| 447 | ||
| 21287 | 448 | fun gen_trade imp exp f ctxt ths = | 
| 20220 | 449 | let val ((_, ths'), ctxt') = imp ths ctxt | 
| 21287 | 450 | in exp ctxt' ctxt (f ctxt' ths') end; | 
| 19926 | 451 | |
| 22601 | 452 | val tradeT = gen_trade importT_thms exportT; | 
| 22568 
ed7aa5a350ef
renamed Variable.import to import_thms (avoid clash with Alice keywords);
 wenzelm parents: 
21799diff
changeset | 453 | val trade = gen_trade (import_thms true) export; | 
| 19926 | 454 | |
| 455 | ||
| 20162 | 456 | (* focus on outermost parameters *) | 
| 20149 | 457 | |
| 458 | fun forall_elim_prop t prop = | |
| 20579 | 459 | Thm.beta_conversion false (Thm.capply (Thm.dest_arg prop) t) | 
| 460 | |> Thm.cprop_of |> Thm.dest_arg; | |
| 20149 | 461 | |
| 462 | fun focus goal ctxt = | |
| 463 | let | |
| 464 | val cert = Thm.cterm_of (Thm.theory_of_cterm goal); | |
| 465 | val t = Thm.term_of goal; | |
| 20162 | 466 | val ps = Term.variant_frees t (Term.strip_all_vars t); (*as they are printed :-*) | 
| 20149 | 467 | val (xs, Ts) = split_list ps; | 
| 20797 
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
 wenzelm parents: 
20579diff
changeset | 468 | val (xs', ctxt') = variant_fixes xs ctxt; | 
| 20220 | 469 | val ps' = ListPair.map (cert o Free) (xs', Ts); | 
| 470 | val goal' = fold forall_elim_prop ps' goal; | |
| 20149 | 471 | in ((ps', goal'), ctxt') end; | 
| 472 | ||
| 20303 | 473 | fun focus_subgoal i st = | 
| 474 | let | |
| 22691 | 475 | val all_vars = Thm.fold_terms Term.add_vars st []; | 
| 20303 | 476 | val no_binds = map (fn (xi, _) => (xi, NONE)) all_vars; | 
| 477 | in | |
| 478 | add_binds no_binds #> | |
| 479 | fold (declare_constraints o Var) all_vars #> | |
| 480 | focus (Thm.cprem_of st i) | |
| 481 | end; | |
| 482 | ||
| 483 | ||
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 484 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 485 | (** implicit polymorphism **) | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 486 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 487 | (* warn_extra_tfrees *) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 488 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 489 | fun warn_extra_tfrees ctxt1 ctxt2 = | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 490 | let | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 491 | fun occs_typ a = Term.exists_subtype (fn TFree (b, _) => a = b | _ => false); | 
| 20162 | 492 | fun occs_free a x = | 
| 493 | (case def_type ctxt1 false (x, ~1) of | |
| 494 | SOME T => if occs_typ a T then I else cons (a, x) | |
| 495 | | NONE => cons (a, x)); | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 496 | |
| 20162 | 497 | val occs1 = type_occs_of ctxt1; | 
| 498 | val occs2 = type_occs_of ctxt2; | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 499 | val extras = Symtab.fold (fn (a, xs) => | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 500 | if Symtab.defined occs1 a then I else fold (occs_free a) xs) occs2 []; | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 501 | val tfrees = map #1 extras |> sort_distinct string_ord; | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 502 | val frees = map #2 extras |> sort_distinct string_ord; | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 503 | in | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 504 | if null extras then () | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 505 |     else warning ("Introduced fixed type variable(s): " ^ commas tfrees ^ " in " ^
 | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 506 | space_implode " or " (map quote frees)) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 507 | end; | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 508 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 509 | |
| 20149 | 510 | (* polymorphic terms *) | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 511 | |
| 24694 | 512 | fun polymorphic_types ctxt ts = | 
| 20003 | 513 | let | 
| 514 | val ctxt' = fold declare_term ts ctxt; | |
| 20162 | 515 | val occs = type_occs_of ctxt; | 
| 516 | val occs' = type_occs_of ctxt'; | |
| 517 | val types = Symtab.fold (fn (a, _) => if Symtab.defined occs a then I else cons a) occs' []; | |
| 24765 | 518 | val idx = maxidx_of ctxt' + 1; | 
| 24694 | 519 | val Ts' = (fold o fold_types o fold_atyps) | 
| 520 | (fn T as TFree _ => | |
| 521 | (case TermSubst.generalizeT types idx T of TVar v => insert (op =) v | _ => I) | |
| 522 | | _ => I) ts []; | |
| 523 | val ts' = map (TermSubst.generalize (types, []) idx) ts; | |
| 524 | in (rev Ts', ts') end; | |
| 525 | ||
| 526 | fun polymorphic ctxt ts = snd (polymorphic_types ctxt ts); | |
| 20003 | 527 | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 528 | end; |