| author | wenzelm |
| Fri, 03 Sep 2021 18:57:33 +0200 | |
| changeset 74220 | c49134ee16c1 |
| parent 73846 | 9447668d1b77 |
| child 74230 | d637611b41bd |
| permissions | -rw-r--r-- |
| 41959 | 1 |
(* Title: Pure/Isar/generic_target.ML |
| 38309 | 2 |
Author: Makarius |
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
||
5 |
Common target infrastructure. |
|
6 |
*) |
|
7 |
||
8 |
signature GENERIC_TARGET = |
|
9 |
sig |
|
|
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
10 |
(*auxiliary*) |
|
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
11 |
val export_abbrev: Proof.context -> |
|
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
12 |
(term -> term) -> term -> term * ((string * sort) list * (term list * term list)) |
| 62764 | 13 |
val check_mixfix: Proof.context -> binding * (string * sort) list -> mixfix -> mixfix |
|
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
14 |
val check_mixfix_global: binding * bool -> mixfix -> mixfix |
|
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
15 |
|
| 60340 | 16 |
(*background primitives*) |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
17 |
val background_foundation: ((binding * typ) * mixfix) * (binding * term) -> |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
18 |
term list * term list -> local_theory -> (term * thm) * local_theory |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
19 |
val background_declaration: declaration -> local_theory -> local_theory |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
20 |
val background_abbrev: binding * term -> term list -> local_theory -> (term * term) * local_theory |
|
71788
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
21 |
val add_foundation_interpretation: (binding * (term * term list) -> Context.generic -> Context.generic) -> |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
22 |
theory -> theory |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
23 |
|
| 60340 | 24 |
(*nested local theories primitives*) |
| 67652 | 25 |
val standard_facts: local_theory -> Proof.context -> Attrib.fact list -> Attrib.fact list |
26 |
val standard_notes: (int * int -> bool) -> string -> Attrib.fact list -> |
|
| 63269 | 27 |
local_theory -> local_theory |
28 |
val standard_declaration: (int * int -> bool) -> |
|
29 |
(morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory |
|
| 60340 | 30 |
val standard_const: (int * int -> bool) -> Syntax.mode -> (binding * mixfix) * term -> |
31 |
local_theory -> local_theory |
|
| 73845 | 32 |
val local_interpretation: Locale.registration -> |
33 |
local_theory -> local_theory |
|
| 60340 | 34 |
|
35 |
(*lifting target primitives to local theory operations*) |
|
| 40782 | 36 |
val define: (((binding * typ) * mixfix) * (binding * term) -> |
37 |
term list * term list -> local_theory -> (term * thm) * local_theory) -> |
|
|
46992
eeea81b86b70
refined Local_Theory.define vs. Local_Theory.define_internal, which allows to pass alternative name to the foundational axiom -- expecially important for 'instantiation' or 'overloading', which loose name information due to Long_Name.base_name cooking etc.;
wenzelm
parents:
46916
diff
changeset
|
38 |
bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory -> |
| 40782 | 39 |
(term * (string * thm)) * local_theory |
|
47250
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
wenzelm
parents:
47249
diff
changeset
|
40 |
val notes: |
| 67652 | 41 |
(string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory) -> |
42 |
string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory |
|
| 60337 | 43 |
val abbrev: (Syntax.mode -> binding * mixfix -> term -> |
| 57160 | 44 |
term list * term list -> local_theory -> local_theory) -> |
| 60337 | 45 |
Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
46 |
|
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
47 |
(*theory target primitives*) |
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
48 |
val theory_target_foundation: ((binding * typ) * mixfix) * (binding * term) -> |
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
49 |
term list * term list -> local_theory -> (term * thm) * local_theory |
| 67652 | 50 |
val theory_target_notes: string -> Attrib.fact list -> Attrib.fact list -> |
51 |
local_theory -> local_theory |
|
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
52 |
val theory_target_abbrev: Syntax.mode -> (binding * mixfix) -> term -> term list * term list -> |
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
53 |
local_theory -> local_theory |
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
54 |
|
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
55 |
(*theory target operations*) |
|
60344
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60342
diff
changeset
|
56 |
val theory_abbrev: Syntax.mode -> (binding * mixfix) * term -> |
|
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60342
diff
changeset
|
57 |
local_theory -> (term * term) * local_theory |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
58 |
val theory_declaration: declaration -> local_theory -> local_theory |
| 73845 | 59 |
val theory_registration: Locale.registration -> local_theory -> local_theory |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
60 |
|
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
61 |
(*locale target primitives*) |
| 67652 | 62 |
val locale_target_notes: string -> string -> Attrib.fact list -> Attrib.fact list -> |
63 |
local_theory -> local_theory |
|
| 63339 | 64 |
val locale_target_abbrev: string -> Syntax.mode -> |
65 |
(binding * mixfix) -> term -> term list * term list -> local_theory -> local_theory |
|
|
57192
180e955711cf
yet another attempt for terminology: foo_target_bar denotes an operation bar operating solely on the target context of target foo, foo_bar denotes a whole stack of operations to accomplish bar for target foo
haftmann
parents:
57191
diff
changeset
|
66 |
val locale_target_declaration: string -> bool -> declaration -> local_theory -> local_theory |
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
67 |
val locale_target_const: string -> (morphism -> bool) -> Syntax.mode -> |
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
68 |
(binding * mixfix) * term -> local_theory -> local_theory |
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
69 |
|
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
70 |
(*locale operations*) |
|
60344
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60342
diff
changeset
|
71 |
val locale_abbrev: string -> Syntax.mode -> (binding * mixfix) * term -> |
|
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60342
diff
changeset
|
72 |
local_theory -> (term * term) * local_theory |
| 57193 | 73 |
val locale_declaration: string -> {syntax: bool, pervasive: bool} -> declaration ->
|
74 |
local_theory -> local_theory |
|
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
75 |
val locale_const: string -> Syntax.mode -> (binding * mixfix) * term -> |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
76 |
local_theory -> local_theory |
| 73845 | 77 |
val locale_dependency: string -> Locale.registration -> |
78 |
local_theory -> local_theory |
|
| 45353 | 79 |
end |
| 38309 | 80 |
|
81 |
structure Generic_Target: GENERIC_TARGET = |
|
82 |
struct |
|
83 |
||
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
84 |
(** consts **) |
|
38312
9dd57db3c0f2
moved extra_tfrees check for mixfix syntax to Generic_Target
haftmann
parents:
38311
diff
changeset
|
85 |
|
|
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
86 |
fun export_abbrev lthy preprocess rhs = |
|
60342
df9b153d866b
separate function to compute exported abbreviation
haftmann
parents:
60341
diff
changeset
|
87 |
let |
|
df9b153d866b
separate function to compute exported abbreviation
haftmann
parents:
60341
diff
changeset
|
88 |
val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy); |
|
df9b153d866b
separate function to compute exported abbreviation
haftmann
parents:
60341
diff
changeset
|
89 |
|
|
df9b153d866b
separate function to compute exported abbreviation
haftmann
parents:
60341
diff
changeset
|
90 |
val rhs' = rhs |
|
df9b153d866b
separate function to compute exported abbreviation
haftmann
parents:
60341
diff
changeset
|
91 |
|> Assumption.export_term lthy (Local_Theory.target_of lthy) |
|
df9b153d866b
separate function to compute exported abbreviation
haftmann
parents:
60341
diff
changeset
|
92 |
|> preprocess; |
| 63339 | 93 |
val term_params = |
94 |
map Free (sort (Variable.fixed_ord lthy o apply2 #1) (Variable.add_fixed lthy rhs' [])); |
|
|
60342
df9b153d866b
separate function to compute exported abbreviation
haftmann
parents:
60341
diff
changeset
|
95 |
val u = fold_rev lambda term_params rhs'; |
|
df9b153d866b
separate function to compute exported abbreviation
haftmann
parents:
60341
diff
changeset
|
96 |
val global_rhs = singleton (Variable.polymorphic thy_ctxt) u; |
|
df9b153d866b
separate function to compute exported abbreviation
haftmann
parents:
60341
diff
changeset
|
97 |
|
|
df9b153d866b
separate function to compute exported abbreviation
haftmann
parents:
60341
diff
changeset
|
98 |
val extra_tfrees = |
|
df9b153d866b
separate function to compute exported abbreviation
haftmann
parents:
60341
diff
changeset
|
99 |
subtract (op =) (Term.add_tfreesT (Term.fastype_of u) []) (Term.add_tfrees u []); |
|
df9b153d866b
separate function to compute exported abbreviation
haftmann
parents:
60341
diff
changeset
|
100 |
val type_params = map (Logic.mk_type o TFree) extra_tfrees; |
|
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
101 |
in (global_rhs, (extra_tfrees, (type_params, term_params))) end; |
|
60342
df9b153d866b
separate function to compute exported abbreviation
haftmann
parents:
60341
diff
changeset
|
102 |
|
|
38312
9dd57db3c0f2
moved extra_tfrees check for mixfix syntax to Generic_Target
haftmann
parents:
38311
diff
changeset
|
103 |
fun check_mixfix ctxt (b, extra_tfrees) mx = |
|
62767
d6b0d35b3aed
relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents:
62764
diff
changeset
|
104 |
if null extra_tfrees then mx |
|
d6b0d35b3aed
relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents:
62764
diff
changeset
|
105 |
else |
|
d6b0d35b3aed
relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents:
62764
diff
changeset
|
106 |
(if Context_Position.is_visible ctxt then |
|
d6b0d35b3aed
relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents:
62764
diff
changeset
|
107 |
warning |
|
d6b0d35b3aed
relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents:
62764
diff
changeset
|
108 |
("Additional type variable(s) in specification of " ^ Binding.print b ^ ": " ^
|
|
d6b0d35b3aed
relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents:
62764
diff
changeset
|
109 |
commas (map (Syntax.string_of_typ ctxt o TFree) (sort_by #1 extra_tfrees)) ^ |
|
d6b0d35b3aed
relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents:
62764
diff
changeset
|
110 |
(if Mixfix.is_empty mx then "" |
|
d6b0d35b3aed
relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents:
62764
diff
changeset
|
111 |
else "\nDropping mixfix syntax " ^ Pretty.string_of (Mixfix.pretty_mixfix mx))) |
|
d6b0d35b3aed
relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents:
62764
diff
changeset
|
112 |
else (); NoSyn); |
|
38312
9dd57db3c0f2
moved extra_tfrees check for mixfix syntax to Generic_Target
haftmann
parents:
38311
diff
changeset
|
113 |
|
|
47291
6a641856a0e9
better drop background syntax if entity depends on parameters;
wenzelm
parents:
47290
diff
changeset
|
114 |
fun check_mixfix_global (b, no_params) mx = |
| 62752 | 115 |
if no_params orelse Mixfix.is_empty mx then mx |
116 |
else |
|
117 |
(warning ("Dropping global mixfix syntax: " ^ Binding.print b ^ " " ^
|
|
|
62767
d6b0d35b3aed
relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents:
62764
diff
changeset
|
118 |
Pretty.string_of (Mixfix.pretty_mixfix mx)); NoSyn); |
|
47291
6a641856a0e9
better drop background syntax if entity depends on parameters;
wenzelm
parents:
47290
diff
changeset
|
119 |
|
|
61701
e89cfc004f18
Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents:
61261
diff
changeset
|
120 |
fun same_const (Const (c, _), Const (c', _)) = c = c' |
|
e89cfc004f18
Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents:
61261
diff
changeset
|
121 |
| same_const (t $ _, t' $ _) = same_const (t, t') |
|
e89cfc004f18
Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents:
61261
diff
changeset
|
122 |
| same_const (_, _) = false; |
|
e89cfc004f18
Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents:
61261
diff
changeset
|
123 |
|
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
124 |
fun const_decl phi_pred prmode ((b, mx), rhs) phi context = |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
125 |
if phi_pred phi then |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
126 |
let |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
127 |
val b' = Morphism.binding phi b; |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
128 |
val rhs' = Morphism.term phi rhs; |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
129 |
val same_shape = Term.aconv_untyped (rhs, rhs'); |
|
61701
e89cfc004f18
Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents:
61261
diff
changeset
|
130 |
val same_stem = same_shape orelse same_const (rhs, rhs'); |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
131 |
val const_alias = |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
132 |
if same_shape then |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
133 |
(case rhs' of |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
134 |
Const (c, T) => |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
135 |
let |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
136 |
val thy = Context.theory_of context; |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
137 |
val ctxt = Context.proof_of context; |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
138 |
in |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
139 |
(case Type_Infer_Context.const_type ctxt c of |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
140 |
SOME T' => if Sign.typ_equiv thy (T, T') then SOME c else NONE |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
141 |
| NONE => NONE) |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
142 |
end |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
143 |
| _ => NONE) |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
144 |
else NONE; |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
145 |
in |
| 60283 | 146 |
(case const_alias of |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
147 |
SOME c => |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
148 |
context |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
149 |
|> Context.mapping (Sign.const_alias b' c) (Proof_Context.const_alias b' c) |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
150 |
|> Morphism.form (Proof_Context.generic_notation true prmode [(rhs', mx)]) |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
151 |
| NONE => |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
152 |
context |
| 62764 | 153 |
|> Proof_Context.generic_add_abbrev Print_Mode.internal |
154 |
(b', Term.close_schematic_term rhs') |
|
|
61701
e89cfc004f18
Refine the supression of abbreviations for morphisms that are not identities.
ballarin
parents:
61261
diff
changeset
|
155 |
|-> (fn (const as Const (c, _), _) => same_stem ? |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
156 |
(Proof_Context.generic_revert_abbrev (#1 prmode) c #> |
| 62764 | 157 |
same_shape ? |
158 |
Morphism.form (Proof_Context.generic_notation true prmode [(const, mx)])))) |
|
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
159 |
end |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
160 |
else context; |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
161 |
|
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
162 |
|
| 60283 | 163 |
|
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
164 |
(** background primitives **) |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
165 |
|
|
71788
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
166 |
structure Foundation_Interpretations = Theory_Data |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
167 |
( |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
168 |
type T = (binding * (term * term list) -> Context.generic -> Context.generic) Inttab.table; |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
169 |
val empty = Inttab.empty; |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
170 |
val extend = I; |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
171 |
val merge = Inttab.merge (K true); |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
172 |
); |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
173 |
|
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
174 |
fun add_foundation_interpretation f = |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
175 |
Foundation_Interpretations.map (Inttab.update_new (serial (), f)); |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
176 |
|
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
177 |
fun foundation_interpretation binding_const_params lthy = |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
178 |
let |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
179 |
val interps = Foundation_Interpretations.get (Proof_Context.theory_of lthy); |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
180 |
val interp = Inttab.fold (fn (_, f) => f binding_const_params) interps; |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
181 |
in |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
182 |
lthy |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
183 |
|> Local_Theory.background_theory (Context.theory_map interp) |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
184 |
|> Local_Theory.map_contexts (K (Context.proof_map interp)) |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
185 |
end; |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
186 |
|
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
187 |
fun background_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) lthy = |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
188 |
let |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
189 |
val params = type_params @ term_params; |
|
71788
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
190 |
val target_params = type_params |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
191 |
@ take_prefix is_Free (Variable.export_terms lthy (Local_Theory.target_of lthy) term_params); |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
192 |
val mx' = check_mixfix_global (b, null params) mx; |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
193 |
|
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
194 |
val (const, lthy2) = lthy |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
195 |
|> Local_Theory.background_theory_result (Sign.declare_const lthy ((b, U), mx')); |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
196 |
val lhs = Term.list_comb (const, params); |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
197 |
|
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
198 |
val ((_, def), lthy3) = lthy2 |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
199 |
|> Local_Theory.background_theory_result |
|
61261
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
wenzelm
parents:
60924
diff
changeset
|
200 |
(Thm.add_def (Proof_Context.defs_context lthy2) false false |
|
71788
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
201 |
(Thm.def_binding_optional b b_def, Logic.mk_equals (lhs, rhs))) |
|
ca3ac5238c41
hooks for foundational terms: protection of foundational terms during simplification
haftmann
parents:
71178
diff
changeset
|
202 |
||> foundation_interpretation (b, (const, target_params)); |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
203 |
in ((lhs, def), lthy3) end; |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
204 |
|
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
205 |
fun background_declaration decl lthy = |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
206 |
let |
| 71178 | 207 |
fun theory_decl context = |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
208 |
Local_Theory.standard_form lthy |
| 71178 | 209 |
(Proof_Context.init_global (Context.theory_of context)) decl context; |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
210 |
in Local_Theory.background_theory (Context.theory_map theory_decl) lthy end; |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
211 |
|
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
212 |
fun background_abbrev (b, global_rhs) params = |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
213 |
Local_Theory.background_theory_result (Sign.add_abbrev Print_Mode.internal (b, global_rhs)) |
|
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58028
diff
changeset
|
214 |
#>> apply2 (fn t => Term.list_comb (Logic.unvarify_global t, params)) |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
215 |
|
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
216 |
|
| 60283 | 217 |
|
| 60340 | 218 |
(** nested local theories primitives **) |
219 |
||
220 |
fun standard_facts lthy ctxt = |
|
221 |
Attrib.transform_facts (Local_Theory.standard_morphism lthy ctxt); |
|
222 |
||
223 |
fun standard_notes pred kind facts lthy = |
|
224 |
Local_Theory.map_contexts (fn level => fn ctxt => |
|
225 |
if pred (Local_Theory.level lthy, level) |
|
226 |
then Attrib.local_notes kind (standard_facts lthy ctxt facts) ctxt |> snd |
|
227 |
else ctxt) lthy; |
|
228 |
||
229 |
fun standard_declaration pred decl lthy = |
|
230 |
Local_Theory.map_contexts (fn level => fn ctxt => |
|
231 |
if pred (Local_Theory.level lthy, level) |
|
232 |
then Context.proof_map (Local_Theory.standard_form lthy ctxt decl) ctxt |
|
233 |
else ctxt) lthy; |
|
234 |
||
235 |
fun standard_const pred prmode ((b, mx), rhs) = |
|
236 |
standard_declaration pred (const_decl (K true) prmode ((b, mx), rhs)); |
|
237 |
||
| 73846 | 238 |
fun standard_registration pred registration lthy = |
239 |
Local_Theory.map_contexts (fn level => |
|
240 |
if pred (Local_Theory.level lthy, level) |
|
241 |
then Context.proof_map (Locale.add_registration registration) |
|
242 |
else I) lthy; |
|
243 |
||
244 |
val local_interpretation = standard_registration (fn (n, level) => level = n - 1); |
|
| 73845 | 245 |
|
246 |
||
| 60340 | 247 |
|
248 |
(** lifting target primitives to local theory operations **) |
|
|
38312
9dd57db3c0f2
moved extra_tfrees check for mixfix syntax to Generic_Target
haftmann
parents:
38311
diff
changeset
|
249 |
|
| 38309 | 250 |
(* define *) |
251 |
||
|
46992
eeea81b86b70
refined Local_Theory.define vs. Local_Theory.define_internal, which allows to pass alternative name to the foundational axiom -- expecially important for 'instantiation' or 'overloading', which loose name information due to Long_Name.base_name cooking etc.;
wenzelm
parents:
46916
diff
changeset
|
252 |
fun define foundation internal ((b, mx), ((b_def, atts), rhs)) lthy = |
| 38309 | 253 |
let |
| 59616 | 254 |
val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy); |
| 38309 | 255 |
|
256 |
(*term and type parameters*) |
|
|
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59616
diff
changeset
|
257 |
val ((defs, _), rhs') = Thm.cterm_of lthy rhs |
| 47276 | 258 |
|> Local_Defs.export_cterm lthy thy_ctxt ||> Thm.term_of; |
| 38309 | 259 |
|
|
47272
ca31cfa509b1
refined define/abbrev: allow extra fixes in aux. context vs. bottom target (NB: export_term expands defined variables, leaving fixed ones);
wenzelm
parents:
47250
diff
changeset
|
260 |
val xs = Variable.add_fixed lthy rhs' []; |
| 38309 | 261 |
val T = Term.fastype_of rhs; |
262 |
val tfreesT = Term.add_tfreesT T (fold (Term.add_tfreesT o #2) xs []); |
|
263 |
val extra_tfrees = rev (subtract (op =) tfreesT (Term.add_tfrees rhs [])); |
|
|
38312
9dd57db3c0f2
moved extra_tfrees check for mixfix syntax to Generic_Target
haftmann
parents:
38311
diff
changeset
|
264 |
val mx' = check_mixfix lthy (b, extra_tfrees) mx; |
| 38309 | 265 |
|
266 |
val type_params = map (Logic.mk_type o TFree) extra_tfrees; |
|
|
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58028
diff
changeset
|
267 |
val term_params = map Free (sort (Variable.fixed_ord lthy o apply2 #1) xs); |
| 38309 | 268 |
val params = type_params @ term_params; |
269 |
||
270 |
val U = map Term.fastype_of params ---> T; |
|
271 |
||
272 |
(*foundation*) |
|
| 45353 | 273 |
val ((lhs', global_def), lthy2) = lthy |
274 |
|> foundation (((b, U), mx'), (b_def, rhs')) (type_params, term_params); |
|
| 38309 | 275 |
|
276 |
(*local definition*) |
|
| 63344 | 277 |
val ([(lhs, (_, local_def))], lthy3) = lthy2 |
| 63346 | 278 |
|> Context_Position.set_visible false |
| 63352 | 279 |
|> Local_Defs.define [((b, NoSyn), (Binding.empty_atts, lhs'))] |
| 63346 | 280 |
||> Context_Position.restore_visible lthy2; |
| 47240 | 281 |
|
282 |
(*result*) |
|
| 47238 | 283 |
val def = |
| 47240 | 284 |
Thm.transitive local_def global_def |
|
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59616
diff
changeset
|
285 |
|> Local_Defs.contract lthy3 defs (Thm.cterm_of lthy3 (Logic.mk_equals (lhs, rhs))); |
| 38315 | 286 |
val ([(res_name, [res])], lthy4) = lthy3 |
| 47080 | 287 |
|> Local_Theory.notes [((if internal then Binding.empty else b_def, atts), [([def], [])])]; |
| 38315 | 288 |
in ((lhs, (res_name, res)), lthy4) end; |
| 38309 | 289 |
|
290 |
||
291 |
(* notes *) |
|
292 |
||
|
47250
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
wenzelm
parents:
47249
diff
changeset
|
293 |
local |
|
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
wenzelm
parents:
47249
diff
changeset
|
294 |
|
| 38309 | 295 |
fun import_export_proof ctxt (name, raw_th) = |
296 |
let |
|
| 59616 | 297 |
val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of ctxt); |
| 38309 | 298 |
|
299 |
(*export assumes/defines*) |
|
|
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54742
diff
changeset
|
300 |
val th = Goal.norm_result ctxt raw_th; |
|
45407
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
wenzelm
parents:
45390
diff
changeset
|
301 |
val ((defs, asms), th') = Local_Defs.export ctxt thy_ctxt th; |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52230
diff
changeset
|
302 |
val asms' = map (rewrite_rule ctxt (Drule.norm_hhf_eqs @ defs)) asms; |
| 38309 | 303 |
|
304 |
(*export fixes*) |
|
305 |
val tfrees = map TFree (Thm.fold_terms Term.add_tfrees th' []); |
|
306 |
val frees = map Free (Thm.fold_terms Term.add_frees th' []); |
|
| 45352 | 307 |
val (th'' :: vs) = |
|
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59616
diff
changeset
|
308 |
(th' :: map (Drule.mk_term o Thm.cterm_of ctxt) (map Logic.mk_type tfrees @ frees)) |
| 38309 | 309 |
|> Variable.export ctxt thy_ctxt |
310 |
|> Drule.zero_var_indexes_list; |
|
311 |
||
312 |
(*thm definition*) |
|
| 70427 | 313 |
val result = Global_Theory.name_thm Global_Theory.official1 name th''; |
| 38309 | 314 |
|
315 |
(*import fixes*) |
|
316 |
val (tvars, vars) = |
|
317 |
chop (length tfrees) (map (Thm.term_of o Drule.dest_term) vs) |
|
318 |
|>> map Logic.dest_type; |
|
319 |
||
|
74220
c49134ee16c1
more scalable data structure (but: rarely used many arguments);
wenzelm
parents:
73846
diff
changeset
|
320 |
val instT = |
|
c49134ee16c1
more scalable data structure (but: rarely used many arguments);
wenzelm
parents:
73846
diff
changeset
|
321 |
fold2 (fn a => fn b => (case a of TVar v => Term_Subst.TVars.add (v, b) | _ => I)) |
|
c49134ee16c1
more scalable data structure (but: rarely used many arguments);
wenzelm
parents:
73846
diff
changeset
|
322 |
tvars tfrees Term_Subst.TVars.empty; |
|
c49134ee16c1
more scalable data structure (but: rarely used many arguments);
wenzelm
parents:
73846
diff
changeset
|
323 |
val cinstT = Term_Subst.TVars.fold (fn (v, T) => cons (v, Thm.ctyp_of ctxt T)) instT []; |
|
c49134ee16c1
more scalable data structure (but: rarely used many arguments);
wenzelm
parents:
73846
diff
changeset
|
324 |
val cinst = |
|
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
60345
diff
changeset
|
325 |
map_filter |
|
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
60345
diff
changeset
|
326 |
(fn (Var (xi, T), t) => |
|
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
60345
diff
changeset
|
327 |
SOME ((xi, Term_Subst.instantiateT instT T), |
|
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
60345
diff
changeset
|
328 |
Thm.cterm_of ctxt (Term.map_types (Term_Subst.instantiateT instT) t)) |
|
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
60345
diff
changeset
|
329 |
| _ => NONE) (vars ~~ frees); |
|
74220
c49134ee16c1
more scalable data structure (but: rarely used many arguments);
wenzelm
parents:
73846
diff
changeset
|
330 |
val result' = Thm.instantiate (cinstT, cinst) result; |
| 38309 | 331 |
|
332 |
(*import assumes/defines*) |
|
333 |
val result'' = |
|
|
45407
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
wenzelm
parents:
45390
diff
changeset
|
334 |
(fold (curry op COMP) asms' result' |
|
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
wenzelm
parents:
45390
diff
changeset
|
335 |
handle THM _ => raise THM ("Failed to re-import result", 0, result' :: asms'))
|
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52230
diff
changeset
|
336 |
|> Local_Defs.contract ctxt defs (Thm.cprop_of th) |
|
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54742
diff
changeset
|
337 |
|> Goal.norm_result ctxt |
| 70427 | 338 |
|> Global_Theory.name_thm Global_Theory.unofficial2 name; |
| 38309 | 339 |
|
340 |
in (result'', result) end; |
|
341 |
||
| 70494 | 342 |
fun bind_name lthy b = |
343 |
(Local_Theory.full_name lthy b, Binding.default_pos_of b); |
|
344 |
||
| 70544 | 345 |
fun map_facts f = map (apsnd (map (apfst (map f)))); |
346 |
||
|
47250
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
wenzelm
parents:
47249
diff
changeset
|
347 |
in |
|
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
wenzelm
parents:
47249
diff
changeset
|
348 |
|
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
349 |
fun notes target_notes kind facts lthy = |
| 38309 | 350 |
let |
351 |
val facts' = facts |
|
| 70494 | 352 |
|> map (fn (a, bs) => |
353 |
(a, Global_Theory.burrow_fact (Global_Theory.name_multi (bind_name lthy (fst a))) bs)) |
|
| 70544 | 354 |
|> map_facts (import_export_proof lthy); |
355 |
val local_facts = map_facts #1 facts'; |
|
356 |
val global_facts = map_facts #2 facts'; |
|
| 38309 | 357 |
in |
358 |
lthy |
|
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
359 |
|> target_notes kind global_facts (Attrib.partial_evaluation lthy local_facts) |
|
47249
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
wenzelm
parents:
47246
diff
changeset
|
360 |
|> Attrib.local_notes kind local_facts |
| 38309 | 361 |
end; |
362 |
||
|
47250
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
wenzelm
parents:
47249
diff
changeset
|
363 |
end; |
|
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
wenzelm
parents:
47249
diff
changeset
|
364 |
|
| 38309 | 365 |
|
366 |
(* abbrev *) |
|
367 |
||
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
368 |
fun abbrev target_abbrev prmode ((b, mx), rhs) lthy = |
| 38309 | 369 |
let |
|
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
370 |
val (global_rhs, (extra_tfrees, (type_params, term_params))) = export_abbrev lthy I rhs; |
|
38312
9dd57db3c0f2
moved extra_tfrees check for mixfix syntax to Generic_Target
haftmann
parents:
38311
diff
changeset
|
371 |
val mx' = check_mixfix lthy (b, extra_tfrees) mx; |
| 38309 | 372 |
in |
373 |
lthy |
|
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
374 |
|> target_abbrev prmode (b, mx') global_rhs (type_params, term_params) |
| 63346 | 375 |
|> Context_Position.set_visible false |
| 57109 | 376 |
|> Proof_Context.add_abbrev Print_Mode.internal (b, rhs) |> snd |
377 |
|> Local_Defs.fixed_abbrev ((b, NoSyn), rhs) |
|
| 62992 | 378 |
||> Context_Position.restore_visible lthy |
| 38309 | 379 |
end; |
380 |
||
|
38341
72dba5bd5f63
moved theory-level target operation fragements to Generic_Target; adjusted bootstrap order
haftmann
parents:
38315
diff
changeset
|
381 |
|
| 60283 | 382 |
|
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
383 |
(** theory target primitives **) |
|
38341
72dba5bd5f63
moved theory-level target operation fragements to Generic_Target; adjusted bootstrap order
haftmann
parents:
38315
diff
changeset
|
384 |
|
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
385 |
fun theory_target_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) = |
|
47284
0e65b6a016dc
clarified background_foundation vs. theory_foundation (with const_declaration);
wenzelm
parents:
47283
diff
changeset
|
386 |
background_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) |
|
57192
180e955711cf
yet another attempt for terminology: foo_target_bar denotes an operation bar operating solely on the target context of target foo, foo_bar denotes a whole stack of operations to accomplish bar for target foo
haftmann
parents:
57191
diff
changeset
|
387 |
#-> (fn (lhs, def) => standard_const (op <>) Syntax.mode_default ((b, mx), lhs) |
| 57057 | 388 |
#> pair (lhs, def)); |
|
47284
0e65b6a016dc
clarified background_foundation vs. theory_foundation (with const_declaration);
wenzelm
parents:
47283
diff
changeset
|
389 |
|
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
390 |
fun theory_target_notes kind global_facts local_facts = |
| 57191 | 391 |
Local_Theory.background_theory (Attrib.global_notes kind global_facts #> snd) |
392 |
#> standard_notes (op <>) kind local_facts; |
|
|
38341
72dba5bd5f63
moved theory-level target operation fragements to Generic_Target; adjusted bootstrap order
haftmann
parents:
38315
diff
changeset
|
393 |
|
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
394 |
fun theory_target_abbrev prmode (b, mx) global_rhs params = |
|
47286
392c4cd97e5c
more uniform theory_abbrev with const_declaration;
wenzelm
parents:
47285
diff
changeset
|
395 |
Local_Theory.background_theory_result |
| 57109 | 396 |
(Sign.add_abbrev (#1 prmode) (b, global_rhs) #-> |
|
47291
6a641856a0e9
better drop background syntax if entity depends on parameters;
wenzelm
parents:
47290
diff
changeset
|
397 |
(fn (lhs, _) => (* FIXME type_params!? *) |
| 63339 | 398 |
Sign.notation true prmode |
399 |
[(lhs, check_mixfix_global (b, null (snd params)) mx)] #> pair lhs)) |
|
| 62764 | 400 |
#-> (fn lhs => |
401 |
standard_const (op <>) prmode |
|
402 |
((b, if null (snd params) then NoSyn else mx), |
|
403 |
Term.list_comb (Logic.unvarify_global lhs, snd params))); |
|
|
38341
72dba5bd5f63
moved theory-level target operation fragements to Generic_Target; adjusted bootstrap order
haftmann
parents:
38315
diff
changeset
|
404 |
|
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
405 |
|
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
406 |
(** theory operations **) |
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
407 |
|
|
60344
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60342
diff
changeset
|
408 |
val theory_abbrev = abbrev theory_target_abbrev; |
|
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60342
diff
changeset
|
409 |
|
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
410 |
fun theory_declaration decl = |
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
411 |
background_declaration decl #> standard_declaration (K true) decl; |
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
412 |
|
| 73846 | 413 |
fun target_registration lthy {inst, mixin, export} =
|
414 |
{inst = inst, mixin = mixin,
|
|
415 |
export = export $> Proof_Context.export_morphism lthy (Local_Theory.target_of lthy)}; |
|
416 |
||
417 |
fun theory_registration registration lthy = |
|
418 |
lthy |
|
419 |
|> (Local_Theory.raw_theory o Context.theory_map) |
|
420 |
(Locale.add_registration (target_registration lthy registration)) |
|
421 |
|> standard_registration (K true) registration; |
|
| 73845 | 422 |
|
| 60283 | 423 |
|
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
424 |
(** locale target primitives **) |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
425 |
|
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
426 |
fun locale_target_notes locale kind global_facts local_facts = |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
427 |
Local_Theory.background_theory |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
428 |
(Attrib.global_notes kind (Attrib.map_facts (K []) global_facts) #> snd) #> |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
429 |
(fn lthy => lthy |> |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
430 |
Local_Theory.target (fn ctxt => ctxt |> |
| 67654 | 431 |
Locale.add_facts locale kind (standard_facts lthy ctxt local_facts))) #> |
| 57191 | 432 |
standard_notes (fn (this, other) => other <> 0 andalso this <> other) kind local_facts; |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
433 |
|
|
57192
180e955711cf
yet another attempt for terminology: foo_target_bar denotes an operation bar operating solely on the target context of target foo, foo_bar denotes a whole stack of operations to accomplish bar for target foo
haftmann
parents:
57191
diff
changeset
|
434 |
fun locale_target_declaration locale syntax decl lthy = lthy |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
435 |
|> Local_Theory.target (fn ctxt => ctxt |> |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
436 |
Locale.add_declaration locale syntax |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
437 |
(Morphism.transform (Local_Theory.standard_morphism lthy ctxt) decl)); |
|
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
438 |
|
|
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
439 |
fun locale_target_const locale phi_pred prmode ((b, mx), rhs) = |
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
440 |
locale_target_declaration locale true (const_decl phi_pred prmode ((b, mx), rhs)) |
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
441 |
|
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
442 |
|
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
443 |
(** locale operations **) |
|
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60340
diff
changeset
|
444 |
|
| 57193 | 445 |
fun locale_declaration locale {syntax, pervasive} decl =
|
446 |
pervasive ? background_declaration decl |
|
447 |
#> locale_target_declaration locale syntax decl |
|
448 |
#> standard_declaration (fn (_, other) => other <> 0) decl; |
|
449 |
||
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
450 |
fun locale_const locale prmode ((b, mx), rhs) = |
|
57192
180e955711cf
yet another attempt for terminology: foo_target_bar denotes an operation bar operating solely on the target context of target foo, foo_bar denotes a whole stack of operations to accomplish bar for target foo
haftmann
parents:
57191
diff
changeset
|
451 |
locale_target_const locale (K true) prmode ((b, mx), rhs) |
|
180e955711cf
yet another attempt for terminology: foo_target_bar denotes an operation bar operating solely on the target context of target foo, foo_bar denotes a whole stack of operations to accomplish bar for target foo
haftmann
parents:
57191
diff
changeset
|
452 |
#> standard_const (fn (this, other) => other <> 0 andalso this <> other) prmode ((b, mx), rhs); |
|
57190
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
haftmann
parents:
57189
diff
changeset
|
453 |
|
| 73846 | 454 |
fun locale_dependency loc registration lthy = |
455 |
lthy |
|
456 |
|> Local_Theory.raw_theory (Locale.add_dependency loc registration) |
|
457 |
|> standard_registration (K true) registration; |
|
| 73845 | 458 |
|
|
60344
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60342
diff
changeset
|
459 |
|
|
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60342
diff
changeset
|
460 |
(** locale abbreviations **) |
|
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60342
diff
changeset
|
461 |
|
|
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60342
diff
changeset
|
462 |
fun locale_target_abbrev locale prmode (b, mx) global_rhs params = |
|
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60342
diff
changeset
|
463 |
background_abbrev (b, global_rhs) (snd params) |
|
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60342
diff
changeset
|
464 |
#-> (fn (lhs, _) => locale_const locale prmode ((b, mx), lhs)); |
|
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60342
diff
changeset
|
465 |
|
|
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60342
diff
changeset
|
466 |
fun locale_abbrev locale = abbrev (locale_target_abbrev locale); |
|
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60342
diff
changeset
|
467 |
|
| 38309 | 468 |
end; |