author | wenzelm |
Wed, 21 Feb 2018 20:13:42 +0100 | |
changeset 67699 | 8e4ff46f807d |
parent 67671 | 857da80611ab |
child 68851 | 6c9825c1e26b |
permissions | -rw-r--r-- |
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38376
diff
changeset
|
1 |
(* Title: Pure/Isar/class_declaration.ML |
24218 | 2 |
Author: Florian Haftmann, TU Muenchen |
3 |
||
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38376
diff
changeset
|
4 |
Declaring classes and subclass relations. |
24218 | 5 |
*) |
6 |
||
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38376
diff
changeset
|
7 |
signature CLASS_DECLARATION = |
24218 | 8 |
sig |
57181
2d13bf9ea77b
dropped obscure and unused ad-hoc before_exit hook for named targets
haftmann
parents:
56921
diff
changeset
|
9 |
val class: binding -> class list -> |
41585
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
wenzelm
parents:
40188
diff
changeset
|
10 |
Element.context_i list -> theory -> string * local_theory |
57181
2d13bf9ea77b
dropped obscure and unused ad-hoc before_exit hook for named targets
haftmann
parents:
56921
diff
changeset
|
11 |
val class_cmd: binding -> xstring list -> |
41585
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
wenzelm
parents:
40188
diff
changeset
|
12 |
Element.context list -> theory -> string * local_theory |
57181
2d13bf9ea77b
dropped obscure and unused ad-hoc before_exit hook for named targets
haftmann
parents:
56921
diff
changeset
|
13 |
val prove_subclass: tactic -> class -> |
41585
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
wenzelm
parents:
40188
diff
changeset
|
14 |
local_theory -> local_theory |
57181
2d13bf9ea77b
dropped obscure and unused ad-hoc before_exit hook for named targets
haftmann
parents:
56921
diff
changeset
|
15 |
val subclass: class -> local_theory -> Proof.state |
2d13bf9ea77b
dropped obscure and unused ad-hoc before_exit hook for named targets
haftmann
parents:
56921
diff
changeset
|
16 |
val subclass_cmd: xstring -> local_theory -> Proof.state |
24218 | 17 |
end; |
18 |
||
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38376
diff
changeset
|
19 |
structure Class_Declaration: CLASS_DECLARATION = |
24218 | 20 |
struct |
21 |
||
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
22 |
(** class definitions **) |
24218 | 23 |
|
24 |
local |
|
25 |
||
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
26 |
(* calculating class-related rules including canonical interpretation *) |
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
27 |
|
29547 | 28 |
fun calculate thy class sups base_sort param_map assm_axiom = |
29 |
let |
|
42360 | 30 |
val empty_ctxt = Proof_Context.init_global thy; |
29547 | 31 |
|
67699
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
wenzelm
parents:
67671
diff
changeset
|
32 |
val certT = Thm.trim_context_ctyp o Thm.global_ctyp_of thy; |
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
wenzelm
parents:
67671
diff
changeset
|
33 |
val cert = Thm.trim_context_cterm o Thm.global_cterm_of thy; |
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
wenzelm
parents:
67671
diff
changeset
|
34 |
|
29547 | 35 |
(* instantiation of canonical interpretation *) |
67699
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
wenzelm
parents:
67671
diff
changeset
|
36 |
val a_tfree = (Name.aT, base_sort); |
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
wenzelm
parents:
67671
diff
changeset
|
37 |
val a_type = TFree a_tfree; |
29627 | 38 |
val param_map_const = (map o apsnd) Const param_map; |
67699
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
wenzelm
parents:
67671
diff
changeset
|
39 |
val param_map_inst = param_map |> map (fn (x, (c, T)) => |
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
wenzelm
parents:
67671
diff
changeset
|
40 |
let val T' = map_atyps (K a_type) T in ((x, T'), cert (Const (c, T'))) end); |
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
wenzelm
parents:
67671
diff
changeset
|
41 |
val const_morph = |
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
wenzelm
parents:
67671
diff
changeset
|
42 |
Element.instantiate_normalize_morphism ([], param_map_inst); |
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
wenzelm
parents:
67671
diff
changeset
|
43 |
val typ_morph = |
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
wenzelm
parents:
67671
diff
changeset
|
44 |
Element.instantiate_normalize_morphism ([(a_tfree, certT (TFree (Name.aT, [class])))], []) |
67450
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
ballarin
parents:
66334
diff
changeset
|
45 |
val (([raw_props], _, [(_, raw_inst_morph)], _, export_morph), _) = empty_ctxt |
29547 | 46 |
|> Expression.cert_goal_expression ([(class, (("", false), |
67450
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
ballarin
parents:
66334
diff
changeset
|
47 |
(Expression.Named param_map_const, [])))], []); |
45431 | 48 |
val (props, inst_morph) = |
49 |
if null param_map |
|
29797 | 50 |
then (raw_props |> map (Morphism.term typ_morph), |
51 |
raw_inst_morph $> typ_morph) |
|
52 |
else (raw_props, raw_inst_morph); (*FIXME proper handling in |
|
53 |
locale.ML / expression.ML would be desirable*) |
|
29547 | 54 |
|
55 |
(* witness for canonical interpretation *) |
|
52636 | 56 |
val some_prop = try the_single props; |
57 |
val some_witn = Option.map (fn prop => |
|
45431 | 58 |
let |
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38376
diff
changeset
|
59 |
val sup_axioms = map_filter (fst o Class.rules thy) sups; |
45431 | 60 |
val loc_intro_tac = |
61 |
(case Locale.intros_of thy class of |
|
62 |
(_, NONE) => all_tac |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58963
diff
changeset
|
63 |
| (_, SOME intro) => ALLGOALS (resolve_tac empty_ctxt [intro])); |
29547 | 64 |
val tac = loc_intro_tac |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
54740
diff
changeset
|
65 |
THEN ALLGOALS (Proof_Context.fact_tac empty_ctxt (sup_axioms @ the_list assm_axiom)); |
52636 | 66 |
in Element.prove_witness empty_ctxt prop tac end) some_prop; |
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54882
diff
changeset
|
67 |
val some_axiom = Option.map (Element.conclude_witness empty_ctxt) some_witn; |
29547 | 68 |
|
69 |
(* canonical interpretation *) |
|
70 |
val base_morph = inst_morph |
|
54740 | 71 |
$> Morphism.binding_morphism "class_binding" (Binding.prefix false (Class.class_prefix class)) |
52636 | 72 |
$> Element.satisfy_morphism (the_list some_witn); |
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38376
diff
changeset
|
73 |
val eq_morph = Element.eq_morphism thy (Class.these_defs thy sups); |
29547 | 74 |
|
75 |
(* assm_intro *) |
|
30344
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
wenzelm
parents:
30335
diff
changeset
|
76 |
fun prove_assm_intro thm = |
29547 | 77 |
let |
31794
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents:
31696
diff
changeset
|
78 |
val ((_, [thm']), _) = Variable.import true [thm] empty_ctxt; |
45431 | 79 |
val const_eq_morph = |
80 |
(case eq_morph of |
|
46856 | 81 |
SOME eq_morph => const_morph $> eq_morph |
45431 | 82 |
| NONE => const_morph); |
36674
d95f39448121
eq_morphism is always optional: avoid trivial morphism for empty list of equations
haftmann
parents:
36672
diff
changeset
|
83 |
val thm'' = Morphism.thm const_eq_morph thm'; |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
54740
diff
changeset
|
84 |
in |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
54740
diff
changeset
|
85 |
Goal.prove_sorry_global thy [] [] (Thm.prop_of thm'') |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
54740
diff
changeset
|
86 |
(fn {context = ctxt, ...} => ALLGOALS (Proof_Context.fact_tac ctxt [thm''])) |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
54740
diff
changeset
|
87 |
end; |
52636 | 88 |
val some_assm_intro = Option.map prove_assm_intro (fst (Locale.intros_of thy class)); |
29547 | 89 |
|
90 |
(* of_class *) |
|
67699
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
wenzelm
parents:
67671
diff
changeset
|
91 |
val of_class_prop_concl = Logic.mk_of_class (a_type, class); |
45431 | 92 |
val of_class_prop = |
52636 | 93 |
(case some_prop of |
45431 | 94 |
NONE => of_class_prop_concl |
29627 | 95 |
| SOME prop => Logic.mk_implies (Morphism.term const_morph |
67699
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
wenzelm
parents:
67671
diff
changeset
|
96 |
((map_types o map_atyps) (K a_type) prop), of_class_prop_concl)); |
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38376
diff
changeset
|
97 |
val sup_of_classes = map (snd o Class.rules thy) sups; |
35021
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
33671
diff
changeset
|
98 |
val loc_axiom_intros = map Drule.export_without_context_open (Locale.axioms_of thy class); |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51551
diff
changeset
|
99 |
val axclass_intro = #intro (Axclass.get_info thy class); |
67699
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
wenzelm
parents:
67671
diff
changeset
|
100 |
val base_sort_trivs = Thm.of_sort (Thm.global_ctyp_of thy a_type, base_sort); |
58957 | 101 |
fun tac ctxt = |
45431 | 102 |
REPEAT (SOMEGOAL |
58957 | 103 |
(match_tac ctxt (axclass_intro :: sup_of_classes @ loc_axiom_intros @ base_sort_trivs) |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58957
diff
changeset
|
104 |
ORELSE' assume_tac ctxt)); |
58957 | 105 |
val of_class = Goal.prove_sorry_global thy [] [] of_class_prop (tac o #context); |
29547 | 106 |
|
52636 | 107 |
in (base_morph, eq_morph, export_morph, some_axiom, some_assm_intro, of_class) end; |
29547 | 108 |
|
29632
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29627
diff
changeset
|
109 |
|
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
110 |
(* reading and processing class specifications *) |
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
111 |
|
36460 | 112 |
fun prep_class_elems prep_decl thy sups raw_elems = |
29632
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29627
diff
changeset
|
113 |
let |
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
114 |
(* user space type system: only permits 'a type variable, improves towards 'a *) |
36460 | 115 |
val algebra = Sign.classes_of thy; |
116 |
val inter_sort = curry (Sorts.inter_sort algebra); |
|
45421 | 117 |
val proto_base_sort = |
118 |
if null sups then Sign.defaultS thy |
|
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38376
diff
changeset
|
119 |
else fold inter_sort (map (Class.base_sort thy) sups) []; |
58294
7f990b3d5189
explicit check phase to guide type inference of class expression towards one single type variable
haftmann
parents:
58293
diff
changeset
|
120 |
val is_param = member (op =) (map fst (Class.these_params thy sups)); |
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
121 |
val base_constraints = (map o apsnd) |
60346 | 122 |
(map_type_tfree (K (TVar ((Name.aT, 0), proto_base_sort))) o fst o fst o snd) |
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38376
diff
changeset
|
123 |
(Class.these_operations thy sups); |
62952 | 124 |
val singleton_fixate = burrow_types (fn Ts => |
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
125 |
let |
58293 | 126 |
val tfrees = fold Term.add_tfreesT Ts []; |
45433
4283f3a57cf5
avoid separate typ_check phases, integrate into main term_check 0 instead (cf. its Syntax.check_typs in Type_Infer_Context.prepare);
wenzelm
parents:
45432
diff
changeset
|
127 |
val inferred_sort = |
58293 | 128 |
(fold o fold_atyps) (fn TVar (_, S) => inter_sort S | _ => I) Ts []; |
45421 | 129 |
val fixate_sort = |
45432
12cc89f1eb0c
clarified singleton_fixate: intersection with supersort is identity, only replace actual type inference parameters;
wenzelm
parents:
45431
diff
changeset
|
130 |
(case tfrees of |
12cc89f1eb0c
clarified singleton_fixate: intersection with supersort is identity, only replace actual type inference parameters;
wenzelm
parents:
45431
diff
changeset
|
131 |
[] => inferred_sort |
45433
4283f3a57cf5
avoid separate typ_check phases, integrate into main term_check 0 instead (cf. its Syntax.check_typs in Type_Infer_Context.prepare);
wenzelm
parents:
45432
diff
changeset
|
132 |
| [(a, S)] => |
4283f3a57cf5
avoid separate typ_check phases, integrate into main term_check 0 instead (cf. its Syntax.check_typs in Type_Infer_Context.prepare);
wenzelm
parents:
45432
diff
changeset
|
133 |
if a <> Name.aT then |
4283f3a57cf5
avoid separate typ_check phases, integrate into main term_check 0 instead (cf. its Syntax.check_typs in Type_Infer_Context.prepare);
wenzelm
parents:
45432
diff
changeset
|
134 |
error ("No type variable other than " ^ Name.aT ^ " allowed in class specification") |
4283f3a57cf5
avoid separate typ_check phases, integrate into main term_check 0 instead (cf. its Syntax.check_typs in Type_Infer_Context.prepare);
wenzelm
parents:
45432
diff
changeset
|
135 |
else if Sorts.sort_le algebra (S, inferred_sort) then S |
45432
12cc89f1eb0c
clarified singleton_fixate: intersection with supersort is identity, only replace actual type inference parameters;
wenzelm
parents:
45431
diff
changeset
|
136 |
else |
12cc89f1eb0c
clarified singleton_fixate: intersection with supersort is identity, only replace actual type inference parameters;
wenzelm
parents:
45431
diff
changeset
|
137 |
error ("Type inference imposes additional sort constraint " ^ |
12cc89f1eb0c
clarified singleton_fixate: intersection with supersort is identity, only replace actual type inference parameters;
wenzelm
parents:
45431
diff
changeset
|
138 |
Syntax.string_of_sort_global thy inferred_sort ^ |
12cc89f1eb0c
clarified singleton_fixate: intersection with supersort is identity, only replace actual type inference parameters;
wenzelm
parents:
45431
diff
changeset
|
139 |
" of type parameter " ^ Name.aT ^ " of sort " ^ |
12cc89f1eb0c
clarified singleton_fixate: intersection with supersort is identity, only replace actual type inference parameters;
wenzelm
parents:
45431
diff
changeset
|
140 |
Syntax.string_of_sort_global thy S) |
12cc89f1eb0c
clarified singleton_fixate: intersection with supersort is identity, only replace actual type inference parameters;
wenzelm
parents:
45431
diff
changeset
|
141 |
| _ => error "Multiple type variables in class specification"); |
12cc89f1eb0c
clarified singleton_fixate: intersection with supersort is identity, only replace actual type inference parameters;
wenzelm
parents:
45431
diff
changeset
|
142 |
val fixateT = TFree (Name.aT, fixate_sort); |
12cc89f1eb0c
clarified singleton_fixate: intersection with supersort is identity, only replace actual type inference parameters;
wenzelm
parents:
45431
diff
changeset
|
143 |
in |
58293 | 144 |
(map o map_atyps) |
145 |
(fn T as TVar (xi, _) => if Type_Infer.is_param xi then fixateT else T | T => T) Ts |
|
62952 | 146 |
end); |
147 |
fun unify_params ts = |
|
40188
eddda8e38360
consider base sorts of superclasses simultaneously during processing of class specification -- avoids clash of different base sorts
haftmann
parents:
38875
diff
changeset
|
148 |
let |
58294
7f990b3d5189
explicit check phase to guide type inference of class expression towards one single type variable
haftmann
parents:
58293
diff
changeset
|
149 |
val param_Ts = (fold o fold_aterms) |
7f990b3d5189
explicit check phase to guide type inference of class expression towards one single type variable
haftmann
parents:
58293
diff
changeset
|
150 |
(fn Free (v, T) => if is_param v then fold_atyps (insert (op =)) T else I | _ => I) ts []; |
58319 | 151 |
val param_namesT = map_filter (try (fst o dest_TVar)) param_Ts; |
152 |
val param_T = if null param_namesT then NONE |
|
58294
7f990b3d5189
explicit check phase to guide type inference of class expression towards one single type variable
haftmann
parents:
58293
diff
changeset
|
153 |
else SOME (case get_first (try dest_TFree) param_Ts of |
7f990b3d5189
explicit check phase to guide type inference of class expression towards one single type variable
haftmann
parents:
58293
diff
changeset
|
154 |
SOME v_sort => TFree v_sort | |
58319 | 155 |
NONE => TVar (hd param_namesT, proto_base_sort)); |
58294
7f990b3d5189
explicit check phase to guide type inference of class expression towards one single type variable
haftmann
parents:
58293
diff
changeset
|
156 |
in case param_T of |
7f990b3d5189
explicit check phase to guide type inference of class expression towards one single type variable
haftmann
parents:
58293
diff
changeset
|
157 |
NONE => ts | |
58319 | 158 |
SOME T => map (subst_TVars (map (rpair T) param_namesT)) ts |
40188
eddda8e38360
consider base sorts of superclasses simultaneously during processing of class specification -- avoids clash of different base sorts
haftmann
parents:
38875
diff
changeset
|
159 |
end; |
29632
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29627
diff
changeset
|
160 |
|
35120 | 161 |
(* preprocessing elements, retrieving base sort from type-checked elements *) |
42402
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42375
diff
changeset
|
162 |
val raw_supexpr = |
67450
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
ballarin
parents:
66334
diff
changeset
|
163 |
(map (fn sup => (sup, (("", false), (Expression.Positional [], [])))) sups, []); |
42402
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42375
diff
changeset
|
164 |
val init_class_body = |
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42375
diff
changeset
|
165 |
fold (Proof_Context.add_const_constraint o apsnd SOME) base_constraints |
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38376
diff
changeset
|
166 |
#> Class.redeclare_operations thy sups |
62952 | 167 |
#> Context.proof_map (Syntax_Phases.term_check 0 "singleton_fixate" (K singleton_fixate)); |
47311
1addbe2a7458
close context elements via Expression.cert/read_declaration;
wenzelm
parents:
46922
diff
changeset
|
168 |
val ((raw_supparams, _, raw_inferred_elems, _), _) = |
42402
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42375
diff
changeset
|
169 |
Proof_Context.init_global thy |
62952 | 170 |
|> Context.proof_map (Syntax_Phases.term_check 0 "unify_params" (K unify_params)) |
32206 | 171 |
|> prep_decl raw_supexpr init_class_body raw_elems; |
38435 | 172 |
fun filter_element (Element.Fixes []) = NONE |
173 |
| filter_element (e as Element.Fixes _) = SOME e |
|
174 |
| filter_element (Element.Constrains []) = NONE |
|
175 |
| filter_element (e as Element.Constrains _) = SOME e |
|
176 |
| filter_element (Element.Assumes []) = NONE |
|
177 |
| filter_element (e as Element.Assumes _) = SOME e |
|
45421 | 178 |
| filter_element (Element.Defines _) = |
179 |
error ("\"defines\" element not allowed in class specification.") |
|
180 |
| filter_element (Element.Notes _) = |
|
67671
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
wenzelm
parents:
67450
diff
changeset
|
181 |
error ("\"notes\" element not allowed in class specification.") |
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
wenzelm
parents:
67450
diff
changeset
|
182 |
| filter_element (Element.Lazy_Notes _) = |
45421 | 183 |
error ("\"notes\" element not allowed in class specification."); |
38435 | 184 |
val inferred_elems = map_filter filter_element raw_inferred_elems; |
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
185 |
fun fold_element_types f (Element.Fixes fxs) = fold (fn (_, SOME T, _) => f T) fxs |
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
186 |
| fold_element_types f (Element.Constrains cnstrs) = fold (f o snd) cnstrs |
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
187 |
| fold_element_types f (Element.Assumes assms) = fold (fold (fn (t, ts) => |
42402
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42375
diff
changeset
|
188 |
fold_types f t #> (fold o fold_types) f ts) o snd) assms; |
45421 | 189 |
val base_sort = |
190 |
if null inferred_elems then proto_base_sort |
|
191 |
else |
|
192 |
(case (fold o fold_element_types) Term.add_tfreesT inferred_elems [] of |
|
193 |
[] => error "No type variable in class specification" |
|
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
194 |
| [(_, sort)] => sort |
45421 | 195 |
| _ => error "Multiple type variables in class specification"); |
32206 | 196 |
val supparams = map (fn ((c, T), _) => |
197 |
(c, map_atyps (K (TFree (Name.aT, base_sort))) T)) raw_supparams; |
|
198 |
val supparam_names = map fst supparams; |
|
199 |
fun mk_param ((c, _), _) = Free (c, (the o AList.lookup (op =) supparams) c); |
|
200 |
val supexpr = (map (fn sup => (sup, (("", false), |
|
67450
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
ballarin
parents:
66334
diff
changeset
|
201 |
(Expression.Positional (map (SOME o mk_param) (Locale.params_of thy sup)), [])))) sups, |
32206 | 202 |
map (fn (c, T) => (Binding.name c, SOME T, NoSyn)) supparams); |
29575
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
203 |
|
32206 | 204 |
in (base_sort, supparam_names, supexpr, inferred_elems) end; |
29575
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
205 |
|
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
206 |
val cert_class_elems = prep_class_elems Expression.cert_declaration; |
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
207 |
val read_class_elems = prep_class_elems Expression.cert_read_declaration; |
29575
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
208 |
|
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
209 |
fun prep_class_spec prep_class prep_class_elems thy raw_supclasses raw_elems = |
29575
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
210 |
let |
46922
3717f3878714
source positions for locale and class expressions;
wenzelm
parents:
46856
diff
changeset
|
211 |
val thy_ctxt = Proof_Context.init_global thy; |
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
212 |
|
29575
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
213 |
(* prepare import *) |
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
214 |
val inter_sort = curry (Sorts.inter_sort (Sign.classes_of thy)); |
46922
3717f3878714
source positions for locale and class expressions;
wenzelm
parents:
46856
diff
changeset
|
215 |
val sups = Sign.minimize_sort thy (map (prep_class thy_ctxt) raw_supclasses); |
45431 | 216 |
val _ = |
217 |
(case filter_out (Class.is_class thy) sups of |
|
218 |
[] => () |
|
219 |
| no_classes => error ("No (proper) classes: " ^ commas_quote no_classes)); |
|
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38376
diff
changeset
|
220 |
val raw_supparams = (map o apsnd) (snd o snd) (Class.these_params thy sups); |
32206 | 221 |
val raw_supparam_names = map fst raw_supparams; |
45431 | 222 |
val _ = |
223 |
if has_duplicates (op =) raw_supparam_names then |
|
224 |
error ("Duplicate parameter(s) in superclasses: " ^ |
|
225 |
(commas_quote (duplicates (op =) raw_supparam_names))) |
|
29575
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
226 |
else (); |
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
227 |
|
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
228 |
(* infer types and base sort *) |
45431 | 229 |
val (base_sort, supparam_names, supexpr, inferred_elems) = prep_class_elems thy sups raw_elems; |
32206 | 230 |
val sup_sort = inter_sort base_sort sups; |
29575
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
231 |
|
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
232 |
(* process elements as class specification *) |
46922
3717f3878714
source positions for locale and class expressions;
wenzelm
parents:
46856
diff
changeset
|
233 |
val class_ctxt = Class.begin sups base_sort thy_ctxt; |
47311
1addbe2a7458
close context elements via Expression.cert/read_declaration;
wenzelm
parents:
46922
diff
changeset
|
234 |
val ((_, _, syntax_elems, _), _) = class_ctxt |
29702 | 235 |
|> Expression.cert_declaration supexpr I inferred_elems; |
45431 | 236 |
fun check_vars e vs = |
237 |
if null vs then |
|
238 |
error ("No type variable in part of specification element " ^ |
|
239 |
Pretty.string_of (Pretty.chunks (Element.pretty_ctxt class_ctxt e))) |
|
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
240 |
else (); |
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
241 |
fun check_element (e as Element.Fixes fxs) = |
45431 | 242 |
List.app (fn (_, SOME T, _) => check_vars e (Term.add_tfreesT T [])) fxs |
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
243 |
| check_element (e as Element.Assumes assms) = |
45431 | 244 |
List.app (fn (_, ts_pss) => |
245 |
List.app (fn (t, _) => check_vars e (Term.add_tfrees t [])) ts_pss) assms |
|
246 |
| check_element _ = (); |
|
247 |
val _ = List.app check_element syntax_elems; |
|
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
248 |
fun fork_syn (Element.Fixes xs) = |
30344
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
wenzelm
parents:
30335
diff
changeset
|
249 |
fold_map (fn (c, ty, syn) => cons (c, syn) #> pair (c, ty, NoSyn)) xs |
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
250 |
#>> Element.Fixes |
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
251 |
| fork_syn x = pair x; |
29575
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
252 |
val (elems, global_syntax) = fold_map fork_syn syntax_elems []; |
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
253 |
|
32713 | 254 |
in (((sups, supparam_names), (sup_sort, base_sort, supexpr)), (elems, global_syntax)) end; |
29575
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
255 |
|
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
256 |
val cert_class_spec = prep_class_spec (K I) cert_class_elems; |
46922
3717f3878714
source positions for locale and class expressions;
wenzelm
parents:
46856
diff
changeset
|
257 |
val read_class_spec = prep_class_spec Proof_Context.read_class read_class_elems; |
29665
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
258 |
|
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
259 |
|
2b956243d123
explicit check for exactly one type variable in class specification elements
haftmann
parents:
29632
diff
changeset
|
260 |
(* class establishment *) |
29575
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
261 |
|
32206 | 262 |
fun add_consts class base_sort sups supparam_names global_syntax thy = |
24968
f9bafc868847
replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents:
24949
diff
changeset
|
263 |
let |
29816 | 264 |
(*FIXME simplify*) |
32206 | 265 |
val supconsts = supparam_names |
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38376
diff
changeset
|
266 |
|> AList.make (snd o the o AList.lookup (op =) (Class.these_params thy sups)) |
25683 | 267 |
|> (map o apsnd o apsnd o map_atyps o K o TFree) (Name.aT, [class]); |
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
268 |
val all_params = Locale.params_of thy class; |
32206 | 269 |
val raw_params = (snd o chop (length supparam_names)) all_params; |
30755
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
wenzelm
parents:
30585
diff
changeset
|
270 |
fun add_const ((raw_c, raw_ty), _) thy = |
25683 | 271 |
let |
30755
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
wenzelm
parents:
30585
diff
changeset
|
272 |
val b = Binding.name raw_c; |
30344
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
wenzelm
parents:
30335
diff
changeset
|
273 |
val c = Sign.full_name thy b; |
25683 | 274 |
val ty = map_atyps (K (TFree (Name.aT, base_sort))) raw_ty; |
275 |
val ty0 = Type.strip_sorts ty; |
|
276 |
val ty' = map_atyps (K (TFree (Name.aT, [class]))) ty0; |
|
30344
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
wenzelm
parents:
30335
diff
changeset
|
277 |
val syn = (the_default NoSyn o AList.lookup Binding.eq_name global_syntax) b; |
25683 | 278 |
in |
279 |
thy |
|
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42360
diff
changeset
|
280 |
|> Sign.declare_const_global ((b, ty0), syn) |
25683 | 281 |
|> snd |
42494 | 282 |
|> pair ((Variable.check_name b, ty), (c, ty')) |
25683 | 283 |
end; |
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
284 |
in |
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
285 |
thy |
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38376
diff
changeset
|
286 |
|> Sign.add_path (Class.class_prefix class) |
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
287 |
|> fold_map add_const raw_params |
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
288 |
||> Sign.restore_naming thy |
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
289 |
|-> (fn params => pair (supconsts @ (map o apfst) fst params, params)) |
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
290 |
end; |
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
291 |
|
32206 | 292 |
fun adjungate_axclass bname class base_sort sups supsort supparam_names global_syntax thy = |
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
293 |
let |
29816 | 294 |
(*FIXME simplify*) |
25683 | 295 |
fun globalize param_map = map_aterms |
296 |
(fn Free (v, ty) => Const ((fst o the o AList.lookup (op =) param_map) v, ty) |
|
297 |
| t => t); |
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
298 |
val raw_pred = Locale.intros_of thy class |
25683 | 299 |
|> fst |
35845
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents:
35669
diff
changeset
|
300 |
|> Option.map (Logic.unvarify_global o Logic.strip_imp_concl o Thm.prop_of); |
45431 | 301 |
fun get_axiom thy = |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51551
diff
changeset
|
302 |
(case #axioms (Axclass.get_info thy class) of |
45431 | 303 |
[] => NONE |
304 |
| [thm] => SOME thm); |
|
24968
f9bafc868847
replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents:
24949
diff
changeset
|
305 |
in |
f9bafc868847
replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents:
24949
diff
changeset
|
306 |
thy |
63352 | 307 |
|> add_consts class base_sort sups supparam_names global_syntax |-> (fn (param_map, params) => |
308 |
Axclass.define_class (bname, supsort) |
|
309 |
(map (fst o snd) params) |
|
310 |
[(Binding.empty_atts, Option.map (globalize param_map) raw_pred |> the_list)] |
|
311 |
#> snd |
|
312 |
#> `get_axiom |
|
313 |
#-> (fn assm_axiom => fold (Sign.add_const_constraint o apsnd SOME o snd) params |
|
314 |
#> pair (param_map, params, assm_axiom))) |
|
24968
f9bafc868847
replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents:
24949
diff
changeset
|
315 |
end; |
f9bafc868847
replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents:
24949
diff
changeset
|
316 |
|
57181
2d13bf9ea77b
dropped obscure and unused ad-hoc before_exit hook for named targets
haftmann
parents:
56921
diff
changeset
|
317 |
fun gen_class prep_class_spec b raw_supclasses raw_elems thy = |
24748 | 318 |
let |
36635
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
haftmann
parents:
36464
diff
changeset
|
319 |
val class = Sign.full_name thy b; |
56921
5bf71b4da706
note of_class rule for type classes in theory: useful to promote class instance proofs to locale interpretation proofs
haftmann
parents:
54883
diff
changeset
|
320 |
val prefix = Binding.qualify true "class"; |
32206 | 321 |
val (((sups, supparam_names), (supsort, base_sort, supexpr)), (elems, global_syntax)) = |
322 |
prep_class_spec thy raw_supclasses raw_elems; |
|
24218 | 323 |
in |
324 |
thy |
|
57181
2d13bf9ea77b
dropped obscure and unused ad-hoc before_exit hook for named targets
haftmann
parents:
56921
diff
changeset
|
325 |
|> Expression.add_locale b (prefix b) supexpr elems |
33671 | 326 |
|> snd |> Local_Theory.exit_global |
36635
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
haftmann
parents:
36464
diff
changeset
|
327 |
|> adjungate_axclass b class base_sort sups supsort supparam_names global_syntax |
29526 | 328 |
|-> (fn (param_map, params, assm_axiom) => |
29547 | 329 |
`(fn thy => calculate thy class sups base_sort param_map assm_axiom) |
52636 | 330 |
#-> (fn (base_morph, eq_morph, export_morph, some_axiom, some_assm_intro, of_class) => |
38107 | 331 |
Context.theory_map (Locale.add_registration (class, base_morph) |
332 |
(Option.map (rpair true) eq_morph) export_morph) |
|
56921
5bf71b4da706
note of_class rule for type classes in theory: useful to promote class instance proofs to locale interpretation proofs
haftmann
parents:
54883
diff
changeset
|
333 |
#> Class.register class sups params base_sort base_morph export_morph some_axiom some_assm_intro of_class |
5bf71b4da706
note of_class rule for type classes in theory: useful to promote class instance proofs to locale interpretation proofs
haftmann
parents:
54883
diff
changeset
|
334 |
#> Global_Theory.store_thm (prefix (Binding.qualified_name (class ^ ".of_class.intro")), of_class))) |
5bf71b4da706
note of_class rule for type classes in theory: useful to promote class instance proofs to locale interpretation proofs
haftmann
parents:
54883
diff
changeset
|
335 |
|> snd |
66334
b210ae666a42
provide explicit variant initializers for regular named target vs. almost-named target
haftmann
parents:
63402
diff
changeset
|
336 |
|> Named_Target.init class |
25038 | 337 |
|> pair class |
24218 | 338 |
end; |
339 |
||
340 |
in |
|
341 |
||
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
342 |
val class = gen_class cert_class_spec; |
26518 | 343 |
val class_cmd = gen_class read_class_spec; |
24218 | 344 |
|
345 |
end; (*local*) |
|
346 |
||
347 |
||
45421 | 348 |
|
29358 | 349 |
(** subclass relations **) |
25462 | 350 |
|
29358 | 351 |
local |
25462 | 352 |
|
57181
2d13bf9ea77b
dropped obscure and unused ad-hoc before_exit hook for named targets
haftmann
parents:
56921
diff
changeset
|
353 |
fun gen_subclass prep_class do_proof raw_sup lthy = |
25462 | 354 |
let |
42360 | 355 |
val thy = Proof_Context.theory_of lthy; |
29558 | 356 |
val proto_sup = prep_class thy raw_sup; |
63268 | 357 |
val proto_sub = |
358 |
(case Named_Target.class_of lthy of |
|
57182 | 359 |
SOME class => class |
63268 | 360 |
| NONE => error "Not in a class target"); |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51551
diff
changeset
|
361 |
val (sub, sup) = Axclass.cert_classrel thy (proto_sub, proto_sup); |
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
362 |
|
67450
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
ballarin
parents:
66334
diff
changeset
|
363 |
val expr = ([(sup, (("", false), (Expression.Positional [], [])))], []); |
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
ballarin
parents:
66334
diff
changeset
|
364 |
val (([props], _, deps, _, export), goal_ctxt) = |
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
365 |
Expression.cert_goal_expression expr lthy; |
29526 | 366 |
val some_prop = try the_single props; |
29558 | 367 |
val some_dep_morph = try the_single (map snd deps); |
368 |
fun after_qed some_wit = |
|
54866
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
369 |
Class.register_subclass (sub, sup) some_dep_morph some_wit export; |
29558 | 370 |
in do_proof after_qed some_prop goal_ctxt end; |
25485 | 371 |
|
29575
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
372 |
fun user_proof after_qed some_prop = |
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
373 |
Element.witness_proof (after_qed o try the_single o the_single) |
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
374 |
[the_list some_prop]; |
25485 | 375 |
|
29575
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
376 |
fun tactic_proof tac after_qed some_prop ctxt = |
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
377 |
after_qed (Option.map |
41d604e59e93
improved and corrected reading of class specs -- still draft version
haftmann
parents:
29558
diff
changeset
|
378 |
(fn prop => Element.prove_witness ctxt prop tac) some_prop) ctxt; |
28666 | 379 |
|
29358 | 380 |
in |
28666 | 381 |
|
57181
2d13bf9ea77b
dropped obscure and unused ad-hoc before_exit hook for named targets
haftmann
parents:
56921
diff
changeset
|
382 |
fun prove_subclass tac = gen_subclass (K I) (tactic_proof tac); |
54882 | 383 |
|
384 |
fun subclass x = gen_subclass (K I) user_proof x; |
|
385 |
fun subclass_cmd x = |
|
386 |
gen_subclass (Proof_Context.read_class o Proof_Context.init_global) user_proof x; |
|
25462 | 387 |
|
29358 | 388 |
end; (*local*) |
389 |
||
24218 | 390 |
end; |