| author | wenzelm | 
| Wed, 06 Apr 2016 17:17:05 +0200 | |
| changeset 62892 | 7485507620b6 | 
| parent 62767 | d6b0d35b3aed | 
| child 62992 | d2e3b3b159d7 | 
| 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  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
21  | 
|
| 60340 | 22  | 
(*nested local theories primitives*)  | 
23  | 
val standard_const: (int * int -> bool) -> Syntax.mode -> (binding * mixfix) * term ->  | 
|
24  | 
local_theory -> local_theory  | 
|
25  | 
||
26  | 
(*lifting target primitives to local theory operations*)  | 
|
| 40782 | 27  | 
val define: (((binding * typ) * mixfix) * (binding * term) ->  | 
28  | 
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
 | 
29  | 
bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory ->  | 
| 40782 | 30  | 
(term * (string * thm)) * local_theory  | 
| 
47250
 
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
 
wenzelm 
parents: 
47249 
diff
changeset
 | 
31  | 
val notes:  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
32  | 
(string -> (Attrib.binding * (thm list * Token.src list) list) list ->  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
33  | 
(Attrib.binding * (thm list * Token.src list) list) list -> local_theory -> local_theory) ->  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
34  | 
string -> (Attrib.binding * (thm list * Token.src list) list) list -> local_theory ->  | 
| 40782 | 35  | 
(string * thm list) list * local_theory  | 
| 60337 | 36  | 
val abbrev: (Syntax.mode -> binding * mixfix -> term ->  | 
| 57160 | 37  | 
term list * term list -> local_theory -> local_theory) ->  | 
| 60337 | 38  | 
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
 | 
39  | 
|
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
40  | 
(*theory target primitives*)  | 
| 
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
41  | 
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
 | 
42  | 
term list * term list -> local_theory -> (term * thm) * local_theory  | 
| 
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
43  | 
val theory_target_notes: string ->  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
44  | 
(Attrib.binding * (thm list * Token.src list) list) list ->  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
45  | 
(Attrib.binding * (thm list * Token.src list) list) list ->  | 
| 40782 | 46  | 
local_theory -> local_theory  | 
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
47  | 
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
 | 
48  | 
local_theory -> local_theory  | 
| 
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
49  | 
|
| 
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
50  | 
(*theory target operations*)  | 
| 
60344
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
51  | 
val theory_abbrev: Syntax.mode -> (binding * mixfix) * term ->  | 
| 
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
52  | 
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
 | 
53  | 
val theory_declaration: declaration -> local_theory -> local_theory  | 
| 
52153
 
f5773a46cf05
more specific structure for registration into theory and dependency onto locale
 
haftmann 
parents: 
47313 
diff
changeset
 | 
54  | 
val theory_registration: string * morphism -> (morphism * bool) option -> morphism ->  | 
| 
 
f5773a46cf05
more specific structure for registration into theory and dependency onto locale
 
haftmann 
parents: 
47313 
diff
changeset
 | 
55  | 
local_theory -> local_theory  | 
| 
57190
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
56  | 
|
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
57  | 
(*locale target primitives*)  | 
| 
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
58  | 
val locale_target_notes: string -> string ->  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
59  | 
(Attrib.binding * (thm list * Token.src list) list) list ->  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
60  | 
(Attrib.binding * (thm list * Token.src list) list) list ->  | 
| 
57190
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
61  | 
local_theory -> local_theory  | 
| 
60344
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
62  | 
val locale_target_abbrev: string -> Syntax.mode -> (binding * mixfix) -> term -> term list * term list ->  | 
| 
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
63  | 
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
 | 
64  | 
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
 | 
65  | 
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
 | 
66  | 
(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
 | 
67  | 
|
| 
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
68  | 
(*locale operations*)  | 
| 
60344
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
69  | 
val locale_abbrev: string -> Syntax.mode -> (binding * mixfix) * term ->  | 
| 
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
70  | 
local_theory -> (term * term) * local_theory  | 
| 57193 | 71  | 
  val locale_declaration: string -> {syntax: bool, pervasive: bool} -> declaration ->
 | 
72  | 
local_theory -> local_theory  | 
|
| 
57190
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
73  | 
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
 | 
74  | 
local_theory -> local_theory  | 
| 
52153
 
f5773a46cf05
more specific structure for registration into theory and dependency onto locale
 
haftmann 
parents: 
47313 
diff
changeset
 | 
75  | 
val locale_dependency: string -> string * morphism -> (morphism * bool) option -> morphism ->  | 
| 
 
f5773a46cf05
more specific structure for registration into theory and dependency onto locale
 
haftmann 
parents: 
47313 
diff
changeset
 | 
76  | 
local_theory -> local_theory  | 
| 45353 | 77  | 
end  | 
| 38309 | 78  | 
|
79  | 
structure Generic_Target: GENERIC_TARGET =  | 
|
80  | 
struct  | 
|
81  | 
||
| 
57190
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
82  | 
(** consts **)  | 
| 
38312
 
9dd57db3c0f2
moved extra_tfrees check for mixfix syntax to Generic_Target
 
haftmann 
parents: 
38311 
diff
changeset
 | 
83  | 
|
| 
60345
 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 
haftmann 
parents: 
60344 
diff
changeset
 | 
84  | 
fun export_abbrev lthy preprocess rhs =  | 
| 
60342
 
df9b153d866b
separate function to compute exported abbreviation
 
haftmann 
parents: 
60341 
diff
changeset
 | 
85  | 
let  | 
| 
 
df9b153d866b
separate function to compute exported abbreviation
 
haftmann 
parents: 
60341 
diff
changeset
 | 
86  | 
val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy);  | 
| 
 
df9b153d866b
separate function to compute exported abbreviation
 
haftmann 
parents: 
60341 
diff
changeset
 | 
87  | 
|
| 
 
df9b153d866b
separate function to compute exported abbreviation
 
haftmann 
parents: 
60341 
diff
changeset
 | 
88  | 
val rhs' = rhs  | 
| 
 
df9b153d866b
separate function to compute exported abbreviation
 
haftmann 
parents: 
60341 
diff
changeset
 | 
89  | 
|> Assumption.export_term lthy (Local_Theory.target_of lthy)  | 
| 
 
df9b153d866b
separate function to compute exported abbreviation
 
haftmann 
parents: 
60341 
diff
changeset
 | 
90  | 
|> preprocess;  | 
| 
 
df9b153d866b
separate function to compute exported abbreviation
 
haftmann 
parents: 
60341 
diff
changeset
 | 
91  | 
val term_params = map Free (sort (Variable.fixed_ord lthy o apply2 #1) (Variable.add_fixed lthy rhs' []));  | 
| 
 
df9b153d866b
separate function to compute exported abbreviation
 
haftmann 
parents: 
60341 
diff
changeset
 | 
92  | 
val u = fold_rev lambda term_params rhs';  | 
| 
 
df9b153d866b
separate function to compute exported abbreviation
 
haftmann 
parents: 
60341 
diff
changeset
 | 
93  | 
val global_rhs = singleton (Variable.polymorphic thy_ctxt) u;  | 
| 
 
df9b153d866b
separate function to compute exported abbreviation
 
haftmann 
parents: 
60341 
diff
changeset
 | 
94  | 
|
| 
 
df9b153d866b
separate function to compute exported abbreviation
 
haftmann 
parents: 
60341 
diff
changeset
 | 
95  | 
val extra_tfrees =  | 
| 
 
df9b153d866b
separate function to compute exported abbreviation
 
haftmann 
parents: 
60341 
diff
changeset
 | 
96  | 
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
 | 
97  | 
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
 | 
98  | 
in (global_rhs, (extra_tfrees, (type_params, term_params))) end;  | 
| 
60342
 
df9b153d866b
separate function to compute exported abbreviation
 
haftmann 
parents: 
60341 
diff
changeset
 | 
99  | 
|
| 
38312
 
9dd57db3c0f2
moved extra_tfrees check for mixfix syntax to Generic_Target
 
haftmann 
parents: 
38311 
diff
changeset
 | 
100  | 
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
 | 
101  | 
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
 | 
102  | 
else  | 
| 
 
d6b0d35b3aed
relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
 
wenzelm 
parents: 
62764 
diff
changeset
 | 
103  | 
(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
 | 
104  | 
warning  | 
| 
 
d6b0d35b3aed
relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
 
wenzelm 
parents: 
62764 
diff
changeset
 | 
105  | 
        ("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
 | 
106  | 
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
 | 
107  | 
(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
 | 
108  | 
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
 | 
109  | 
else (); NoSyn);  | 
| 
38312
 
9dd57db3c0f2
moved extra_tfrees check for mixfix syntax to Generic_Target
 
haftmann 
parents: 
38311 
diff
changeset
 | 
110  | 
|
| 
47291
 
6a641856a0e9
better drop background syntax if entity depends on parameters;
 
wenzelm 
parents: 
47290 
diff
changeset
 | 
111  | 
fun check_mixfix_global (b, no_params) mx =  | 
| 62752 | 112  | 
if no_params orelse Mixfix.is_empty mx then mx  | 
113  | 
else  | 
|
114  | 
    (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
 | 
115  | 
Pretty.string_of (Mixfix.pretty_mixfix mx)); NoSyn);  | 
| 
47291
 
6a641856a0e9
better drop background syntax if entity depends on parameters;
 
wenzelm 
parents: 
47290 
diff
changeset
 | 
116  | 
|
| 
61701
 
e89cfc004f18
Refine the supression of abbreviations for morphisms that are not identities.
 
ballarin 
parents: 
61261 
diff
changeset
 | 
117  | 
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
 | 
118  | 
| 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
 | 
119  | 
| same_const (_, _) = false;  | 
| 
 
e89cfc004f18
Refine the supression of abbreviations for morphisms that are not identities.
 
ballarin 
parents: 
61261 
diff
changeset
 | 
120  | 
|
| 
57190
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
121  | 
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
 | 
122  | 
if phi_pred phi then  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
123  | 
let  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
124  | 
val b' = Morphism.binding phi b;  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
125  | 
val rhs' = Morphism.term phi rhs;  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
126  | 
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
 | 
127  | 
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
 | 
128  | 
val const_alias =  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
129  | 
if same_shape then  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
130  | 
(case rhs' of  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
131  | 
Const (c, T) =>  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
132  | 
let  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
133  | 
val thy = Context.theory_of context;  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
134  | 
val ctxt = Context.proof_of context;  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
135  | 
in  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
136  | 
(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
 | 
137  | 
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
 | 
138  | 
| NONE => NONE)  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
139  | 
end  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
140  | 
| _ => NONE)  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
141  | 
else NONE;  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
142  | 
in  | 
| 60283 | 143  | 
(case const_alias of  | 
| 
57190
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
144  | 
SOME c =>  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
145  | 
context  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
146  | 
|> 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
 | 
147  | 
|> 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
 | 
148  | 
| NONE =>  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
149  | 
context  | 
| 62764 | 150  | 
|> Proof_Context.generic_add_abbrev Print_Mode.internal  | 
151  | 
(b', Term.close_schematic_term rhs')  | 
|
| 
61701
 
e89cfc004f18
Refine the supression of abbreviations for morphisms that are not identities.
 
ballarin 
parents: 
61261 
diff
changeset
 | 
152  | 
|-> (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
 | 
153  | 
(Proof_Context.generic_revert_abbrev (#1 prmode) c #>  | 
| 62764 | 154  | 
same_shape ?  | 
155  | 
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
 | 
156  | 
end  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
157  | 
else context;  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
158  | 
|
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
159  | 
|
| 60283 | 160  | 
|
| 
57190
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
161  | 
(** background primitives **)  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
162  | 
|
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
163  | 
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
 | 
164  | 
let  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
165  | 
val params = type_params @ term_params;  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
166  | 
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
 | 
167  | 
|
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
168  | 
val (const, lthy2) = lthy  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
169  | 
|> 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
 | 
170  | 
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
 | 
171  | 
|
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
172  | 
val ((_, def), lthy3) = lthy2  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
173  | 
|> Local_Theory.background_theory_result  | 
| 
61261
 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 
wenzelm 
parents: 
60924 
diff
changeset
 | 
174  | 
(Thm.add_def (Proof_Context.defs_context lthy2) false false  | 
| 
57190
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
175  | 
(Thm.def_binding_optional b b_def, Logic.mk_equals (lhs, rhs)));  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
176  | 
in ((lhs, def), lthy3) end;  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
177  | 
|
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
178  | 
fun background_declaration decl lthy =  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
179  | 
let  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
180  | 
val theory_decl =  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
181  | 
Local_Theory.standard_form lthy  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
182  | 
(Proof_Context.init_global (Proof_Context.theory_of lthy)) decl;  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
183  | 
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
 | 
184  | 
|
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
185  | 
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
 | 
186  | 
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
 | 
187  | 
#>> 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
 | 
188  | 
|
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
189  | 
|
| 60283 | 190  | 
|
| 60340 | 191  | 
(** nested local theories primitives **)  | 
192  | 
||
193  | 
fun standard_facts lthy ctxt =  | 
|
194  | 
Attrib.transform_facts (Local_Theory.standard_morphism lthy ctxt);  | 
|
195  | 
||
196  | 
fun standard_notes pred kind facts lthy =  | 
|
197  | 
Local_Theory.map_contexts (fn level => fn ctxt =>  | 
|
198  | 
if pred (Local_Theory.level lthy, level)  | 
|
199  | 
then Attrib.local_notes kind (standard_facts lthy ctxt facts) ctxt |> snd  | 
|
200  | 
else ctxt) lthy;  | 
|
201  | 
||
202  | 
fun standard_declaration pred decl lthy =  | 
|
203  | 
Local_Theory.map_contexts (fn level => fn ctxt =>  | 
|
204  | 
if pred (Local_Theory.level lthy, level)  | 
|
205  | 
then Context.proof_map (Local_Theory.standard_form lthy ctxt decl) ctxt  | 
|
206  | 
else ctxt) lthy;  | 
|
207  | 
||
208  | 
fun standard_const pred prmode ((b, mx), rhs) =  | 
|
209  | 
standard_declaration pred (const_decl (K true) prmode ((b, mx), rhs));  | 
|
210  | 
||
211  | 
||
212  | 
(** lifting target primitives to local theory operations **)  | 
|
| 
38312
 
9dd57db3c0f2
moved extra_tfrees check for mixfix syntax to Generic_Target
 
haftmann 
parents: 
38311 
diff
changeset
 | 
213  | 
|
| 38309 | 214  | 
(* define *)  | 
215  | 
||
| 
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
 | 
216  | 
fun define foundation internal ((b, mx), ((b_def, atts), rhs)) lthy =  | 
| 38309 | 217  | 
let  | 
| 59616 | 218  | 
val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy);  | 
| 38309 | 219  | 
|
220  | 
(*term and type parameters*)  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59616 
diff
changeset
 | 
221  | 
val ((defs, _), rhs') = Thm.cterm_of lthy rhs  | 
| 47276 | 222  | 
|> Local_Defs.export_cterm lthy thy_ctxt ||> Thm.term_of;  | 
| 38309 | 223  | 
|
| 
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
 | 
224  | 
val xs = Variable.add_fixed lthy rhs' [];  | 
| 38309 | 225  | 
val T = Term.fastype_of rhs;  | 
226  | 
val tfreesT = Term.add_tfreesT T (fold (Term.add_tfreesT o #2) xs []);  | 
|
227  | 
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
 | 
228  | 
val mx' = check_mixfix lthy (b, extra_tfrees) mx;  | 
| 38309 | 229  | 
|
230  | 
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
 | 
231  | 
val term_params = map Free (sort (Variable.fixed_ord lthy o apply2 #1) xs);  | 
| 38309 | 232  | 
val params = type_params @ term_params;  | 
233  | 
||
234  | 
val U = map Term.fastype_of params ---> T;  | 
|
235  | 
||
236  | 
(*foundation*)  | 
|
| 45353 | 237  | 
val ((lhs', global_def), lthy2) = lthy  | 
238  | 
|> foundation (((b, U), mx'), (b_def, rhs')) (type_params, term_params);  | 
|
| 38309 | 239  | 
|
240  | 
(*local definition*)  | 
|
| 38315 | 241  | 
val ((lhs, local_def), lthy3) = lthy2  | 
| 38309 | 242  | 
|> Local_Defs.add_def ((b, NoSyn), lhs');  | 
| 47240 | 243  | 
|
244  | 
(*result*)  | 
|
| 47238 | 245  | 
val def =  | 
| 47240 | 246  | 
Thm.transitive local_def global_def  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59616 
diff
changeset
 | 
247  | 
|> Local_Defs.contract lthy3 defs (Thm.cterm_of lthy3 (Logic.mk_equals (lhs, rhs)));  | 
| 38315 | 248  | 
val ([(res_name, [res])], lthy4) = lthy3  | 
| 47080 | 249  | 
|> Local_Theory.notes [((if internal then Binding.empty else b_def, atts), [([def], [])])];  | 
| 38315 | 250  | 
in ((lhs, (res_name, res)), lthy4) end;  | 
| 38309 | 251  | 
|
252  | 
||
253  | 
(* notes *)  | 
|
254  | 
||
| 
47250
 
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
 
wenzelm 
parents: 
47249 
diff
changeset
 | 
255  | 
local  | 
| 
 
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
 
wenzelm 
parents: 
47249 
diff
changeset
 | 
256  | 
|
| 38309 | 257  | 
fun import_export_proof ctxt (name, raw_th) =  | 
258  | 
let  | 
|
| 59616 | 259  | 
val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of ctxt);  | 
| 38309 | 260  | 
|
261  | 
(*export assumes/defines*)  | 
|
| 
54883
 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 
wenzelm 
parents: 
54742 
diff
changeset
 | 
262  | 
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
 | 
263  | 
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
 | 
264  | 
val asms' = map (rewrite_rule ctxt (Drule.norm_hhf_eqs @ defs)) asms;  | 
| 38309 | 265  | 
|
266  | 
(*export fixes*)  | 
|
267  | 
val tfrees = map TFree (Thm.fold_terms Term.add_tfrees th' []);  | 
|
268  | 
val frees = map Free (Thm.fold_terms Term.add_frees th' []);  | 
|
| 45352 | 269  | 
val (th'' :: vs) =  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59616 
diff
changeset
 | 
270  | 
(th' :: map (Drule.mk_term o Thm.cterm_of ctxt) (map Logic.mk_type tfrees @ frees))  | 
| 38309 | 271  | 
|> Variable.export ctxt thy_ctxt  | 
272  | 
|> Drule.zero_var_indexes_list;  | 
|
273  | 
||
274  | 
(*thm definition*)  | 
|
| 
45413
 
117ff038f8f7
disabled Thm.compress (again) -- costs for building tables tend to be higher than potential benefit;
 
wenzelm 
parents: 
45407 
diff
changeset
 | 
275  | 
val result = Global_Theory.name_thm true true name th'';  | 
| 38309 | 276  | 
|
277  | 
(*import fixes*)  | 
|
278  | 
val (tvars, vars) =  | 
|
279  | 
chop (length tfrees) (map (Thm.term_of o Drule.dest_term) vs)  | 
|
280  | 
|>> map Logic.dest_type;  | 
|
281  | 
||
282  | 
val instT = map_filter (fn (TVar v, T) => SOME (v, T) | _ => NONE) (tvars ~~ tfrees);  | 
|
| 
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
 | 
283  | 
val inst =  | 
| 
 
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
 | 
284  | 
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
 | 
285  | 
(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
 | 
286  | 
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
 | 
287  | 
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
 | 
288  | 
| _ => NONE) (vars ~~ frees);  | 
| 
 
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
 | 
289  | 
val result' = Thm.instantiate (map (apsnd (Thm.ctyp_of ctxt)) instT, inst) result;  | 
| 38309 | 290  | 
|
291  | 
(*import assumes/defines*)  | 
|
292  | 
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
 | 
293  | 
(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
 | 
294  | 
        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
 | 
295  | 
|> Local_Defs.contract ctxt defs (Thm.cprop_of th)  | 
| 
54883
 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 
wenzelm 
parents: 
54742 
diff
changeset
 | 
296  | 
|> Goal.norm_result ctxt  | 
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
38831 
diff
changeset
 | 
297  | 
|> Global_Theory.name_thm false false name;  | 
| 38309 | 298  | 
|
299  | 
in (result'', result) end;  | 
|
300  | 
||
| 
47250
 
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
 
wenzelm 
parents: 
47249 
diff
changeset
 | 
301  | 
in  | 
| 
 
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
 
wenzelm 
parents: 
47249 
diff
changeset
 | 
302  | 
|
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
303  | 
fun notes target_notes kind facts lthy =  | 
| 38309 | 304  | 
let  | 
305  | 
val facts' = facts  | 
|
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
38831 
diff
changeset
 | 
306  | 
|> map (fn (a, bs) => (a, Global_Theory.burrow_fact (Global_Theory.name_multi  | 
| 38309 | 307  | 
(Local_Theory.full_name lthy (fst a))) bs))  | 
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
38831 
diff
changeset
 | 
308  | 
|> Global_Theory.map_facts (import_export_proof lthy);  | 
| 
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
38831 
diff
changeset
 | 
309  | 
val local_facts = Global_Theory.map_facts #1 facts';  | 
| 
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
38831 
diff
changeset
 | 
310  | 
val global_facts = Global_Theory.map_facts #2 facts';  | 
| 38309 | 311  | 
in  | 
312  | 
lthy  | 
|
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
313  | 
|> 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
 | 
314  | 
|> Attrib.local_notes kind local_facts  | 
| 38309 | 315  | 
end;  | 
316  | 
||
| 
47250
 
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
 
wenzelm 
parents: 
47249 
diff
changeset
 | 
317  | 
end;  | 
| 
 
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
 
wenzelm 
parents: 
47249 
diff
changeset
 | 
318  | 
|
| 38309 | 319  | 
|
320  | 
(* abbrev *)  | 
|
321  | 
||
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
322  | 
fun abbrev target_abbrev prmode ((b, mx), rhs) lthy =  | 
| 38309 | 323  | 
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
 | 
324  | 
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
 | 
325  | 
val mx' = check_mixfix lthy (b, extra_tfrees) mx;  | 
| 38309 | 326  | 
in  | 
327  | 
lthy  | 
|
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
328  | 
|> target_abbrev prmode (b, mx') global_rhs (type_params, term_params)  | 
| 57109 | 329  | 
|> Proof_Context.add_abbrev Print_Mode.internal (b, rhs) |> snd  | 
330  | 
|> Local_Defs.fixed_abbrev ((b, NoSyn), rhs)  | 
|
| 38309 | 331  | 
end;  | 
332  | 
||
| 
38341
 
72dba5bd5f63
moved theory-level target operation fragements to Generic_Target; adjusted bootstrap order
 
haftmann 
parents: 
38315 
diff
changeset
 | 
333  | 
|
| 60283 | 334  | 
|
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
335  | 
(** theory target primitives **)  | 
| 
38341
 
72dba5bd5f63
moved theory-level target operation fragements to Generic_Target; adjusted bootstrap order
 
haftmann 
parents: 
38315 
diff
changeset
 | 
336  | 
|
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
337  | 
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
 | 
338  | 
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
 | 
339  | 
#-> (fn (lhs, def) => standard_const (op <>) Syntax.mode_default ((b, mx), lhs)  | 
| 57057 | 340  | 
#> pair (lhs, def));  | 
| 
47284
 
0e65b6a016dc
clarified background_foundation vs. theory_foundation (with const_declaration);
 
wenzelm 
parents: 
47283 
diff
changeset
 | 
341  | 
|
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
342  | 
fun theory_target_notes kind global_facts local_facts =  | 
| 57191 | 343  | 
Local_Theory.background_theory (Attrib.global_notes kind global_facts #> snd)  | 
344  | 
#> 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
 | 
345  | 
|
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
346  | 
fun theory_target_abbrev prmode (b, mx) global_rhs params =  | 
| 
47286
 
392c4cd97e5c
more uniform theory_abbrev with const_declaration;
 
wenzelm 
parents: 
47285 
diff
changeset
 | 
347  | 
Local_Theory.background_theory_result  | 
| 57109 | 348  | 
(Sign.add_abbrev (#1 prmode) (b, global_rhs) #->  | 
| 
47291
 
6a641856a0e9
better drop background syntax if entity depends on parameters;
 
wenzelm 
parents: 
47290 
diff
changeset
 | 
349  | 
(fn (lhs, _) => (* FIXME type_params!? *)  | 
| 57160 | 350  | 
Sign.notation true prmode [(lhs, check_mixfix_global (b, null (snd params)) mx)] #> pair lhs))  | 
| 62764 | 351  | 
#-> (fn lhs =>  | 
352  | 
standard_const (op <>) prmode  | 
|
353  | 
((b, if null (snd params) then NoSyn else mx),  | 
|
354  | 
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
 | 
355  | 
|
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
356  | 
|
| 
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
357  | 
(** theory operations **)  | 
| 
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
358  | 
|
| 
60344
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
359  | 
val theory_abbrev = abbrev theory_target_abbrev;  | 
| 
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
360  | 
|
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
361  | 
fun theory_declaration decl =  | 
| 
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
362  | 
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
 | 
363  | 
|
| 
57190
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
364  | 
val theory_registration =  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
365  | 
Local_Theory.raw_theory o Context.theory_map ooo Locale.add_registration;  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
366  | 
|
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
367  | 
|
| 60283 | 368  | 
|
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
369  | 
(** locale target primitives **)  | 
| 
57190
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
370  | 
|
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
371  | 
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
 | 
372  | 
Local_Theory.background_theory  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
373  | 
(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
 | 
374  | 
(fn lthy => lthy |>  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
375  | 
Local_Theory.target (fn ctxt => ctxt |>  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
376  | 
Locale.add_thmss locale kind (standard_facts lthy ctxt local_facts))) #>  | 
| 57191 | 377  | 
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
 | 
378  | 
|
| 
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
 | 
379  | 
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
 | 
380  | 
|> Local_Theory.target (fn ctxt => ctxt |>  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
381  | 
Locale.add_declaration locale syntax  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
382  | 
(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
 | 
383  | 
|
| 
60341
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
384  | 
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
 | 
385  | 
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
 | 
386  | 
|
| 
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
387  | 
|
| 
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
388  | 
(** locale operations **)  | 
| 
 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 
haftmann 
parents: 
60340 
diff
changeset
 | 
389  | 
|
| 57193 | 390  | 
fun locale_declaration locale {syntax, pervasive} decl =
 | 
391  | 
pervasive ? background_declaration decl  | 
|
392  | 
#> locale_target_declaration locale syntax decl  | 
|
393  | 
#> standard_declaration (fn (_, other) => other <> 0) decl;  | 
|
394  | 
||
| 
57190
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
395  | 
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
 | 
396  | 
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
 | 
397  | 
#> 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
 | 
398  | 
|
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
399  | 
fun locale_dependency locale dep_morph mixin export =  | 
| 
 
05ad9aae4537
recovered structure of module, which got somehow convoluted due to incremental modifications
 
haftmann 
parents: 
57189 
diff
changeset
 | 
400  | 
(Local_Theory.raw_theory ooo Locale.add_dependency locale) dep_morph mixin export  | 
| 
57926
 
59b2572e8e93
load local_theory.ML before attrib.ML, with subtle change of semantics due to canonical Local_Theory.map_contexts instead of private Local_Theory.map_top;
 
wenzelm 
parents: 
57193 
diff
changeset
 | 
401  | 
#> Locale.activate_fragment_nonbrittle dep_morph mixin export;  | 
| 
47279
 
4bab649dedf0
clarified standard_declaration vs. theory_declaration;
 
wenzelm 
parents: 
47276 
diff
changeset
 | 
402  | 
|
| 
60344
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
403  | 
|
| 
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
404  | 
(** locale abbreviations **)  | 
| 
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
405  | 
|
| 
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
406  | 
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
 | 
407  | 
background_abbrev (b, global_rhs) (snd params)  | 
| 
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
408  | 
#-> (fn (lhs, _) => locale_const locale prmode ((b, mx), lhs));  | 
| 
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
409  | 
|
| 
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
410  | 
fun locale_abbrev locale = abbrev (locale_target_abbrev locale);  | 
| 
 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 
haftmann 
parents: 
60342 
diff
changeset
 | 
411  | 
|
| 38309 | 412  | 
end;  |