author | wenzelm |
Sat, 30 Nov 2024 22:33:21 +0100 | |
changeset 81519 | cdc43c0fdbfc |
parent 81518 | f012cbfd96d0 |
child 81521 | 1bfad73ab115 |
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 names_of: Proof.context -> Name.context |
10 |
val binds_of: Proof.context -> (typ * term) Vartab.table |
|
24765 | 11 |
val maxidx_of: Proof.context -> int |
20303 | 12 |
val constraints_of: Proof.context -> typ Vartab.table * sort Vartab.table |
13 |
val is_declared: Proof.context -> string -> bool |
|
42494 | 14 |
val check_name: binding -> string |
20303 | 15 |
val default_type: Proof.context -> string -> typ option |
16 |
val def_type: Proof.context -> bool -> indexname -> typ option |
|
17 |
val def_sort: Proof.context -> indexname -> sort option |
|
59646 | 18 |
val declare_maxidx: int -> Proof.context -> Proof.context |
27280 | 19 |
val declare_names: term -> Proof.context -> Proof.context |
81519 | 20 |
val variant_names: Proof.context -> (string * 'a) list -> (string * 'a) list |
20303 | 21 |
val declare_constraints: term -> Proof.context -> Proof.context |
62012
12d3edd62932
more precise context -- potentially relevant for Eisbach dummy thm;
wenzelm
parents:
61949
diff
changeset
|
22 |
val declare_internal: term -> Proof.context -> Proof.context |
20303 | 23 |
val declare_term: term -> Proof.context -> Proof.context |
27280 | 24 |
val declare_typ: typ -> Proof.context -> Proof.context |
20303 | 25 |
val declare_prf: Proofterm.proof -> Proof.context -> Proof.context |
26 |
val declare_thm: thm -> Proof.context -> Proof.context |
|
60401 | 27 |
val bind_term: indexname * term -> Proof.context -> Proof.context |
28 |
val unbind_term: indexname -> Proof.context -> Proof.context |
|
29 |
val maybe_bind_term: indexname * term option -> Proof.context -> Proof.context |
|
20303 | 30 |
val expand_binds: Proof.context -> term -> term |
25325 | 31 |
val lookup_const: Proof.context -> string -> string option |
25316
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
wenzelm
parents:
25051
diff
changeset
|
32 |
val is_const: Proof.context -> string -> bool |
25325 | 33 |
val declare_const: string * string -> Proof.context -> Proof.context |
55635 | 34 |
val next_bound: string * typ -> Proof.context -> term * Proof.context |
55014
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
35 |
val revert_bounds: Proof.context -> term -> term |
59798 | 36 |
val is_body: Proof.context -> bool |
37 |
val set_body: bool -> Proof.context -> Proof.context |
|
38 |
val restore_body: Proof.context -> Proof.context -> Proof.context |
|
59790 | 39 |
val improper_fixes: Proof.context -> Proof.context |
40 |
val restore_proper_fixes: Proof.context -> Proof.context -> Proof.context |
|
41 |
val is_improper: Proof.context -> string -> bool |
|
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
42 |
val is_fixed: Proof.context -> string -> bool |
59846 | 43 |
val is_newly_fixed: Proof.context -> Proof.context -> string -> bool |
70586 | 44 |
val fixed_ord: Proof.context -> string ord |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
45 |
val intern_fixed: Proof.context -> string -> string |
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
46 |
val lookup_fixed: Proof.context -> string -> string option |
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
47 |
val revert_fixed: Proof.context -> string -> string |
62987
dc8a8a7559e7
highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents:
62955
diff
changeset
|
48 |
val markup_fixed: Proof.context -> string -> Markup.T |
dc8a8a7559e7
highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents:
62955
diff
changeset
|
49 |
val markup: Proof.context -> string -> Markup.T |
dc8a8a7559e7
highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents:
62955
diff
changeset
|
50 |
val markup_entity_def: Proof.context -> string -> Markup.T |
dc8a8a7559e7
highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents:
62955
diff
changeset
|
51 |
val dest_fixes: Proof.context -> (string * string) list |
42482 | 52 |
val add_fixed_names: Proof.context -> term -> string list -> string list |
53 |
val add_fixed: Proof.context -> term -> (string * typ) list -> (string * typ) list |
|
59846 | 54 |
val add_newly_fixed: Proof.context -> Proof.context -> |
55 |
term -> (string * typ) list -> (string * typ) list |
|
42482 | 56 |
val add_free_names: Proof.context -> term -> string list -> string list |
57 |
val add_frees: Proof.context -> term -> (string * typ) list -> (string * typ) list |
|
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
58 |
val add_fixes_binding: binding list -> Proof.context -> string list * Proof.context |
20303 | 59 |
val add_fixes: string list -> Proof.context -> string list * Proof.context |
60 |
val add_fixes_direct: string list -> Proof.context -> Proof.context |
|
70733 | 61 |
val add_fixes_implicit: term -> Proof.context -> Proof.context |
59796 | 62 |
val fix_dummy_patterns: term -> Proof.context -> term * Proof.context |
20797
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents:
20579
diff
changeset
|
63 |
val variant_fixes: string list -> Proof.context -> string list * Proof.context |
60822 | 64 |
val gen_all: Proof.context -> thm -> thm |
20303 | 65 |
val export_terms: Proof.context -> Proof.context -> term list -> term list |
66 |
val exportT_terms: Proof.context -> Proof.context -> term list -> term list |
|
67 |
val exportT: Proof.context -> Proof.context -> thm list -> thm list |
|
68 |
val export_prf: Proof.context -> Proof.context -> Proofterm.proof -> Proofterm.proof |
|
69 |
val export: Proof.context -> Proof.context -> thm list -> thm list |
|
21522 | 70 |
val export_morphism: Proof.context -> Proof.context -> morphism |
59796 | 71 |
val invent_types: sort list -> Proof.context -> (string * sort) list * Proof.context |
74266 | 72 |
val importT_inst: term list -> Proof.context -> typ TVars.table * Proof.context |
20303 | 73 |
val import_inst: bool -> term list -> Proof.context -> |
74266 | 74 |
(typ TVars.table * term Vars.table) * Proof.context |
81222
b61abd1e5027
notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents:
78062
diff
changeset
|
75 |
val import_inst_revert: typ TVars.table * term Vars.table -> typ TFrees.table * term Frees.table |
20303 | 76 |
val importT_terms: term list -> Proof.context -> term list * Proof.context |
77 |
val import_terms: bool -> term list -> Proof.context -> term list * Proof.context |
|
74266 | 78 |
val importT: thm list -> Proof.context -> (ctyp TVars.table * thm list) * Proof.context |
20303 | 79 |
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:
30756
diff
changeset
|
80 |
val import: bool -> thm list -> Proof.context -> |
74266 | 81 |
((ctyp TVars.table * cterm Vars.table) * thm list) * Proof.context |
70304 | 82 |
val import_vars: Proof.context -> thm -> thm |
21287 | 83 |
val tradeT: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list |
84 |
val trade: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list |
|
74525
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
85 |
val dest_abs: term -> Proof.context -> ((string * typ) * term) * Proof.context |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
86 |
val dest_abs_cterm: cterm -> Proof.context -> (cterm * cterm) * Proof.context |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
87 |
val dest_all: term -> Proof.context -> ((string * typ) * term) * Proof.context |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
88 |
val dest_all_cterm: cterm -> Proof.context -> (cterm * cterm) * Proof.context |
60707
e96b7be56d44
SUBPROOF and Subgoal.FOCUS combinators use anonymous quasi-bound variables (like the Simplifier);
wenzelm
parents:
60695
diff
changeset
|
89 |
val is_bound_focus: Proof.context -> bool |
e96b7be56d44
SUBPROOF and Subgoal.FOCUS combinators use anonymous quasi-bound variables (like the Simplifier);
wenzelm
parents:
60695
diff
changeset
|
90 |
val set_bound_focus: bool -> Proof.context -> Proof.context |
e96b7be56d44
SUBPROOF and Subgoal.FOCUS combinators use anonymous quasi-bound variables (like the Simplifier);
wenzelm
parents:
60695
diff
changeset
|
91 |
val restore_bound_focus: Proof.context -> Proof.context -> Proof.context |
60695
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
92 |
val focus_params: binding list option -> term -> Proof.context -> |
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
93 |
(string list * (string * typ) list) * Proof.context |
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
94 |
val focus: binding list option -> term -> Proof.context -> |
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
95 |
((string * (string * typ)) list * term) * Proof.context |
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
96 |
val focus_cterm: binding list option -> cterm -> Proof.context -> |
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
97 |
((string * cterm) list * cterm) * Proof.context |
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
98 |
val focus_subgoal: binding list option -> int -> thm -> Proof.context -> |
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
99 |
((string * cterm) list * cterm) * Proof.context |
20303 | 100 |
val warn_extra_tfrees: Proof.context -> Proof.context -> unit |
24694 | 101 |
val polymorphic_types: Proof.context -> term list -> (indexname * sort) list * term list |
20303 | 102 |
val polymorphic: Proof.context -> term list -> term list |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
103 |
end; |
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 |
structure Variable: VARIABLE = |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
106 |
struct |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
107 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
108 |
(** local context data **) |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
109 |
|
59790 | 110 |
type fixes = (string * bool) Name_Space.table; |
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49688
diff
changeset
|
111 |
val empty_fixes: fixes = Name_Space.empty_table Markup.fixedN; |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
112 |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
113 |
datatype data = Data of |
59798 | 114 |
{names: Name.context, (*type/term variable names*) |
25325 | 115 |
consts: string Symtab.table, (*consts within the local scope*) |
55014
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
116 |
bounds: int * ((string * typ) * string) list, (*next index, internal name, type, external name*) |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
117 |
fixes: fixes, (*term fixes -- global name space, intern ~> extern*) |
20102 | 118 |
binds: (typ * term) Vartab.table, (*term bindings*) |
20162 | 119 |
type_occs: string list Symtab.table, (*type variables -- possibly within term variables*) |
24765 | 120 |
maxidx: int, (*maximum var index*) |
20162 | 121 |
constraints: |
20102 | 122 |
typ Vartab.table * (*type constraints*) |
20162 | 123 |
sort Vartab.table}; (*default sorts*) |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
124 |
|
61508 | 125 |
fun make_data (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) = |
59798 | 126 |
Data {names = names, consts = consts, bounds = bounds, fixes = fixes, binds = binds, |
61508 | 127 |
type_occs = type_occs, maxidx = maxidx, constraints = constraints}; |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
128 |
|
59150 | 129 |
val empty_data = |
59798 | 130 |
make_data (Name.context, Symtab.empty, (0, []), empty_fixes, Vartab.empty, |
61508 | 131 |
Symtab.empty, ~1, (Vartab.empty, Vartab.empty)); |
59150 | 132 |
|
33519 | 133 |
structure Data = Proof_Data |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
134 |
( |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
135 |
type T = data; |
59150 | 136 |
fun init _ = empty_data; |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
137 |
); |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
138 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
139 |
fun map_data f = |
61508 | 140 |
Data.map (fn Data {names, consts, bounds, fixes, binds, type_occs, maxidx, constraints} => |
141 |
make_data (f (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints))); |
|
25316
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
wenzelm
parents:
25051
diff
changeset
|
142 |
|
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
wenzelm
parents:
25051
diff
changeset
|
143 |
fun map_names f = |
61508 | 144 |
map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => |
145 |
(f names, consts, bounds, fixes, binds, type_occs, maxidx, constraints)); |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
146 |
|
25325 | 147 |
fun map_consts f = |
61508 | 148 |
map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => |
149 |
(names, f consts, bounds, fixes, binds, type_occs, maxidx, constraints)); |
|
55014
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
150 |
|
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
151 |
fun map_bounds f = |
61508 | 152 |
map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => |
153 |
(names, consts, f bounds, fixes, binds, type_occs, maxidx, constraints)); |
|
20162 | 154 |
|
25316
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
wenzelm
parents:
25051
diff
changeset
|
155 |
fun map_fixes f = |
61508 | 156 |
map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => |
157 |
(names, consts, bounds, f fixes, binds, type_occs, maxidx, constraints)); |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
158 |
|
25316
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
wenzelm
parents:
25051
diff
changeset
|
159 |
fun map_binds f = |
61508 | 160 |
map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => |
161 |
(names, consts, bounds, fixes, f binds, type_occs, maxidx, constraints)); |
|
24765 | 162 |
|
25316
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
wenzelm
parents:
25051
diff
changeset
|
163 |
fun map_type_occs f = |
61508 | 164 |
map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => |
165 |
(names, consts, bounds, fixes, binds, f type_occs, maxidx, constraints)); |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
166 |
|
25316
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
wenzelm
parents:
25051
diff
changeset
|
167 |
fun map_maxidx f = |
61508 | 168 |
map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => |
169 |
(names, consts, bounds, fixes, binds, type_occs, f maxidx, constraints)); |
|
20102 | 170 |
|
25316
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
wenzelm
parents:
25051
diff
changeset
|
171 |
fun map_constraints f = |
61508 | 172 |
map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => |
173 |
(names, consts, bounds, fixes, binds, type_occs, maxidx, f constraints)); |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
174 |
|
45650 | 175 |
fun rep_data ctxt = Data.get ctxt |> (fn Data rep => rep); |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
176 |
|
20102 | 177 |
val names_of = #names o rep_data; |
178 |
val fixes_of = #fixes o rep_data; |
|
56025 | 179 |
val fixes_space = Name_Space.space_of_table o fixes_of; |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
180 |
val binds_of = #binds o rep_data; |
20162 | 181 |
val type_occs_of = #type_occs o rep_data; |
24765 | 182 |
val maxidx_of = #maxidx o rep_data; |
20162 | 183 |
val constraints_of = #constraints o rep_data; |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
184 |
|
20162 | 185 |
val is_declared = Name.is_declared o names_of; |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
186 |
|
47021
f35f654f297d
clarified Binding.name_of vs Name_Space.base_name vs Variable.check_name (see also 9bd8d4addd6e, 3305f573294e);
wenzelm
parents:
47005
diff
changeset
|
187 |
val check_name = Name_Space.base_name o tap Binding.check; |
42357 | 188 |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
189 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
190 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
191 |
(** declarations **) |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
192 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
193 |
(* default sorts and types *) |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
194 |
|
20162 | 195 |
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
|
196 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
197 |
fun def_type ctxt pattern xi = |
20162 | 198 |
let val {binds, constraints = (types, _), ...} = rep_data ctxt in |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
199 |
(case Vartab.lookup types xi of |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
200 |
NONE => |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
201 |
if pattern then NONE |
39290
44e4d8dfd6bf
load type_infer.ML later -- proper context for Type_Infer.infer_types;
wenzelm
parents:
38831
diff
changeset
|
202 |
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
|
203 |
| some => some) |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
204 |
end; |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
205 |
|
20162 | 206 |
val def_sort = Vartab.lookup o #2 o constraints_of; |
207 |
||
208 |
||
59646 | 209 |
(* maxidx *) |
210 |
||
211 |
val declare_maxidx = map_maxidx o Integer.max; |
|
212 |
||
213 |
||
81519 | 214 |
(* type/term names *) |
20162 | 215 |
|
24765 | 216 |
fun declare_type_names t = |
81505 | 217 |
map_names (Term.declare_tfree_names t) #> |
24765 | 218 |
map_maxidx (fold_types Term.maxidx_typ t); |
20162 | 219 |
|
220 |
fun declare_names t = |
|
221 |
declare_type_names t #> |
|
81505 | 222 |
map_names (Term.declare_free_names t) #> |
24765 | 223 |
map_maxidx (Term.maxidx_term t); |
20162 | 224 |
|
81519 | 225 |
fun variant_names ctxt xs = Name.variant_names (names_of ctxt) xs; |
226 |
||
20162 | 227 |
|
228 |
(* type occurrences *) |
|
229 |
||
24719
21d1cdab2d8c
declare_constraints: declare (fix) type variables within constraints, but not terms themselves;
wenzelm
parents:
24694
diff
changeset
|
230 |
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:
24694
diff
changeset
|
231 |
|
22671 | 232 |
val decl_type_occs = fold_term_types |
20162 | 233 |
(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:
24694
diff
changeset
|
234 |
| _ => decl_type_occsT); |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
235 |
|
24719
21d1cdab2d8c
declare_constraints: declare (fix) type variables within constraints, but not terms themselves;
wenzelm
parents:
24694
diff
changeset
|
236 |
val declare_type_occsT = map_type_occs o fold_types decl_type_occsT; |
22671 | 237 |
val declare_type_occs = map_type_occs o decl_type_occs; |
238 |
||
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
239 |
|
20162 | 240 |
(* constraints *) |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
241 |
|
49685
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
242 |
fun constrain_tvar (xi, raw_S) = |
49688 | 243 |
let val S = #2 (Term_Position.decode_positionS raw_S) |
49685
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
244 |
in if S = dummyS then Vartab.delete_safe xi else Vartab.update (xi, S) end; |
21355 | 245 |
|
20162 | 246 |
fun declare_constraints t = map_constraints (fn (types, sorts) => |
247 |
let |
|
248 |
val types' = fold_aterms |
|
62955
2fd4378caca2
back to dummy constraints (amending dd2914250ca7): important for Syntax_Phases.get_free/is_declared;
wenzelm
parents:
62771
diff
changeset
|
249 |
(fn Free (x, T) => Vartab.update ((x, ~1), T) |
2fd4378caca2
back to dummy constraints (amending dd2914250ca7): important for Syntax_Phases.get_free/is_declared;
wenzelm
parents:
62771
diff
changeset
|
250 |
| Var v => Vartab.update v |
20162 | 251 |
| _ => I) t types; |
45426 | 252 |
val sorts' = (fold_types o fold_atyps) |
21355 | 253 |
(fn TFree (x, S) => constrain_tvar ((x, ~1), S) |
254 |
| TVar v => constrain_tvar v |
|
45426 | 255 |
| _ => I) t sorts; |
20162 | 256 |
in (types', sorts') end) |
24719
21d1cdab2d8c
declare_constraints: declare (fix) type variables within constraints, but not terms themselves;
wenzelm
parents:
24694
diff
changeset
|
257 |
#> declare_type_occsT t |
22711 | 258 |
#> declare_type_names t; |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
259 |
|
20162 | 260 |
|
261 |
(* common declarations *) |
|
262 |
||
263 |
fun declare_internal t = |
|
264 |
declare_names t #> |
|
28625 | 265 |
declare_type_occs t #> |
61508 | 266 |
Thm.declare_term_sorts t; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
267 |
|
20162 | 268 |
fun declare_term t = |
269 |
declare_internal t #> |
|
270 |
declare_constraints t; |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
271 |
|
27280 | 272 |
val declare_typ = declare_term o Logic.mk_type; |
273 |
||
70843
cc987440d776
more compact XML: separate environment for free variables;
wenzelm
parents:
70733
diff
changeset
|
274 |
val declare_prf = |
cc987440d776
more compact XML: separate environment for free variables;
wenzelm
parents:
70733
diff
changeset
|
275 |
Proofterm.fold_proof_terms_types declare_internal (declare_internal o Logic.mk_type); |
20303 | 276 |
|
74241
eb265f54e3ce
more efficient operations: traverse hyps only when required;
wenzelm
parents:
74240
diff
changeset
|
277 |
val declare_thm = Thm.fold_terms {hyps = true} declare_internal; |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
278 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
279 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
280 |
|
20303 | 281 |
(** term bindings **) |
282 |
||
60401 | 283 |
fun bind_term ((x, i), t) = |
284 |
let |
|
285 |
val u = Term.close_schematic_term t; |
|
286 |
val U = Term.fastype_of u; |
|
287 |
in declare_term u #> map_binds (Vartab.update ((x, i), (U, u))) end; |
|
288 |
||
289 |
val unbind_term = map_binds o Vartab.delete_safe; |
|
290 |
||
291 |
fun maybe_bind_term (xi, SOME t) = bind_term (xi, t) |
|
292 |
| maybe_bind_term (xi, NONE) = unbind_term xi; |
|
20303 | 293 |
|
294 |
fun expand_binds ctxt = |
|
295 |
let |
|
296 |
val binds = binds_of ctxt; |
|
21799 | 297 |
val get = fn Var (xi, _) => Vartab.lookup binds xi | _ => NONE; |
298 |
in Envir.beta_norm o Envir.expand_term get end; |
|
20303 | 299 |
|
300 |
||
301 |
||
25325 | 302 |
(** consts **) |
25316
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
wenzelm
parents:
25051
diff
changeset
|
303 |
|
25325 | 304 |
val lookup_const = Symtab.lookup o #consts o rep_data; |
305 |
val is_const = is_some oo lookup_const; |
|
25316
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
wenzelm
parents:
25051
diff
changeset
|
306 |
|
25325 | 307 |
val declare_fixed = map_consts o Symtab.delete_safe; |
308 |
val declare_const = map_consts o Symtab.update; |
|
25316
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
wenzelm
parents:
25051
diff
changeset
|
309 |
|
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
wenzelm
parents:
25051
diff
changeset
|
310 |
|
17c183417f93
added is_const/declare_const for local scope of fixes/consts;
wenzelm
parents:
25051
diff
changeset
|
311 |
|
55014
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
312 |
(** bounds **) |
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
313 |
|
74532 | 314 |
fun inc_bound (a, T) ctxt = |
55014
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
315 |
let |
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
316 |
val b = Name.bound (#1 (#bounds (rep_data ctxt))); |
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
317 |
val ctxt' = ctxt |> map_bounds (fn (next, bounds) => (next + 1, ((b, T), a) :: bounds)); |
55635 | 318 |
in (Free (b, T), ctxt') end; |
55014
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
319 |
|
74532 | 320 |
fun next_bound a ctxt = |
321 |
let val (x as Free (b, _), ctxt') = inc_bound a ctxt |
|
322 |
in if Name.is_declared (names_of ctxt') b then inc_bound a ctxt' else (x, ctxt') end; |
|
323 |
||
55014
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
324 |
fun revert_bounds ctxt t = |
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
325 |
(case #2 (#bounds (rep_data ctxt)) of |
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
326 |
[] => t |
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
327 |
| bounds => |
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
328 |
let |
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
329 |
val names = Term.declare_term_names t (names_of ctxt); |
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
330 |
val xs = rev (#1 (fold_map Name.variant (rev (map #2 bounds)) names)); |
74539
f07761ee5a7f
more robust Variable.revert_bounds (see also b12f2cef3ee5);
wenzelm
parents:
74532
diff
changeset
|
331 |
fun substs (((b, T), _), x') = |
f07761ee5a7f
more robust Variable.revert_bounds (see also b12f2cef3ee5);
wenzelm
parents:
74532
diff
changeset
|
332 |
let fun subst U = (Free (b, U), Syntax_Trans.mark_bound_abs (x', U)) |
f07761ee5a7f
more robust Variable.revert_bounds (see also b12f2cef3ee5);
wenzelm
parents:
74532
diff
changeset
|
333 |
in [subst T, subst (Type_Annotation.ignore_type T)] end; |
f07761ee5a7f
more robust Variable.revert_bounds (see also b12f2cef3ee5);
wenzelm
parents:
74532
diff
changeset
|
334 |
in Term.subst_atomic (maps substs (bounds ~~ xs)) t end); |
55014
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
335 |
|
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
336 |
|
a93f496f6c30
general notion of auxiliary bounds within context;
wenzelm
parents:
54740
diff
changeset
|
337 |
|
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
338 |
(** fixes **) |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
339 |
|
59798 | 340 |
(* inner body mode *) |
341 |
||
69575 | 342 |
val inner_body = Config.declare_bool ("inner_body", \<^here>) (K false); |
69576 | 343 |
val is_body = Config.apply inner_body; |
59798 | 344 |
val set_body = Config.put inner_body; |
69576 | 345 |
val restore_body = set_body o is_body; |
59798 | 346 |
|
347 |
||
59790 | 348 |
(* proper mode *) |
349 |
||
69575 | 350 |
val proper_fixes = Config.declare_bool ("proper_fixes", \<^here>) (K true); |
59790 | 351 |
val improper_fixes = Config.put proper_fixes false; |
69576 | 352 |
val restore_proper_fixes = Config.put proper_fixes o Config.apply proper_fixes; |
59790 | 353 |
|
354 |
fun is_improper ctxt x = |
|
59884 | 355 |
(case Name_Space.lookup (fixes_of ctxt) x of |
356 |
SOME (_, proper) => not proper |
|
59790 | 357 |
| NONE => false); |
358 |
||
359 |
||
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
360 |
(* specialized name space *) |
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
361 |
|
59883 | 362 |
val is_fixed = Name_Space.defined o fixes_of; |
59846 | 363 |
fun is_newly_fixed inner outer = is_fixed inner andf (not o is_fixed outer); |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
364 |
|
42493
01430341fc79
more informative markup for fixed variables (via name space entry);
wenzelm
parents:
42488
diff
changeset
|
365 |
val fixed_ord = Name_Space.entry_ord o fixes_space; |
01430341fc79
more informative markup for fixed variables (via name space entry);
wenzelm
parents:
42488
diff
changeset
|
366 |
val intern_fixed = Name_Space.intern o fixes_space; |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
367 |
|
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
368 |
fun lookup_fixed ctxt x = |
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
369 |
let val x' = intern_fixed ctxt x |
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
370 |
in if is_fixed ctxt x' then SOME x' else NONE end; |
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
371 |
|
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
372 |
fun revert_fixed ctxt x = |
59884 | 373 |
(case Name_Space.lookup (fixes_of ctxt) x of |
374 |
SOME (x', _) => if intern_fixed ctxt x' = x then x' else x |
|
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
375 |
| NONE => x); |
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
376 |
|
45472 | 377 |
fun markup_fixed ctxt x = |
378 |
Name_Space.markup (fixes_space ctxt) x |
|
379 |
|> Markup.name (revert_fixed ctxt x); |
|
380 |
||
62987
dc8a8a7559e7
highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents:
62955
diff
changeset
|
381 |
fun markup ctxt x = |
dc8a8a7559e7
highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents:
62955
diff
changeset
|
382 |
if is_improper ctxt x then Markup.improper |
dc8a8a7559e7
highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents:
62955
diff
changeset
|
383 |
else if Name.is_skolem x then Markup.skolem |
dc8a8a7559e7
highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents:
62955
diff
changeset
|
384 |
else Markup.free; |
dc8a8a7559e7
highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents:
62955
diff
changeset
|
385 |
|
dc8a8a7559e7
highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents:
62955
diff
changeset
|
386 |
val markup_entity_def = Name_Space.markup_def o fixes_space; |
dc8a8a7559e7
highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents:
62955
diff
changeset
|
387 |
|
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
388 |
fun dest_fixes ctxt = |
59790 | 389 |
Name_Space.fold_table (fn (x, (y, _)) => cons (y, x)) (fixes_of ctxt) [] |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58668
diff
changeset
|
390 |
|> sort (Name_Space.entry_ord (fixes_space ctxt) o apply2 #2); |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
391 |
|
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
392 |
|
42482 | 393 |
(* collect variables *) |
394 |
||
395 |
fun add_free_names ctxt = |
|
396 |
fold_aterms (fn Free (x, _) => not (is_fixed ctxt x) ? insert (op =) x | _ => I); |
|
397 |
||
398 |
fun add_frees ctxt = |
|
399 |
fold_aterms (fn Free (x, T) => not (is_fixed ctxt x) ? insert (op =) (x, T) | _ => I); |
|
400 |
||
401 |
fun add_fixed_names ctxt = |
|
402 |
fold_aterms (fn Free (x, _) => is_fixed ctxt x ? insert (op =) x | _ => I); |
|
403 |
||
404 |
fun add_fixed ctxt = |
|
405 |
fold_aterms (fn Free (x, T) => is_fixed ctxt x ? insert (op =) (x, T) | _ => I); |
|
406 |
||
59846 | 407 |
fun add_newly_fixed ctxt' ctxt = |
408 |
fold_aterms (fn Free (x, T) => is_newly_fixed ctxt' ctxt x ? insert (op =) (x, T) | _ => I); |
|
409 |
||
42482 | 410 |
|
411 |
(* declarations *) |
|
412 |
||
20084 | 413 |
local |
414 |
||
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
415 |
fun err_dups dups = |
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
416 |
error ("Duplicate fixed variable(s): " ^ commas (map Binding.print dups)); |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
417 |
|
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
418 |
fun new_fixed ((x, x'), pos) ctxt = |
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
419 |
if is_some (lookup_fixed ctxt x') then err_dups [Binding.make (x, pos)] |
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
420 |
else |
59790 | 421 |
let |
422 |
val proper = Config.get ctxt proper_fixes; |
|
61949 | 423 |
val context = Context.Proof ctxt |
424 |
|> Name_Space.map_naming (K Name_Space.global_naming) |
|
425 |
|> Context_Position.set_visible_generic false; |
|
59790 | 426 |
in |
47005
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents:
46869
diff
changeset
|
427 |
ctxt |
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents:
46869
diff
changeset
|
428 |
|> map_fixes |
59790 | 429 |
(Name_Space.define context true (Binding.make (x', pos), (x, proper)) #> snd #> |
77979
a12c48fbf10f
back to more elementary concept of aliases as adhoc change of accesses, but now with "suppress" information (see also 31ea5c1f874d);
wenzelm
parents:
77970
diff
changeset
|
430 |
x <> "" ? Name_Space.alias_table Name_Space.global_naming (Binding.make (x, pos)) x') |
47005
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents:
46869
diff
changeset
|
431 |
|> declare_fixed x |
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents:
46869
diff
changeset
|
432 |
|> declare_constraints (Syntax.free x') |
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents:
46869
diff
changeset
|
433 |
end; |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
434 |
|
71316 | 435 |
fun new_fixes names' args = |
20102 | 436 |
map_names (K names') #> |
71316 | 437 |
fold new_fixed args #> |
438 |
pair (map (#2 o #1) args); |
|
20084 | 439 |
|
81518 | 440 |
fun variant ctxt raw_xs = |
441 |
let |
|
442 |
val names = names_of ctxt; |
|
443 |
val xs = map (fn x => Name.clean x |> Name.is_internal x ? Name.internal) raw_xs; |
|
444 |
val (xs', names') = fold_map Name.variant xs names |>> (is_body ctxt ? map Name.skolem); |
|
445 |
in (names', xs ~~ xs') end; |
|
446 |
||
20084 | 447 |
in |
448 |
||
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
449 |
fun add_fixes_binding bs ctxt = |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
450 |
let |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
451 |
val _ = |
55948 | 452 |
(case filter (Name.is_skolem o Binding.name_of) bs of |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
453 |
[] => () |
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
454 |
| bads => error ("Illegal internal Skolem constant(s): " ^ commas (map Binding.print bads))); |
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
455 |
val _ = |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58668
diff
changeset
|
456 |
(case duplicates (op = o apply2 Binding.name_of) bs of |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
457 |
[] => () |
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
458 |
| dups => err_dups dups); |
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
459 |
|
42494 | 460 |
val xs = map check_name bs; |
20102 | 461 |
val names = names_of ctxt; |
20084 | 462 |
val (xs', names') = |
43326
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents:
42495
diff
changeset
|
463 |
if is_body ctxt then fold_map Name.variant xs names |>> map Name.skolem |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
464 |
else (xs, fold Name.declare xs names); |
71316 | 465 |
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:
19899
diff
changeset
|
466 |
|
71317
e58bc223f46c
more robust bound_fixes: external name does appear in "fixes" name space and could clash with direct fixes (e.g. via Subgoal.FOCUS within a global context);
wenzelm
parents:
71316
diff
changeset
|
467 |
fun variant_fixes xs ctxt = |
81518 | 468 |
let val (names', vs) = variant ctxt xs; |
71317
e58bc223f46c
more robust bound_fixes: external name does appear in "fixes" name space and could clash with direct fixes (e.g. via Subgoal.FOCUS within a global context);
wenzelm
parents:
71316
diff
changeset
|
469 |
in ctxt |> new_fixes names' (map (rpair Position.none) vs) end; |
e58bc223f46c
more robust bound_fixes: external name does appear in "fixes" name space and could clash with direct fixes (e.g. via Subgoal.FOCUS within a global context);
wenzelm
parents:
71316
diff
changeset
|
470 |
|
e58bc223f46c
more robust bound_fixes: external name does appear in "fixes" name space and could clash with direct fixes (e.g. via Subgoal.FOCUS within a global context);
wenzelm
parents:
71316
diff
changeset
|
471 |
fun bound_fixes xs ctxt = |
e58bc223f46c
more robust bound_fixes: external name does appear in "fixes" name space and could clash with direct fixes (e.g. via Subgoal.FOCUS within a global context);
wenzelm
parents:
71316
diff
changeset
|
472 |
let |
81518 | 473 |
val (names', vs) = variant ctxt (map #1 xs); |
71317
e58bc223f46c
more robust bound_fixes: external name does appear in "fixes" name space and could clash with direct fixes (e.g. via Subgoal.FOCUS within a global context);
wenzelm
parents:
71316
diff
changeset
|
474 |
val (ys, ctxt') = fold_map next_bound (map2 (fn (x', _) => fn (_, T) => (x', T)) vs xs) ctxt; |
e58bc223f46c
more robust bound_fixes: external name does appear in "fixes" name space and could clash with direct fixes (e.g. via Subgoal.FOCUS within a global context);
wenzelm
parents:
71316
diff
changeset
|
475 |
val fixes = map2 (fn (x, _) => fn Free (y, _) => ((x, y), Position.none)) vs ys; |
e58bc223f46c
more robust bound_fixes: external name does appear in "fixes" name space and could clash with direct fixes (e.g. via Subgoal.FOCUS within a global context);
wenzelm
parents:
71316
diff
changeset
|
476 |
in ctxt' |> new_fixes names' fixes end; |
20084 | 477 |
|
478 |
end; |
|
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
479 |
|
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
480 |
val add_fixes = add_fixes_binding o map Binding.name; |
20251 | 481 |
|
482 |
fun add_fixes_direct xs ctxt = ctxt |
|
483 |
|> set_body false |
|
484 |
|> (snd o add_fixes xs) |
|
485 |
|> restore_body ctxt; |
|
486 |
||
70733 | 487 |
fun add_fixes_implicit t ctxt = ctxt |
488 |
|> not (is_body ctxt) ? add_fixes_direct (rev (add_free_names ctxt t [])); |
|
489 |
||
59796 | 490 |
|
491 |
(* dummy patterns *) |
|
492 |
||
493 |
fun fix_dummy_patterns (Const ("Pure.dummy_pattern", T)) ctxt = |
|
61923
a10cc7fb1841
clarified context policy to allow multiple dummies;
wenzelm
parents:
61508
diff
changeset
|
494 |
let val ([x], ctxt') = ctxt |> set_body true |> add_fixes [Name.uu_] ||> restore_body ctxt |
59796 | 495 |
in (Free (x, T), ctxt') end |
496 |
| fix_dummy_patterns (Abs (x, T, b)) ctxt = |
|
497 |
let val (b', ctxt') = fix_dummy_patterns b ctxt |
|
498 |
in (Abs (x, T, b'), ctxt') end |
|
499 |
| fix_dummy_patterns (t $ u) ctxt = |
|
500 |
let |
|
501 |
val (t', ctxt') = fix_dummy_patterns t ctxt; |
|
502 |
val (u', ctxt'') = fix_dummy_patterns u ctxt'; |
|
503 |
in (t' $ u', ctxt'') end |
|
504 |
| fix_dummy_patterns a ctxt = (a, ctxt); |
|
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
505 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
506 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
507 |
|
22671 | 508 |
(** export -- generalize type/term variables (beware of closure sizes) **) |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
509 |
|
60823 | 510 |
fun gen_all ctxt th = |
511 |
let |
|
512 |
val i = Thm.maxidx_thm th (maxidx_of ctxt) + 1; |
|
513 |
fun gen (x, T) = Thm.forall_elim (Thm.cterm_of ctxt (Var ((x, i), T))); |
|
514 |
in fold gen (Drule.outer_params (Thm.prop_of th)) th end; |
|
60822 | 515 |
|
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
516 |
fun export_inst inner outer = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
517 |
let |
20162 | 518 |
val declared_outer = is_declared outer; |
59846 | 519 |
val still_fixed = not o is_newly_fixed inner outer; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
520 |
|
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42482
diff
changeset
|
521 |
val gen_fixes = |
74266 | 522 |
Names.build (fixes_of inner |> Name_Space.fold_table (fn (y, _) => |
523 |
not (is_fixed outer y) ? Names.add_set y)); |
|
22671 | 524 |
|
525 |
val type_occs_inner = type_occs_of inner; |
|
526 |
fun gen_fixesT ts = |
|
74266 | 527 |
Names.build (fold decl_type_occs ts type_occs_inner |> Symtab.fold (fn (a, xs) => |
20162 | 528 |
if declared_outer a orelse exists still_fixed xs |
74266 | 529 |
then I else Names.add_set a)); |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
530 |
in (gen_fixesT, gen_fixes) end; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
531 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
532 |
fun exportT_inst inner outer = #1 (export_inst inner outer); |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
533 |
|
22671 | 534 |
fun exportT_terms inner outer = |
59645
f89464e9ffa0
clarified Variable.export: observe maxidx of target context;
wenzelm
parents:
59623
diff
changeset
|
535 |
let |
f89464e9ffa0
clarified Variable.export: observe maxidx of target context;
wenzelm
parents:
59623
diff
changeset
|
536 |
val mk_tfrees = exportT_inst inner outer; |
f89464e9ffa0
clarified Variable.export: observe maxidx of target context;
wenzelm
parents:
59623
diff
changeset
|
537 |
val maxidx = maxidx_of outer; |
f89464e9ffa0
clarified Variable.export: observe maxidx of target context;
wenzelm
parents:
59623
diff
changeset
|
538 |
in |
22671 | 539 |
fn ts => ts |> map |
74266 | 540 |
(Term_Subst.generalize (mk_tfrees ts, Names.empty) |
59645
f89464e9ffa0
clarified Variable.export: observe maxidx of target context;
wenzelm
parents:
59623
diff
changeset
|
541 |
(fold (Term.fold_types Term.maxidx_typ) ts maxidx + 1)) |
22671 | 542 |
end; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
543 |
|
22671 | 544 |
fun export_terms inner outer = |
59645
f89464e9ffa0
clarified Variable.export: observe maxidx of target context;
wenzelm
parents:
59623
diff
changeset
|
545 |
let |
f89464e9ffa0
clarified Variable.export: observe maxidx of target context;
wenzelm
parents:
59623
diff
changeset
|
546 |
val (mk_tfrees, tfrees) = export_inst inner outer; |
f89464e9ffa0
clarified Variable.export: observe maxidx of target context;
wenzelm
parents:
59623
diff
changeset
|
547 |
val maxidx = maxidx_of outer; |
f89464e9ffa0
clarified Variable.export: observe maxidx of target context;
wenzelm
parents:
59623
diff
changeset
|
548 |
in |
22671 | 549 |
fn ts => ts |> map |
74201 | 550 |
(Term_Subst.generalize (mk_tfrees ts, tfrees) |
59645
f89464e9ffa0
clarified Variable.export: observe maxidx of target context;
wenzelm
parents:
59623
diff
changeset
|
551 |
(fold Term.maxidx_term ts maxidx + 1)) |
22671 | 552 |
end; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
553 |
|
20303 | 554 |
fun export_prf inner outer prf = |
555 |
let |
|
22671 | 556 |
val (mk_tfrees, frees) = export_inst (declare_prf prf inner) outer; |
74201 | 557 |
val tfrees = mk_tfrees []; |
59645
f89464e9ffa0
clarified Variable.export: observe maxidx of target context;
wenzelm
parents:
59623
diff
changeset
|
558 |
val maxidx = maxidx_of outer; |
f89464e9ffa0
clarified Variable.export: observe maxidx of target context;
wenzelm
parents:
59623
diff
changeset
|
559 |
val idx = Proofterm.maxidx_proof prf maxidx + 1; |
74201 | 560 |
val gen_term = Term_Subst.generalize_same (tfrees, frees) idx; |
36620
e6bb250402b5
simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
wenzelm
parents:
36610
diff
changeset
|
561 |
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:
36610
diff
changeset
|
562 |
in Same.commit (Proofterm.map_proof_terms_same gen_term gen_typ) prf end; |
20303 | 563 |
|
22671 | 564 |
|
59645
f89464e9ffa0
clarified Variable.export: observe maxidx of target context;
wenzelm
parents:
59623
diff
changeset
|
565 |
fun gen_export (mk_tfrees, frees) maxidx ths = |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
566 |
let |
22671 | 567 |
val tfrees = mk_tfrees (map Thm.full_prop_of ths); |
59645
f89464e9ffa0
clarified Variable.export: observe maxidx of target context;
wenzelm
parents:
59623
diff
changeset
|
568 |
val idx = fold Thm.maxidx_thm ths maxidx + 1; |
74201 | 569 |
in map (Thm.generalize (tfrees, frees) idx) ths end; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
570 |
|
74266 | 571 |
fun exportT inner outer = gen_export (exportT_inst inner outer, Names.empty) (maxidx_of outer); |
59645
f89464e9ffa0
clarified Variable.export: observe maxidx of target context;
wenzelm
parents:
59623
diff
changeset
|
572 |
fun export inner outer = gen_export (export_inst inner outer) (maxidx_of outer); |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
573 |
|
21522 | 574 |
fun export_morphism inner outer = |
575 |
let |
|
576 |
val fact = export inner outer; |
|
577 |
val term = singleton (export_terms inner outer); |
|
578 |
val typ = Logic.type_map term; |
|
54740 | 579 |
in |
78062
edb195122938
support for context within morphism (for background theory);
wenzelm
parents:
78050
diff
changeset
|
580 |
Morphism.morphism "Variable.export" |
edb195122938
support for context within morphism (for background theory);
wenzelm
parents:
78050
diff
changeset
|
581 |
{binding = [], typ = [K typ], term = [K term], fact = [K fact]} |
54740 | 582 |
end; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
583 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
584 |
|
24765 | 585 |
|
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
586 |
(** import -- fix schematic type/term variables **) |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
587 |
|
59796 | 588 |
fun invent_types Ss ctxt = |
589 |
let |
|
81512
c1aa8a61ee65
tuned: more direct use of Name.context operations;
wenzelm
parents:
81505
diff
changeset
|
590 |
val tfrees = Name.invent_types (names_of ctxt) Ss; |
59796 | 591 |
val ctxt' = fold (declare_constraints o Logic.mk_type o TFree) tfrees ctxt; |
592 |
in (tfrees, ctxt') end; |
|
593 |
||
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
594 |
fun importT_inst ts ctxt = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
595 |
let |
74281 | 596 |
val tvars = TVars.build (fold TVars.add_tvars ts) |> TVars.list_set; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
597 |
val (tfrees, ctxt') = invent_types (map #2 tvars) ctxt; |
74266 | 598 |
val instT = TVars.build (fold2 (fn a => fn b => TVars.add (a, TFree b)) tvars tfrees); |
74220
c49134ee16c1
more scalable data structure (but: rarely used many arguments);
wenzelm
parents:
74201
diff
changeset
|
599 |
in (instT, ctxt') end; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
600 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
601 |
fun import_inst is_open ts ctxt = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
602 |
let |
26714
4773b832f1b1
variant_fixes: preserve internal state, mark skolem only for body mode;
wenzelm
parents:
25573
diff
changeset
|
603 |
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:
19899
diff
changeset
|
604 |
val (instT, ctxt') = importT_inst ts ctxt; |
74281 | 605 |
val vars = |
606 |
Vars.build (fold Vars.add_vars ts) |> Vars.list_set |
|
607 |
|> map (apsnd (Term_Subst.instantiateT instT)); |
|
74220
c49134ee16c1
more scalable data structure (but: rarely used many arguments);
wenzelm
parents:
74201
diff
changeset
|
608 |
val (ys, ctxt'') = variant_fixes (map (ren o #1 o #1) vars) ctxt'; |
74266 | 609 |
val inst = Vars.build (fold2 (fn (x, T) => fn y => Vars.add ((x, T), Free (y, T))) vars ys); |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
610 |
in ((instT, inst), ctxt'') end; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
611 |
|
81222
b61abd1e5027
notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents:
78062
diff
changeset
|
612 |
fun import_inst_revert (instT, inst) = |
b61abd1e5027
notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents:
78062
diff
changeset
|
613 |
let |
b61abd1e5027
notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents:
78062
diff
changeset
|
614 |
val instT' = TFrees.build (instT |> TVars.fold (fn (v, TFree w) => TFrees.add (w, TVar v))); |
b61abd1e5027
notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents:
78062
diff
changeset
|
615 |
val instantiateT' = Term_Subst.instantiateT_frees instT'; |
b61abd1e5027
notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents:
78062
diff
changeset
|
616 |
val inst' = |
b61abd1e5027
notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents:
78062
diff
changeset
|
617 |
Frees.build (inst |> Vars.fold (fn ((xi, T), Free (y, U)) => |
b61abd1e5027
notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents:
78062
diff
changeset
|
618 |
Frees.add ((y, instantiateT' U), Var (xi, instantiateT' T)))); |
b61abd1e5027
notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents:
78062
diff
changeset
|
619 |
in (instT', inst') end; |
b61abd1e5027
notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents:
78062
diff
changeset
|
620 |
|
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
621 |
fun importT_terms ts ctxt = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
622 |
let val (instT, ctxt') = importT_inst ts ctxt |
74266 | 623 |
in (map (Term_Subst.instantiate (instT, Vars.empty)) ts, ctxt') end; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
624 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
625 |
fun import_terms is_open ts ctxt = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
626 |
let val (inst, ctxt') = import_inst is_open ts ctxt |
31977 | 627 |
in (map (Term_Subst.instantiate inst) ts, ctxt') end; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
628 |
|
31794
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents:
30756
diff
changeset
|
629 |
fun importT ths ctxt = |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
630 |
let |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
631 |
val (instT, ctxt') = importT_inst (map Thm.full_prop_of ths) ctxt; |
74266 | 632 |
val instT' = TVars.map (K (Thm.ctyp_of ctxt')) instT; |
74282 | 633 |
val ths' = map (Thm.instantiate (instT', Vars.empty)) ths; |
32280
4fb3f426052a
Variable.importT/import: return full instantiations, tuned;
wenzelm
parents:
32199
diff
changeset
|
634 |
in ((instT', ths'), ctxt') end; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
635 |
|
20303 | 636 |
fun import_prf is_open prf ctxt = |
637 |
let |
|
70843
cc987440d776
more compact XML: separate environment for free variables;
wenzelm
parents:
70733
diff
changeset
|
638 |
val ts = rev (Proofterm.fold_proof_terms_types cons (cons o Logic.mk_type) prf []); |
20303 | 639 |
val (insts, ctxt') = import_inst is_open ts ctxt; |
640 |
in (Proofterm.instantiate insts prf, ctxt') end; |
|
641 |
||
31794
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents:
30756
diff
changeset
|
642 |
fun import is_open ths ctxt = |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
643 |
let |
60805 | 644 |
val ((instT, inst), ctxt') = import_inst is_open (map Thm.full_prop_of ths) ctxt; |
74266 | 645 |
val instT' = TVars.map (K (Thm.ctyp_of ctxt')) instT; |
646 |
val inst' = Vars.map (K (Thm.cterm_of ctxt')) inst; |
|
74282 | 647 |
val ths' = map (Thm.instantiate (instT', inst')) ths; |
74220
c49134ee16c1
more scalable data structure (but: rarely used many arguments);
wenzelm
parents:
74201
diff
changeset
|
648 |
in (((instT', inst'), ths'), ctxt') end; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
649 |
|
70304 | 650 |
fun import_vars ctxt th = |
651 |
let val ((_, [th']), _) = ctxt |> set_body false |> import true [th]; |
|
652 |
in th' end; |
|
653 |
||
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
654 |
|
19926 | 655 |
(* import/export *) |
656 |
||
21287 | 657 |
fun gen_trade imp exp f ctxt ths = |
20220 | 658 |
let val ((_, ths'), ctxt') = imp ths ctxt |
21287 | 659 |
in exp ctxt' ctxt (f ctxt' ths') end; |
19926 | 660 |
|
31794
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents:
30756
diff
changeset
|
661 |
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:
30756
diff
changeset
|
662 |
val trade = gen_trade (import true) export; |
19926 | 663 |
|
664 |
||
74525
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
665 |
(* destruct binders *) |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
666 |
|
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
667 |
local |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
668 |
|
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
669 |
fun gen_dest_abs exn dest term_of arg ctxt = |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
670 |
(case term_of arg of |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
671 |
Abs (a, T, _) => |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
672 |
let |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
673 |
val (x, ctxt') = yield_singleton bound_fixes (a, T) ctxt; |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
674 |
val res = dest x arg handle Term.USED_FREE _ => |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
675 |
raise Fail ("Bad context: clash of fresh free for bound: " ^ |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
676 |
Syntax.string_of_term ctxt (Free (x, T)) ^ " vs. " ^ |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
677 |
Syntax.string_of_term ctxt' (Free (x, T))); |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
678 |
in (res, ctxt') end |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
679 |
| _ => raise exn ("dest_abs", [arg])); |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
680 |
|
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
681 |
in |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
682 |
|
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
683 |
val dest_abs = gen_dest_abs TERM Term.dest_abs_fresh I; |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
684 |
val dest_abs_cterm = gen_dest_abs CTERM Thm.dest_abs_fresh Thm.term_of; |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
685 |
|
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
686 |
fun dest_all t ctxt = |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
687 |
(case t of |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
688 |
Const ("Pure.all", _) $ u => dest_abs u ctxt |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
689 |
| _ => raise TERM ("dest_all", [t])); |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
690 |
|
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
691 |
fun dest_all_cterm ct ctxt = |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
692 |
(case Thm.term_of ct of |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
693 |
Const ("Pure.all", _) $ _ => dest_abs_cterm (Thm.dest_arg ct) ctxt |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
694 |
| _ => raise CTERM ("dest_all", [ct])); |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
695 |
|
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
696 |
end; |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
697 |
|
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74282
diff
changeset
|
698 |
|
67721 | 699 |
(* focus on outermost parameters: \<And>x y z. B *) |
42495
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
wenzelm
parents:
42494
diff
changeset
|
700 |
|
69575 | 701 |
val bound_focus = Config.declare_bool ("bound_focus", \<^here>) (K false); |
69576 | 702 |
val is_bound_focus = Config.apply bound_focus; |
60707
e96b7be56d44
SUBPROOF and Subgoal.FOCUS combinators use anonymous quasi-bound variables (like the Simplifier);
wenzelm
parents:
60695
diff
changeset
|
703 |
val set_bound_focus = Config.put bound_focus; |
69576 | 704 |
val restore_bound_focus = set_bound_focus o is_bound_focus; |
60707
e96b7be56d44
SUBPROOF and Subgoal.FOCUS combinators use anonymous quasi-bound variables (like the Simplifier);
wenzelm
parents:
60695
diff
changeset
|
705 |
|
60695
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
706 |
fun focus_params bindings t ctxt = |
42495
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
wenzelm
parents:
42494
diff
changeset
|
707 |
let |
60707
e96b7be56d44
SUBPROOF and Subgoal.FOCUS combinators use anonymous quasi-bound variables (like the Simplifier);
wenzelm
parents:
60695
diff
changeset
|
708 |
val ps = Term.variant_frees t (Term.strip_all_vars t); (*as they are printed :-*) |
e96b7be56d44
SUBPROOF and Subgoal.FOCUS combinators use anonymous quasi-bound variables (like the Simplifier);
wenzelm
parents:
60695
diff
changeset
|
709 |
val (xs, Ts) = split_list ps; |
60695
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
710 |
val (xs', ctxt') = |
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
711 |
(case bindings of |
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
712 |
SOME bs => ctxt |> set_body true |> add_fixes_binding bs ||> restore_body ctxt |
60707
e96b7be56d44
SUBPROOF and Subgoal.FOCUS combinators use anonymous quasi-bound variables (like the Simplifier);
wenzelm
parents:
60695
diff
changeset
|
713 |
| NONE => if is_bound_focus ctxt then bound_fixes ps ctxt else variant_fixes xs ctxt); |
e96b7be56d44
SUBPROOF and Subgoal.FOCUS combinators use anonymous quasi-bound variables (like the Simplifier);
wenzelm
parents:
60695
diff
changeset
|
714 |
val ps' = xs' ~~ Ts; |
e96b7be56d44
SUBPROOF and Subgoal.FOCUS combinators use anonymous quasi-bound variables (like the Simplifier);
wenzelm
parents:
60695
diff
changeset
|
715 |
val ctxt'' = ctxt' |> fold (declare_constraints o Free) ps'; |
e96b7be56d44
SUBPROOF and Subgoal.FOCUS combinators use anonymous quasi-bound variables (like the Simplifier);
wenzelm
parents:
60695
diff
changeset
|
716 |
in ((xs, ps'), ctxt'') end; |
42495
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
wenzelm
parents:
42494
diff
changeset
|
717 |
|
60695
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
718 |
fun focus bindings t ctxt = |
42495
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
wenzelm
parents:
42494
diff
changeset
|
719 |
let |
60695
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
720 |
val ((xs, ps), ctxt') = focus_params bindings t ctxt; |
42495
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
wenzelm
parents:
42494
diff
changeset
|
721 |
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:
42494
diff
changeset
|
722 |
in (((xs ~~ ps), t'), ctxt') end; |
20149 | 723 |
|
724 |
fun forall_elim_prop t prop = |
|
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
45666
diff
changeset
|
725 |
Thm.beta_conversion false (Thm.apply (Thm.dest_arg prop) t) |
20579 | 726 |
|> Thm.cprop_of |> Thm.dest_arg; |
20149 | 727 |
|
60695
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
728 |
fun focus_cterm bindings goal ctxt = |
20149 | 729 |
let |
60695
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
730 |
val ((xs, ps), ctxt') = focus_params bindings (Thm.term_of goal) ctxt; |
59623 | 731 |
val ps' = map (Thm.cterm_of ctxt' o Free) ps; |
20220 | 732 |
val goal' = fold forall_elim_prop ps' goal; |
42495
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
wenzelm
parents:
42494
diff
changeset
|
733 |
in ((xs ~~ ps', goal'), ctxt') end; |
20149 | 734 |
|
60695
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
735 |
fun focus_subgoal bindings i st = |
74241
eb265f54e3ce
more efficient operations: traverse hyps only when required;
wenzelm
parents:
74240
diff
changeset
|
736 |
let |
74266 | 737 |
val all_vars = Vars.build (Thm.fold_terms {hyps = false} Vars.add_vars st); |
74241
eb265f54e3ce
more efficient operations: traverse hyps only when required;
wenzelm
parents:
74240
diff
changeset
|
738 |
in |
74266 | 739 |
Vars.fold (unbind_term o #1 o #1) all_vars #> |
740 |
Vars.fold (declare_constraints o Var o #1) all_vars #> |
|
60695
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
741 |
focus_cterm bindings (Thm.cprem_of st i) |
20303 | 742 |
end; |
743 |
||
744 |
||
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
745 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
746 |
(** implicit polymorphism **) |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
747 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
748 |
(* warn_extra_tfrees *) |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
749 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
750 |
fun warn_extra_tfrees ctxt1 ctxt2 = |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
751 |
let |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
752 |
fun occs_typ a = Term.exists_subtype (fn TFree (b, _) => a = b | _ => false); |
20162 | 753 |
fun occs_free a x = |
754 |
(case def_type ctxt1 false (x, ~1) of |
|
755 |
SOME T => if occs_typ a T then I else cons (a, x) |
|
756 |
| NONE => cons (a, x)); |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
757 |
|
20162 | 758 |
val occs1 = type_occs_of ctxt1; |
759 |
val occs2 = type_occs_of ctxt2; |
|
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
760 |
val extras = Symtab.fold (fn (a, xs) => |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
761 |
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
|
762 |
val tfrees = map #1 extras |> sort_distinct string_ord; |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
763 |
val frees = map #2 extras |> sort_distinct string_ord; |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
764 |
in |
38831
4933a3dfd745
more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents:
37145
diff
changeset
|
765 |
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
|
766 |
else warning ("Introduced fixed type variable(s): " ^ commas tfrees ^ " in " ^ |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
767 |
space_implode " or " (map quote frees)) |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
768 |
end; |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
769 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
770 |
|
20149 | 771 |
(* polymorphic terms *) |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
772 |
|
24694 | 773 |
fun polymorphic_types ctxt ts = |
20003 | 774 |
let |
775 |
val ctxt' = fold declare_term ts ctxt; |
|
20162 | 776 |
val occs = type_occs_of ctxt; |
777 |
val occs' = type_occs_of ctxt'; |
|
74200
17090e27aae9
more scalable data structure (but: rarely used with > 5 arguments);
wenzelm
parents:
71317
diff
changeset
|
778 |
val types = |
74266 | 779 |
Names.build (occs' |> Symtab.fold (fn (a, _) => |
780 |
if Symtab.defined occs a then I else Names.add_set a)); |
|
24765 | 781 |
val idx = maxidx_of ctxt' + 1; |
24694 | 782 |
val Ts' = (fold o fold_types o fold_atyps) |
783 |
(fn T as TFree _ => |
|
31977 | 784 |
(case Term_Subst.generalizeT types idx T of TVar v => insert (op =) v | _ => I) |
24694 | 785 |
| _ => I) ts []; |
74266 | 786 |
val ts' = map (Term_Subst.generalize (types, Names.empty) idx) ts; |
24694 | 787 |
in (rev Ts', ts') end; |
788 |
||
789 |
fun polymorphic ctxt ts = snd (polymorphic_types ctxt ts); |
|
20003 | 790 |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
791 |
end; |