| author | blanchet | 
| Sun, 01 May 2011 18:37:24 +0200 | |
| changeset 42567 | d012947edd36 | 
| parent 42495 | 1af81b70cf09 | 
| child 43326 | 47cf4bc789aa | 
| 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 | Author: Makarius | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 3 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 4 | Fixed type/term variables and polymorphic term abbreviations. | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 5 | *) | 
| 
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 | signature VARIABLE = | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 8 | sig | 
| 20303 | 9 | val is_body: Proof.context -> bool | 
| 10 | val set_body: bool -> Proof.context -> Proof.context | |
| 11 | val restore_body: Proof.context -> Proof.context -> Proof.context | |
| 12 | val names_of: Proof.context -> Name.context | |
| 13 | val binds_of: Proof.context -> (typ * term) Vartab.table | |
| 24765 | 14 | val maxidx_of: Proof.context -> int | 
| 28625 | 15 | val sorts_of: Proof.context -> sort list | 
| 20303 | 16 | val constraints_of: Proof.context -> typ Vartab.table * sort Vartab.table | 
| 17 | val is_declared: Proof.context -> string -> bool | |
| 42494 | 18 | val check_name: binding -> string | 
| 20303 | 19 | val default_type: Proof.context -> string -> typ option | 
| 20 | val def_type: Proof.context -> bool -> indexname -> typ option | |
| 21 | val def_sort: Proof.context -> indexname -> sort option | |
| 27280 | 22 | val declare_names: term -> Proof.context -> Proof.context | 
| 20303 | 23 | val declare_constraints: term -> Proof.context -> Proof.context | 
| 24 | val declare_term: term -> Proof.context -> Proof.context | |
| 27280 | 25 | val declare_typ: typ -> Proof.context -> Proof.context | 
| 20303 | 26 | val declare_prf: Proofterm.proof -> Proof.context -> Proof.context | 
| 27 | val declare_thm: thm -> Proof.context -> Proof.context | |
| 36603 
d5d6111761a6
renamed Variable.thm_context to Variable.global_thm_context to emphasize that this is not the real thing;
 wenzelm parents: 
33957diff
changeset | 28 | val global_thm_context: thm -> Proof.context | 
| 20303 | 29 | val variant_frees: Proof.context -> term list -> (string * 'a) list -> (string * 'a) list | 
| 30756 | 30 | val bind_term: indexname * term option -> Proof.context -> Proof.context | 
| 20303 | 31 | val expand_binds: Proof.context -> term -> term | 
| 25325 | 32 | 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 | 33 | val is_const: Proof.context -> string -> bool | 
| 25325 | 34 | val declare_const: string * string -> Proof.context -> Proof.context | 
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 35 | val is_fixed: Proof.context -> string -> bool | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 36 | val newly_fixed: Proof.context -> Proof.context -> string -> bool | 
| 42493 
01430341fc79
more informative markup for fixed variables (via name space entry);
 wenzelm parents: 
42488diff
changeset | 37 | val fixed_ord: Proof.context -> string * string -> order | 
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 38 | val intern_fixed: Proof.context -> string -> string | 
| 42493 
01430341fc79
more informative markup for fixed variables (via name space entry);
 wenzelm parents: 
42488diff
changeset | 39 | val markup_fixed: Proof.context -> string -> Markup.T | 
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 40 | val lookup_fixed: Proof.context -> string -> string option | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 41 | val revert_fixed: Proof.context -> string -> string | 
| 42482 | 42 | val add_fixed_names: Proof.context -> term -> string list -> string list | 
| 43 | val add_fixed: Proof.context -> term -> (string * typ) list -> (string * typ) list | |
| 44 | val add_free_names: Proof.context -> term -> string list -> string list | |
| 45 | val add_frees: Proof.context -> term -> (string * typ) list -> (string * typ) list | |
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 46 | val add_fixes_binding: binding list -> Proof.context -> string list * Proof.context | 
| 20303 | 47 | val add_fixes: string list -> Proof.context -> string list * Proof.context | 
| 48 | val add_fixes_direct: string list -> Proof.context -> Proof.context | |
| 21369 | 49 | val auto_fixes: term -> Proof.context -> Proof.context | 
| 20797 
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
 wenzelm parents: 
20579diff
changeset | 50 | val variant_fixes: string list -> Proof.context -> string list * Proof.context | 
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 51 | val dest_fixes: Proof.context -> (string * string) list | 
| 20303 | 52 | val invent_types: sort list -> Proof.context -> (string * sort) list * Proof.context | 
| 53 | val export_terms: Proof.context -> Proof.context -> term list -> term list | |
| 54 | val exportT_terms: Proof.context -> Proof.context -> term list -> term list | |
| 55 | val exportT: Proof.context -> Proof.context -> thm list -> thm list | |
| 56 | val export_prf: Proof.context -> Proof.context -> Proofterm.proof -> Proofterm.proof | |
| 57 | val export: Proof.context -> Proof.context -> thm list -> thm list | |
| 21522 | 58 | val export_morphism: Proof.context -> Proof.context -> morphism | 
| 20303 | 59 | val importT_inst: term list -> Proof.context -> ((indexname * sort) * typ) list * Proof.context | 
| 60 | val import_inst: bool -> term list -> Proof.context -> | |
| 61 | (((indexname * sort) * typ) list * ((indexname * typ) * term) list) * Proof.context | |
| 62 | val importT_terms: term list -> Proof.context -> term list * Proof.context | |
| 63 | val import_terms: bool -> term list -> Proof.context -> term list * Proof.context | |
| 32280 
4fb3f426052a
Variable.importT/import: return full instantiations, tuned;
 wenzelm parents: 
32199diff
changeset | 64 | val importT: thm list -> Proof.context -> ((ctyp * ctyp) list * thm list) * Proof.context | 
| 20303 | 65 | val import_prf: bool -> Proofterm.proof -> Proof.context -> Proofterm.proof * Proof.context | 
| 31794 
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
 wenzelm parents: 
30756diff
changeset | 66 | val import: bool -> thm list -> Proof.context -> | 
| 32280 
4fb3f426052a
Variable.importT/import: return full instantiations, tuned;
 wenzelm parents: 
32199diff
changeset | 67 | (((ctyp * ctyp) list * (cterm * cterm) list) * thm list) * Proof.context | 
| 21287 | 68 | val tradeT: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list | 
| 69 | val trade: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list | |
| 42495 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 70 | val focus: term -> Proof.context -> ((string * (string * typ)) list * term) * Proof.context | 
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 71 | val focus_cterm: cterm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context | 
| 32199 | 72 | val focus_subgoal: int -> thm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context | 
| 20303 | 73 | val warn_extra_tfrees: Proof.context -> Proof.context -> unit | 
| 24694 | 74 | val polymorphic_types: Proof.context -> term list -> (indexname * sort) list * term list | 
| 20303 | 75 | val polymorphic: Proof.context -> term list -> term list | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 76 | end; | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 77 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 78 | structure Variable: VARIABLE = | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 79 | struct | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 80 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 81 | (** local context data **) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 82 | |
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 83 | type fixes = string Name_Space.table; | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 84 | val empty_fixes: fixes = Name_Space.empty_table "fixed variable"; | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 85 | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 86 | datatype data = Data of | 
| 20102 | 87 |  {is_body: bool,                        (*inner body mode*)
 | 
| 20162 | 88 | names: Name.context, (*type/term variable names*) | 
| 25325 | 89 | consts: string Symtab.table, (*consts within the local scope*) | 
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 90 | fixes: fixes, (*term fixes -- global name space, intern ~> extern*) | 
| 20102 | 91 | binds: (typ * term) Vartab.table, (*term bindings*) | 
| 20162 | 92 | type_occs: string list Symtab.table, (*type variables -- possibly within term variables*) | 
| 24765 | 93 | maxidx: int, (*maximum var index*) | 
| 40124 | 94 | sorts: sort Ord_List.T, (*declared sort occurrences*) | 
| 20162 | 95 | constraints: | 
| 20102 | 96 | typ Vartab.table * (*type constraints*) | 
| 20162 | 97 | sort Vartab.table}; (*default sorts*) | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 98 | |
| 28625 | 99 | fun make_data (is_body, names, consts, fixes, binds, type_occs, maxidx, sorts, constraints) = | 
| 25325 | 100 |   Data {is_body = is_body, names = names, consts = consts, fixes = fixes, binds = binds,
 | 
| 28625 | 101 | type_occs = type_occs, maxidx = maxidx, sorts = sorts, constraints = constraints}; | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 102 | |
| 33519 | 103 | structure Data = Proof_Data | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 104 | ( | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 105 | type T = data; | 
| 32784 | 106 | fun init _ = | 
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 107 | make_data (false, Name.context, Symtab.empty, empty_fixes, Vartab.empty, | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 108 | Symtab.empty, ~1, [], (Vartab.empty, Vartab.empty)); | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 109 | ); | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 110 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 111 | fun map_data f = | 
| 28625 | 112 |   Data.map (fn Data {is_body, names, consts, fixes, binds, type_occs, maxidx, sorts, constraints} =>
 | 
| 113 | make_data (f (is_body, names, consts, fixes, binds, type_occs, maxidx, sorts, constraints))); | |
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 114 | |
| 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 115 | fun map_names f = | 
| 28625 | 116 | map_data (fn (is_body, names, consts, fixes, binds, type_occs, maxidx, sorts, constraints) => | 
| 117 | (is_body, f names, consts, fixes, binds, type_occs, maxidx, sorts, constraints)); | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 118 | |
| 25325 | 119 | fun map_consts f = | 
| 28625 | 120 | map_data (fn (is_body, names, consts, fixes, binds, type_occs, maxidx, sorts, constraints) => | 
| 121 | (is_body, names, f consts, fixes, binds, type_occs, maxidx, sorts, constraints)); | |
| 20162 | 122 | |
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 123 | fun map_fixes f = | 
| 28625 | 124 | map_data (fn (is_body, names, consts, fixes, binds, type_occs, maxidx, sorts, constraints) => | 
| 125 | (is_body, names, consts, f fixes, binds, type_occs, maxidx, sorts, constraints)); | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 126 | |
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 127 | fun map_binds f = | 
| 28625 | 128 | map_data (fn (is_body, names, consts, fixes, binds, type_occs, maxidx, sorts, constraints) => | 
| 129 | (is_body, names, consts, fixes, f binds, type_occs, maxidx, sorts, constraints)); | |
| 24765 | 130 | |
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 131 | fun map_type_occs f = | 
| 28625 | 132 | map_data (fn (is_body, names, consts, fixes, binds, type_occs, maxidx, sorts, constraints) => | 
| 133 | (is_body, names, consts, fixes, binds, f type_occs, maxidx, sorts, constraints)); | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 134 | |
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 135 | fun map_maxidx f = | 
| 28625 | 136 | map_data (fn (is_body, names, consts, fixes, binds, type_occs, maxidx, sorts, constraints) => | 
| 137 | (is_body, names, consts, fixes, binds, type_occs, f maxidx, sorts, constraints)); | |
| 138 | ||
| 139 | fun map_sorts f = | |
| 140 | map_data (fn (is_body, names, consts, fixes, binds, type_occs, maxidx, sorts, constraints) => | |
| 141 | (is_body, names, consts, fixes, binds, type_occs, maxidx, f sorts, constraints)); | |
| 20102 | 142 | |
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 143 | fun map_constraints f = | 
| 28625 | 144 | map_data (fn (is_body, names, consts, fixes, binds, type_occs, maxidx, sorts, constraints) => | 
| 145 | (is_body, names, consts, fixes, binds, type_occs, maxidx, sorts, f constraints)); | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 146 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 147 | fun rep_data ctxt = Data.get ctxt |> (fn Data args => args); | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 148 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 149 | val is_body = #is_body o rep_data; | 
| 24765 | 150 | |
| 28625 | 151 | fun set_body b = | 
| 152 | map_data (fn (_, names, consts, fixes, binds, type_occs, maxidx, sorts, constraints) => | |
| 153 | (b, names, consts, fixes, binds, type_occs, maxidx, sorts, constraints)); | |
| 24765 | 154 | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 155 | fun restore_body ctxt = set_body (is_body ctxt); | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 156 | |
| 20102 | 157 | val names_of = #names o rep_data; | 
| 158 | val fixes_of = #fixes o rep_data; | |
| 42493 
01430341fc79
more informative markup for fixed variables (via name space entry);
 wenzelm parents: 
42488diff
changeset | 159 | val fixes_space = #1 o fixes_of; | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 160 | val binds_of = #binds o rep_data; | 
| 20162 | 161 | val type_occs_of = #type_occs o rep_data; | 
| 24765 | 162 | val maxidx_of = #maxidx o rep_data; | 
| 28625 | 163 | val sorts_of = #sorts o rep_data; | 
| 20162 | 164 | val constraints_of = #constraints o rep_data; | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 165 | |
| 20162 | 166 | val is_declared = Name.is_declared o names_of; | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 167 | |
| 42494 | 168 | val check_name = Long_Name.base_name o Name_Space.full_name Name_Space.default_naming; | 
| 42357 | 169 | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 170 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 171 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 172 | (** declarations **) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 173 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 174 | (* default sorts and types *) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 175 | |
| 20162 | 176 | 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 | 177 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 178 | fun def_type ctxt pattern xi = | 
| 20162 | 179 |   let val {binds, constraints = (types, _), ...} = rep_data ctxt in
 | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 180 | (case Vartab.lookup types xi of | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 181 | NONE => | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 182 | if pattern then NONE | 
| 39290 
44e4d8dfd6bf
load type_infer.ML later -- proper context for Type_Infer.infer_types;
 wenzelm parents: 
38831diff
changeset | 183 | else Vartab.lookup binds xi |> Option.map (Type.mark_polymorphic o #1) | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 184 | | some => some) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 185 | end; | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 186 | |
| 20162 | 187 | val def_sort = Vartab.lookup o #2 o constraints_of; | 
| 188 | ||
| 189 | ||
| 190 | (* names *) | |
| 191 | ||
| 24765 | 192 | fun declare_type_names t = | 
| 29279 | 193 | map_names (fold_types (fold_atyps Term.declare_typ_names) t) #> | 
| 24765 | 194 | map_maxidx (fold_types Term.maxidx_typ t); | 
| 20162 | 195 | |
| 196 | fun declare_names t = | |
| 197 | declare_type_names t #> | |
| 29279 | 198 | map_names (fold_aterms Term.declare_term_frees t) #> | 
| 24765 | 199 | map_maxidx (Term.maxidx_term t); | 
| 20162 | 200 | |
| 201 | ||
| 202 | (* type occurrences *) | |
| 203 | ||
| 24719 
21d1cdab2d8c
declare_constraints: declare (fix) type variables within constraints, but not terms themselves;
 wenzelm parents: 
24694diff
changeset | 204 | 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 | 205 | |
| 22671 | 206 | val decl_type_occs = fold_term_types | 
| 20162 | 207 | (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 | 208 | | _ => decl_type_occsT); | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 209 | |
| 24719 
21d1cdab2d8c
declare_constraints: declare (fix) type variables within constraints, but not terms themselves;
 wenzelm parents: 
24694diff
changeset | 210 | val declare_type_occsT = map_type_occs o fold_types decl_type_occsT; | 
| 22671 | 211 | val declare_type_occs = map_type_occs o decl_type_occs; | 
| 212 | ||
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 213 | |
| 20162 | 214 | (* constraints *) | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 215 | |
| 21355 | 216 | fun constrain_tvar (xi, S) = | 
| 217 | if S = dummyS then Vartab.delete_safe xi else Vartab.update (xi, S); | |
| 218 | ||
| 20162 | 219 | fun declare_constraints t = map_constraints (fn (types, sorts) => | 
| 220 | let | |
| 221 | val types' = fold_aterms | |
| 222 | (fn Free (x, T) => Vartab.update ((x, ~1), T) | |
| 223 | | Var v => Vartab.update v | |
| 224 | | _ => I) t types; | |
| 225 | val sorts' = fold_types (fold_atyps | |
| 21355 | 226 | (fn TFree (x, S) => constrain_tvar ((x, ~1), S) | 
| 227 | | TVar v => constrain_tvar v | |
| 20162 | 228 | | _ => I)) t sorts; | 
| 229 | in (types', sorts') end) | |
| 24719 
21d1cdab2d8c
declare_constraints: declare (fix) type variables within constraints, but not terms themselves;
 wenzelm parents: 
24694diff
changeset | 230 | #> declare_type_occsT t | 
| 22711 | 231 | #> declare_type_names t; | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 232 | |
| 20162 | 233 | |
| 234 | (* common declarations *) | |
| 235 | ||
| 236 | fun declare_internal t = | |
| 237 | declare_names t #> | |
| 28625 | 238 | declare_type_occs t #> | 
| 239 | map_sorts (Sorts.insert_term t); | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 240 | |
| 20162 | 241 | fun declare_term t = | 
| 242 | declare_internal t #> | |
| 243 | declare_constraints t; | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 244 | |
| 27280 | 245 | val declare_typ = declare_term o Logic.mk_type; | 
| 246 | ||
| 20303 | 247 | val declare_prf = Proofterm.fold_proof_terms declare_internal (declare_internal o Logic.mk_type); | 
| 248 | ||
| 22691 | 249 | val declare_thm = Thm.fold_terms declare_internal; | 
| 42360 | 250 | fun global_thm_context th = declare_thm th (Proof_Context.init_global (Thm.theory_of_thm th)); | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 251 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 252 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 253 | (* renaming term/type frees *) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 254 | |
| 20162 | 255 | fun variant_frees ctxt ts frees = | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 256 | let | 
| 20162 | 257 | val names = names_of (fold declare_names ts ctxt); | 
| 20084 | 258 | val xs = fst (Name.variants (map #1 frees) names); | 
| 259 | in xs ~~ map snd frees end; | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 260 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 261 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 262 | |
| 20303 | 263 | (** term bindings **) | 
| 264 | ||
| 30756 | 265 | fun bind_term (xi, NONE) = map_binds (Vartab.delete_safe xi) | 
| 266 | | bind_term ((x, i), SOME t) = | |
| 20303 | 267 | let | 
| 25051 | 268 | val u = Term.close_schematic_term t; | 
| 269 | val U = Term.fastype_of u; | |
| 270 | in declare_term u #> map_binds (Vartab.update ((x, i), (U, u))) end; | |
| 20303 | 271 | |
| 272 | fun expand_binds ctxt = | |
| 273 | let | |
| 274 | val binds = binds_of ctxt; | |
| 21799 | 275 | val get = fn Var (xi, _) => Vartab.lookup binds xi | _ => NONE; | 
| 276 | in Envir.beta_norm o Envir.expand_term get end; | |
| 20303 | 277 | |
| 278 | ||
| 279 | ||
| 25325 | 280 | (** consts **) | 
| 25316 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 281 | |
| 25325 | 282 | val lookup_const = Symtab.lookup o #consts o rep_data; | 
| 283 | 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 | 284 | |
| 25325 | 285 | val declare_fixed = map_consts o Symtab.delete_safe; | 
| 286 | 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 | 287 | |
| 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 288 | |
| 
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
 wenzelm parents: 
25051diff
changeset | 289 | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 290 | (** fixes **) | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 291 | |
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 292 | (* specialized name space *) | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 293 | |
| 42493 
01430341fc79
more informative markup for fixed variables (via name space entry);
 wenzelm parents: 
42488diff
changeset | 294 | fun is_fixed ctxt x = can (Name_Space.the_entry (fixes_space ctxt)) x; | 
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 295 | fun newly_fixed inner outer x = is_fixed inner x andalso not (is_fixed outer x); | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 296 | |
| 42493 
01430341fc79
more informative markup for fixed variables (via name space entry);
 wenzelm parents: 
42488diff
changeset | 297 | val fixed_ord = Name_Space.entry_ord o fixes_space; | 
| 
01430341fc79
more informative markup for fixed variables (via name space entry);
 wenzelm parents: 
42488diff
changeset | 298 | val intern_fixed = Name_Space.intern o fixes_space; | 
| 
01430341fc79
more informative markup for fixed variables (via name space entry);
 wenzelm parents: 
42488diff
changeset | 299 | val markup_fixed = Name_Space.markup o fixes_space; | 
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 300 | |
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 301 | fun lookup_fixed ctxt x = | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 302 | let val x' = intern_fixed ctxt x | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 303 | in if is_fixed ctxt x' then SOME x' else NONE end; | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 304 | |
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 305 | fun revert_fixed ctxt x = | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 306 | (case Symtab.lookup (#2 (fixes_of ctxt)) x of | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 307 | SOME x' => if intern_fixed ctxt x' = x then x' else x | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 308 | | NONE => x); | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 309 | |
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 310 | fun dest_fixes ctxt = | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 311 | let val (space, tab) = fixes_of ctxt | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 312 | in sort (Name_Space.entry_ord space o pairself #2) (map swap (Symtab.dest tab)) end; | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 313 | |
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 314 | |
| 42482 | 315 | (* collect variables *) | 
| 316 | ||
| 317 | fun add_free_names ctxt = | |
| 318 | fold_aterms (fn Free (x, _) => not (is_fixed ctxt x) ? insert (op =) x | _ => I); | |
| 319 | ||
| 320 | fun add_frees ctxt = | |
| 321 | fold_aterms (fn Free (x, T) => not (is_fixed ctxt x) ? insert (op =) (x, T) | _ => I); | |
| 322 | ||
| 323 | fun add_fixed_names ctxt = | |
| 324 | fold_aterms (fn Free (x, _) => is_fixed ctxt x ? insert (op =) x | _ => I); | |
| 325 | ||
| 326 | fun add_fixed ctxt = | |
| 327 | fold_aterms (fn Free (x, T) => is_fixed ctxt x ? insert (op =) (x, T) | _ => I); | |
| 328 | ||
| 329 | ||
| 330 | (* declarations *) | |
| 331 | ||
| 20084 | 332 | local | 
| 333 | ||
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 334 | fun err_dups dups = | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 335 |   error ("Duplicate fixed variable(s): " ^ commas (map Binding.print dups));
 | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 336 | |
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 337 | fun new_fixed ((x, x'), pos) ctxt = | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 338 | if is_some (lookup_fixed ctxt x') then err_dups [Binding.make (x, pos)] | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 339 | else | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 340 | ctxt | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 341 | |> map_fixes | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 342 | (Name_Space.define ctxt true Name_Space.default_naming (Binding.make (x', pos), x) #> snd #>> | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 343 | Name_Space.alias Name_Space.default_naming (Binding.make (x, pos)) x') | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 344 | |> declare_fixed x | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 345 | |> declare_constraints (Syntax.free x'); | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 346 | |
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 347 | fun new_fixes names' xs xs' ps = | 
| 20102 | 348 | map_names (K names') #> | 
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 349 | fold new_fixed ((xs ~~ xs') ~~ ps) #> | 
| 20084 | 350 | pair xs'; | 
| 351 | ||
| 352 | in | |
| 353 | ||
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 354 | fun add_fixes_binding bs ctxt = | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 355 | let | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 356 | val _ = | 
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 357 | (case filter (can Name.dest_skolem o Binding.name_of) bs of | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 358 | [] => () | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 359 |       | bads => error ("Illegal internal Skolem constant(s): " ^ commas (map Binding.print bads)));
 | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 360 | val _ = | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 361 | (case duplicates (op = o pairself Binding.name_of) bs of | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 362 | [] => () | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 363 | | dups => err_dups dups); | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 364 | |
| 42494 | 365 | val xs = map check_name bs; | 
| 20102 | 366 | val names = names_of ctxt; | 
| 20084 | 367 | val (xs', names') = | 
| 20149 | 368 | if is_body ctxt then Name.variants xs names |>> map Name.skolem | 
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 369 | else (xs, fold Name.declare xs names); | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 370 | in ctxt |> new_fixes names' xs xs' (map Binding.pos_of bs) end; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 371 | |
| 20797 
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
 wenzelm parents: 
20579diff
changeset | 372 | fun variant_fixes raw_xs ctxt = | 
| 20084 | 373 | let | 
| 20102 | 374 | val names = names_of ctxt; | 
| 26714 
4773b832f1b1
variant_fixes: preserve internal state, mark skolem only for body mode;
 wenzelm parents: 
25573diff
changeset | 375 | val xs = map (fn x => Name.clean x |> can Name.dest_internal x ? Name.internal) raw_xs; | 
| 
4773b832f1b1
variant_fixes: preserve internal state, mark skolem only for body mode;
 wenzelm parents: 
25573diff
changeset | 376 | val (xs', names') = Name.variants xs names |>> (is_body ctxt ? map Name.skolem); | 
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 377 | in ctxt |> new_fixes names' xs xs' (replicate (length xs) Position.none) end; | 
| 20084 | 378 | |
| 379 | end; | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 380 | |
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 381 | val add_fixes = add_fixes_binding o map Binding.name; | 
| 20251 | 382 | |
| 383 | fun add_fixes_direct xs ctxt = ctxt | |
| 384 | |> set_body false | |
| 385 | |> (snd o add_fixes xs) | |
| 386 | |> restore_body ctxt; | |
| 387 | ||
| 42482 | 388 | fun auto_fixes t ctxt = ctxt | 
| 389 | |> not (is_body ctxt) ? add_fixes_direct (rev (add_free_names ctxt t [])) | |
| 20251 | 390 | |> declare_term t; | 
| 391 | ||
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 392 | fun invent_types Ss ctxt = | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 393 | let | 
| 24848 | 394 | val tfrees = Name.invents (names_of ctxt) Name.aT (length Ss) ~~ Ss; | 
| 20162 | 395 | 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 | 396 | in (tfrees, ctxt') end; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 397 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 398 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 399 | |
| 22671 | 400 | (** 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 | 401 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 402 | fun export_inst inner outer = | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 403 | let | 
| 20162 | 404 | val declared_outer = is_declared outer; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 405 | |
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 406 | val gen_fixes = | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 407 | Symtab.fold (fn (y, _) => not (is_fixed outer y) ? cons y) | 
| 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42482diff
changeset | 408 | (#2 (fixes_of inner)) []; | 
| 20162 | 409 | val still_fixed = not o member (op =) gen_fixes; | 
| 22671 | 410 | |
| 411 | val type_occs_inner = type_occs_of inner; | |
| 412 | fun gen_fixesT ts = | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 413 | Symtab.fold (fn (a, xs) => | 
| 20162 | 414 | if declared_outer a orelse exists still_fixed xs | 
| 22671 | 415 | 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 | 416 | in (gen_fixesT, gen_fixes) end; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 417 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 418 | 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 | 419 | |
| 22671 | 420 | fun exportT_terms inner outer = | 
| 421 | let val mk_tfrees = exportT_inst inner outer in | |
| 422 | fn ts => ts |> map | |
| 31977 | 423 | (Term_Subst.generalize (mk_tfrees ts, []) | 
| 22671 | 424 | (fold (Term.fold_types Term.maxidx_typ) ts ~1 + 1)) | 
| 425 | end; | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 426 | |
| 22671 | 427 | fun export_terms inner outer = | 
| 428 | let val (mk_tfrees, tfrees) = export_inst inner outer in | |
| 429 | fn ts => ts |> map | |
| 31977 | 430 | (Term_Subst.generalize (mk_tfrees ts, tfrees) | 
| 22671 | 431 | (fold Term.maxidx_term ts ~1 + 1)) | 
| 432 | end; | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 433 | |
| 20303 | 434 | fun export_prf inner outer prf = | 
| 435 | let | |
| 22671 | 436 | val (mk_tfrees, frees) = export_inst (declare_prf prf inner) outer; | 
| 437 | val tfrees = mk_tfrees []; | |
| 20303 | 438 | val idx = Proofterm.maxidx_proof prf ~1 + 1; | 
| 36620 
e6bb250402b5
simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
 wenzelm parents: 
36610diff
changeset | 439 | val gen_term = Term_Subst.generalize_same (tfrees, frees) idx; | 
| 
e6bb250402b5
simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
 wenzelm parents: 
36610diff
changeset | 440 | val gen_typ = Term_Subst.generalizeT_same tfrees idx; | 
| 
e6bb250402b5
simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
 wenzelm parents: 
36610diff
changeset | 441 | in Same.commit (Proofterm.map_proof_terms_same gen_term gen_typ) prf end; | 
| 20303 | 442 | |
| 22671 | 443 | |
| 444 | fun gen_export (mk_tfrees, frees) ths = | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 445 | let | 
| 22671 | 446 | val tfrees = mk_tfrees (map Thm.full_prop_of ths); | 
| 447 | val maxidx = fold Thm.maxidx_thm ths ~1; | |
| 448 | 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 | 449 | |
| 22671 | 450 | fun exportT inner outer = gen_export (exportT_inst inner outer, []); | 
| 451 | 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 | 452 | |
| 21522 | 453 | fun export_morphism inner outer = | 
| 454 | let | |
| 455 | val fact = export inner outer; | |
| 456 | val term = singleton (export_terms inner outer); | |
| 457 | val typ = Logic.type_map term; | |
| 29605 | 458 |   in Morphism.morphism {binding = I, typ = typ, term = term, fact = fact} end;
 | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 459 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 460 | |
| 24765 | 461 | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 462 | (** import -- fix schematic type/term variables **) | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 463 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 464 | fun importT_inst ts ctxt = | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 465 | let | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 466 | val tvars = rev (fold Term.add_tvars ts []); | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 467 | val (tfrees, ctxt') = invent_types (map #2 tvars) ctxt; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 468 | in (tvars ~~ map TFree tfrees, ctxt') end; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 469 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 470 | fun import_inst is_open ts ctxt = | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 471 | let | 
| 26714 
4773b832f1b1
variant_fixes: preserve internal state, mark skolem only for body mode;
 wenzelm parents: 
25573diff
changeset | 472 | val ren = Name.clean #> (if is_open then I else Name.internal); | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 473 | val (instT, ctxt') = importT_inst ts ctxt; | 
| 31977 | 474 | val vars = map (apsnd (Term_Subst.instantiateT instT)) (rev (fold Term.add_vars ts [])); | 
| 20797 
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
 wenzelm parents: 
20579diff
changeset | 475 | 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 | 476 | val inst = vars ~~ map Free (xs ~~ map #2 vars); | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 477 | in ((instT, inst), ctxt'') end; | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 478 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 479 | fun importT_terms ts ctxt = | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 480 | let val (instT, ctxt') = importT_inst ts ctxt | 
| 31977 | 481 | in (map (Term_Subst.instantiate (instT, [])) ts, ctxt') end; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 482 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 483 | fun import_terms is_open ts ctxt = | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 484 | let val (inst, ctxt') = import_inst is_open ts ctxt | 
| 31977 | 485 | in (map (Term_Subst.instantiate inst) ts, ctxt') end; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 486 | |
| 31794 
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
 wenzelm parents: 
30756diff
changeset | 487 | fun importT ths ctxt = | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 488 | let | 
| 42360 | 489 | val thy = Proof_Context.theory_of ctxt; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 490 | val (instT, ctxt') = importT_inst (map Thm.full_prop_of ths) ctxt; | 
| 32280 
4fb3f426052a
Variable.importT/import: return full instantiations, tuned;
 wenzelm parents: 
32199diff
changeset | 491 | val insts' as (instT', _) = Thm.certify_inst thy (instT, []); | 
| 
4fb3f426052a
Variable.importT/import: return full instantiations, tuned;
 wenzelm parents: 
32199diff
changeset | 492 | val ths' = map (Thm.instantiate insts') ths; | 
| 
4fb3f426052a
Variable.importT/import: return full instantiations, tuned;
 wenzelm parents: 
32199diff
changeset | 493 | in ((instT', ths'), ctxt') end; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 494 | |
| 20303 | 495 | fun import_prf is_open prf ctxt = | 
| 496 | let | |
| 497 | val ts = rev (Proofterm.fold_proof_terms cons (cons o Logic.mk_type) prf []); | |
| 498 | val (insts, ctxt') = import_inst is_open ts ctxt; | |
| 499 | in (Proofterm.instantiate insts prf, ctxt') end; | |
| 500 | ||
| 31794 
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
 wenzelm parents: 
30756diff
changeset | 501 | fun import is_open ths ctxt = | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 502 | let | 
| 42360 | 503 | val thy = Proof_Context.theory_of ctxt; | 
| 32280 
4fb3f426052a
Variable.importT/import: return full instantiations, tuned;
 wenzelm parents: 
32199diff
changeset | 504 | val (insts, ctxt') = import_inst is_open (map Thm.full_prop_of ths) ctxt; | 
| 
4fb3f426052a
Variable.importT/import: return full instantiations, tuned;
 wenzelm parents: 
32199diff
changeset | 505 | val insts' = Thm.certify_inst thy insts; | 
| 
4fb3f426052a
Variable.importT/import: return full instantiations, tuned;
 wenzelm parents: 
32199diff
changeset | 506 | val ths' = map (Thm.instantiate insts') ths; | 
| 
4fb3f426052a
Variable.importT/import: return full instantiations, tuned;
 wenzelm parents: 
32199diff
changeset | 507 | in ((insts', ths'), ctxt') end; | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 508 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 509 | |
| 19926 | 510 | (* import/export *) | 
| 511 | ||
| 21287 | 512 | fun gen_trade imp exp f ctxt ths = | 
| 20220 | 513 | let val ((_, ths'), ctxt') = imp ths ctxt | 
| 21287 | 514 | in exp ctxt' ctxt (f ctxt' ths') end; | 
| 19926 | 515 | |
| 31794 
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
 wenzelm parents: 
30756diff
changeset | 516 | val tradeT = gen_trade importT exportT; | 
| 
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
 wenzelm parents: 
30756diff
changeset | 517 | val trade = gen_trade (import true) export; | 
| 19926 | 518 | |
| 519 | ||
| 42495 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 520 | (* focus on outermost parameters: !!x y z. B *) | 
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 521 | |
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 522 | fun focus_params t ctxt = | 
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 523 | let | 
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 524 | val (xs, Ts) = | 
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 525 | split_list (Term.variant_frees t (Term.strip_all_vars t)); (*as they are printed :-*) | 
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 526 | val (xs', ctxt') = variant_fixes xs ctxt; | 
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 527 | val ps = xs' ~~ Ts; | 
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 528 | val ctxt'' = ctxt' |> fold (declare_constraints o Free) ps; | 
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 529 | in ((xs, ps), ctxt'') end; | 
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 530 | |
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 531 | fun focus t ctxt = | 
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 532 | let | 
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 533 | val ((xs, ps), ctxt') = focus_params t ctxt; | 
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 534 | val t' = Term.subst_bounds (rev (map Free ps), Term.strip_all_body t); | 
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 535 | in (((xs ~~ ps), t'), ctxt') end; | 
| 20149 | 536 | |
| 537 | fun forall_elim_prop t prop = | |
| 20579 | 538 | Thm.beta_conversion false (Thm.capply (Thm.dest_arg prop) t) | 
| 539 | |> Thm.cprop_of |> Thm.dest_arg; | |
| 20149 | 540 | |
| 42495 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 541 | fun focus_cterm goal ctxt = | 
| 20149 | 542 | let | 
| 543 | val cert = Thm.cterm_of (Thm.theory_of_cterm goal); | |
| 42495 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 544 | val ((xs, ps), ctxt') = focus_params (Thm.term_of goal) ctxt; | 
| 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 545 | val ps' = map (cert o Free) ps; | 
| 20220 | 546 | val goal' = fold forall_elim_prop ps' goal; | 
| 42495 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 547 | in ((xs ~~ ps', goal'), ctxt') end; | 
| 20149 | 548 | |
| 20303 | 549 | fun focus_subgoal i st = | 
| 550 | let | |
| 22691 | 551 | val all_vars = Thm.fold_terms Term.add_vars st []; | 
| 20303 | 552 | val no_binds = map (fn (xi, _) => (xi, NONE)) all_vars; | 
| 553 | in | |
| 30756 | 554 | fold bind_term no_binds #> | 
| 20303 | 555 | fold (declare_constraints o Var) all_vars #> | 
| 42495 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42494diff
changeset | 556 | focus_cterm (Thm.cprem_of st i) | 
| 20303 | 557 | end; | 
| 558 | ||
| 559 | ||
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 560 | |
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 561 | (** implicit polymorphism **) | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 562 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 563 | (* warn_extra_tfrees *) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 564 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 565 | fun warn_extra_tfrees ctxt1 ctxt2 = | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 566 | let | 
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 567 | fun occs_typ a = Term.exists_subtype (fn TFree (b, _) => a = b | _ => false); | 
| 20162 | 568 | fun occs_free a x = | 
| 569 | (case def_type ctxt1 false (x, ~1) of | |
| 570 | SOME T => if occs_typ a T then I else cons (a, x) | |
| 571 | | NONE => cons (a, x)); | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 572 | |
| 20162 | 573 | val occs1 = type_occs_of ctxt1; | 
| 574 | val occs2 = type_occs_of ctxt2; | |
| 19911 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 575 | val extras = Symtab.fold (fn (a, xs) => | 
| 
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
 wenzelm parents: 
19899diff
changeset | 576 | 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 | 577 | val tfrees = map #1 extras |> sort_distinct string_ord; | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 578 | val frees = map #2 extras |> sort_distinct string_ord; | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 579 | in | 
| 38831 
4933a3dfd745
more careful treatment of context visibility flag wrt. spurious warnings;
 wenzelm parents: 
37145diff
changeset | 580 | if null extras orelse not (Context_Position.is_visible ctxt2) then () | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 581 |     else warning ("Introduced fixed type variable(s): " ^ commas tfrees ^ " in " ^
 | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 582 | space_implode " or " (map quote frees)) | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 583 | end; | 
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 584 | |
| 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 585 | |
| 20149 | 586 | (* polymorphic terms *) | 
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 587 | |
| 24694 | 588 | fun polymorphic_types ctxt ts = | 
| 20003 | 589 | let | 
| 590 | val ctxt' = fold declare_term ts ctxt; | |
| 20162 | 591 | val occs = type_occs_of ctxt; | 
| 592 | val occs' = type_occs_of ctxt'; | |
| 593 | val types = Symtab.fold (fn (a, _) => if Symtab.defined occs a then I else cons a) occs' []; | |
| 24765 | 594 | val idx = maxidx_of ctxt' + 1; | 
| 24694 | 595 | val Ts' = (fold o fold_types o fold_atyps) | 
| 596 | (fn T as TFree _ => | |
| 31977 | 597 | (case Term_Subst.generalizeT types idx T of TVar v => insert (op =) v | _ => I) | 
| 24694 | 598 | | _ => I) ts []; | 
| 31977 | 599 | val ts' = map (Term_Subst.generalize (types, []) idx) ts; | 
| 24694 | 600 | in (rev Ts', ts') end; | 
| 601 | ||
| 602 | fun polymorphic ctxt ts = snd (polymorphic_types ctxt ts); | |
| 20003 | 603 | |
| 19899 
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
 wenzelm parents: diff
changeset | 604 | end; |