| author | wenzelm | 
| Tue, 09 Jan 2024 17:38:50 +0100 | |
| changeset 79454 | 6b6e9af552f5 | 
| parent 79438 | 032ca41f590a | 
| child 81116 | 0fb1e2dd4122 | 
| 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  | 
| 72605 | 9  | 
val class: binding -> Bundle.name list -> 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  | 
| 
72536
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
11  | 
val class_cmd: binding -> (xstring * Position.T) list -> 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  | 
|
| 70309 | 30  | 
val thy_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;  | 
| 78041 | 34  | 
val conclude_witness = Thm.trim_context o Element.conclude_witness thy_ctxt;  | 
| 
67699
 
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
35  | 
|
| 29547 | 36  | 
(* 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
 | 
37  | 
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
 | 
38  | 
val a_type = TFree a_tfree;  | 
| 29627 | 39  | 
val param_map_const = (map o apsnd) Const param_map;  | 
| 74282 | 40  | 
val param_map_inst =  | 
41  | 
Frees.build (param_map |> fold (fn (x, (c, T)) =>  | 
|
42  | 
let val T' = map_atyps (K a_type) T  | 
|
43  | 
in Frees.add ((x, T'), cert (Const (c, T'))) end));  | 
|
| 
67699
 
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
44  | 
val const_morph =  | 
| 74282 | 45  | 
Element.instantiate_normalize_morphism (TFrees.empty, param_map_inst);  | 
| 
67699
 
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
46  | 
val typ_morph =  | 
| 74282 | 47  | 
Element.instantiate_normalize_morphism  | 
| 77879 | 48  | 
(TFrees.make1 (a_tfree, certT (Term.aT [class])), Frees.empty)  | 
| 70309 | 49  | 
val (([raw_props], _, [(_, raw_inst_morph)], _, export_morph), _) = thy_ctxt  | 
| 29547 | 50  | 
      |> Expression.cert_goal_expression ([(class, (("", false),
 | 
| 
67450
 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 
ballarin 
parents: 
66334 
diff
changeset
 | 
51  | 
(Expression.Named param_map_const, [])))], []);  | 
| 45431 | 52  | 
val (props, inst_morph) =  | 
53  | 
if null param_map  | 
|
| 29797 | 54  | 
then (raw_props |> map (Morphism.term typ_morph),  | 
55  | 
raw_inst_morph $> typ_morph)  | 
|
56  | 
else (raw_props, raw_inst_morph); (*FIXME proper handling in  | 
|
57  | 
locale.ML / expression.ML would be desirable*)  | 
|
| 29547 | 58  | 
|
59  | 
(* witness for canonical interpretation *)  | 
|
| 52636 | 60  | 
val some_prop = try the_single props;  | 
61  | 
val some_witn = Option.map (fn prop =>  | 
|
| 45431 | 62  | 
let  | 
| 
38379
 
67d71449e85b
more convenient split of class modules: class and class_declaration
 
haftmann 
parents: 
38376 
diff
changeset
 | 
63  | 
val sup_axioms = map_filter (fst o Class.rules thy) sups;  | 
| 45431 | 64  | 
val loc_intro_tac =  | 
65  | 
(case Locale.intros_of thy class of  | 
|
66  | 
(_, NONE) => all_tac  | 
|
| 70309 | 67  | 
| (_, SOME intro) => ALLGOALS (resolve_tac thy_ctxt [intro]));  | 
| 29547 | 68  | 
val tac = loc_intro_tac  | 
| 70309 | 69  | 
THEN ALLGOALS (Proof_Context.fact_tac thy_ctxt (sup_axioms @ the_list assm_axiom));  | 
70  | 
in Element.prove_witness thy_ctxt prop tac end) some_prop;  | 
|
| 78041 | 71  | 
val some_axiom = Option.map conclude_witness some_witn;  | 
| 29547 | 72  | 
|
73  | 
(* canonical interpretation *)  | 
|
74  | 
val base_morph = inst_morph  | 
|
| 54740 | 75  | 
$> Morphism.binding_morphism "class_binding" (Binding.prefix false (Class.class_prefix class))  | 
| 52636 | 76  | 
$> Element.satisfy_morphism (the_list some_witn);  | 
| 
78065
 
11d6a1e62841
more careful treatment of set_context / reset_context for persistent morphisms;
 
wenzelm 
parents: 
78041 
diff
changeset
 | 
77  | 
val eq_morph =  | 
| 
78453
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
78  | 
Element.eq_morphism thy_ctxt (Class.these_defs thy sups)  | 
| 
78065
 
11d6a1e62841
more careful treatment of set_context / reset_context for persistent morphisms;
 
wenzelm 
parents: 
78041 
diff
changeset
 | 
79  | 
|> Option.map (Morphism.set_context thy);  | 
| 29547 | 80  | 
|
81  | 
(* assm_intro *)  | 
|
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30335 
diff
changeset
 | 
82  | 
fun prove_assm_intro thm =  | 
| 29547 | 83  | 
let  | 
| 70309 | 84  | 
val ((_, [thm']), _) = Variable.import true [thm] thy_ctxt;  | 
| 45431 | 85  | 
val const_eq_morph =  | 
86  | 
(case eq_morph of  | 
|
| 46856 | 87  | 
SOME eq_morph => const_morph $> eq_morph  | 
| 45431 | 88  | 
| NONE => const_morph);  | 
| 
36674
 
d95f39448121
eq_morphism is always optional: avoid trivial morphism for empty list of equations
 
haftmann 
parents: 
36672 
diff
changeset
 | 
89  | 
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
 | 
90  | 
in  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54740 
diff
changeset
 | 
91  | 
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
 | 
92  | 
          (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
 | 
93  | 
end;  | 
| 52636 | 94  | 
val some_assm_intro = Option.map prove_assm_intro (fst (Locale.intros_of thy class));  | 
| 29547 | 95  | 
|
96  | 
(* of_class *)  | 
|
| 
67699
 
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
97  | 
val of_class_prop_concl = Logic.mk_of_class (a_type, class);  | 
| 45431 | 98  | 
val of_class_prop =  | 
| 52636 | 99  | 
(case some_prop of  | 
| 45431 | 100  | 
NONE => of_class_prop_concl  | 
| 29627 | 101  | 
| SOME prop => Logic.mk_implies (Morphism.term const_morph  | 
| 79438 | 102  | 
((map_types o Term.map_atyps_same) (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
 | 
103  | 
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
 | 
104  | 
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
 | 
105  | 
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
 | 
106  | 
val base_sort_trivs = Thm.of_sort (Thm.global_ctyp_of thy a_type, base_sort);  | 
| 58957 | 107  | 
fun tac ctxt =  | 
| 45431 | 108  | 
REPEAT (SOMEGOAL  | 
| 58957 | 109  | 
(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
 | 
110  | 
ORELSE' assume_tac ctxt));  | 
| 
71018
 
d32ed8927a42
more robust expose_proofs corresponding to register_proofs/consolidate_theory;
 
wenzelm 
parents: 
71017 
diff
changeset
 | 
111  | 
val of_class = Goal.prove_sorry_global thy [] [] of_class_prop (tac o #context);  | 
| 52636 | 112  | 
in (base_morph, eq_morph, export_morph, some_axiom, some_assm_intro, of_class) end;  | 
| 29547 | 113  | 
|
| 
29632
 
c3d576157244
fixed reading of class specs: declare class operations in context
 
haftmann 
parents: 
29627 
diff
changeset
 | 
114  | 
|
| 
29665
 
2b956243d123
explicit check for exactly one type variable in class specification elements
 
haftmann 
parents: 
29632 
diff
changeset
 | 
115  | 
(* reading and processing class specifications *)  | 
| 
 
2b956243d123
explicit check for exactly one type variable in class specification elements
 
haftmann 
parents: 
29632 
diff
changeset
 | 
116  | 
|
| 
72536
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
117  | 
fun prep_class_elems prep_decl ctxt sups raw_elems =  | 
| 
29632
 
c3d576157244
fixed reading of class specs: declare class operations in context
 
haftmann 
parents: 
29627 
diff
changeset
 | 
118  | 
let  | 
| 
29665
 
2b956243d123
explicit check for exactly one type variable in class specification elements
 
haftmann 
parents: 
29632 
diff
changeset
 | 
119  | 
(* user space type system: only permits 'a type variable, improves towards 'a *)  | 
| 
72536
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
120  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 36460 | 121  | 
val algebra = Sign.classes_of thy;  | 
122  | 
val inter_sort = curry (Sorts.inter_sort algebra);  | 
|
| 45421 | 123  | 
val proto_base_sort =  | 
124  | 
if null sups then Sign.defaultS thy  | 
|
| 
38379
 
67d71449e85b
more convenient split of class modules: class and class_declaration
 
haftmann 
parents: 
38376 
diff
changeset
 | 
125  | 
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
 | 
126  | 
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
 | 
127  | 
val base_constraints = (map o apsnd)  | 
| 60346 | 128  | 
(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
 | 
129  | 
(Class.these_operations thy sups);  | 
| 62952 | 130  | 
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
 | 
131  | 
let  | 
| 58293 | 132  | 
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
 | 
133  | 
val inferred_sort =  | 
| 58293 | 134  | 
(fold o fold_atyps) (fn TVar (_, S) => inter_sort S | _ => I) Ts [];  | 
| 45421 | 135  | 
val fixate_sort =  | 
| 
45432
 
12cc89f1eb0c
clarified singleton_fixate: intersection with supersort is identity, only replace actual type inference parameters;
 
wenzelm 
parents: 
45431 
diff
changeset
 | 
136  | 
(case tfrees of  | 
| 
 
12cc89f1eb0c
clarified singleton_fixate: intersection with supersort is identity, only replace actual type inference parameters;
 
wenzelm 
parents: 
45431 
diff
changeset
 | 
137  | 
[] => 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
 | 
138  | 
| [(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
 | 
139  | 
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
 | 
140  | 
                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
 | 
141  | 
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
 | 
142  | 
else  | 
| 
 
12cc89f1eb0c
clarified singleton_fixate: intersection with supersort is identity, only replace actual type inference parameters;
 
wenzelm 
parents: 
45431 
diff
changeset
 | 
143  | 
                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
 | 
144  | 
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
 | 
145  | 
" 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
 | 
146  | 
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
 | 
147  | 
| _ => error "Multiple type variables in class specification");  | 
| 70387 | 148  | 
val fixateT = Term.aT fixate_sort;  | 
| 
45432
 
12cc89f1eb0c
clarified singleton_fixate: intersection with supersort is identity, only replace actual type inference parameters;
 
wenzelm 
parents: 
45431 
diff
changeset
 | 
149  | 
in  | 
| 79438 | 150  | 
(Same.commit o Same.map o Term.map_atyps_same)  | 
151  | 
(fn TVar (xi, _) => if Type_Infer.is_param xi then fixateT else raise Same.SAME  | 
|
152  | 
| _ => raise Same.SAME) Ts  | 
|
| 62952 | 153  | 
end);  | 
154  | 
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
 | 
155  | 
let  | 
| 
58294
 
7f990b3d5189
explicit check phase to guide type inference of class expression towards one single type variable
 
haftmann 
parents: 
58293 
diff
changeset
 | 
156  | 
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
 | 
157  | 
(fn Free (v, T) => if is_param v then fold_atyps (insert (op =)) T else I | _ => I) ts [];  | 
| 58319 | 158  | 
val param_namesT = map_filter (try (fst o dest_TVar)) param_Ts;  | 
159  | 
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
 | 
160  | 
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
 | 
161  | 
SOME v_sort => TFree v_sort |  | 
| 58319 | 162  | 
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
 | 
163  | 
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
 | 
164  | 
NONE => ts |  | 
| 58319 | 165  | 
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
 | 
166  | 
end;  | 
| 
29632
 
c3d576157244
fixed reading of class specs: declare class operations in context
 
haftmann 
parents: 
29627 
diff
changeset
 | 
167  | 
|
| 35120 | 168  | 
(* 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
 | 
169  | 
val raw_supexpr =  | 
| 
67450
 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 
ballarin 
parents: 
66334 
diff
changeset
 | 
170  | 
      (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
 | 
171  | 
val init_class_body =  | 
| 
 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 
wenzelm 
parents: 
42375 
diff
changeset
 | 
172  | 
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
 | 
173  | 
#> Class.redeclare_operations thy sups  | 
| 62952 | 174  | 
#> 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
 | 
175  | 
val ((raw_supparams, _, raw_inferred_elems, _), _) =  | 
| 
72536
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
176  | 
ctxt  | 
| 62952 | 177  | 
|> Context.proof_map (Syntax_Phases.term_check 0 "unify_params" (K unify_params))  | 
| 32206 | 178  | 
|> prep_decl raw_supexpr init_class_body raw_elems;  | 
| 38435 | 179  | 
fun filter_element (Element.Fixes []) = NONE  | 
180  | 
| filter_element (e as Element.Fixes _) = SOME e  | 
|
181  | 
| filter_element (Element.Constrains []) = NONE  | 
|
182  | 
| filter_element (e as Element.Constrains _) = SOME e  | 
|
183  | 
| filter_element (Element.Assumes []) = NONE  | 
|
184  | 
| filter_element (e as Element.Assumes _) = SOME e  | 
|
| 45421 | 185  | 
| filter_element (Element.Defines _) =  | 
186  | 
          error ("\"defines\" element not allowed in class specification.")
 | 
|
187  | 
| 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
 | 
188  | 
          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
 | 
189  | 
| filter_element (Element.Lazy_Notes _) =  | 
| 45421 | 190  | 
          error ("\"notes\" element not allowed in class specification.");
 | 
| 38435 | 191  | 
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
 | 
192  | 
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
 | 
193  | 
| 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
 | 
194  | 
| 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
 | 
195  | 
fold_types f t #> (fold o fold_types) f ts) o snd) assms;  | 
| 45421 | 196  | 
val base_sort =  | 
197  | 
if null inferred_elems then proto_base_sort  | 
|
198  | 
else  | 
|
199  | 
(case (fold o fold_element_types) Term.add_tfreesT inferred_elems [] of  | 
|
200  | 
[] => 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
 | 
201  | 
| [(_, sort)] => sort  | 
| 45421 | 202  | 
| _ => error "Multiple type variables in class specification");  | 
| 32206 | 203  | 
val supparams = map (fn ((c, T), _) =>  | 
| 70387 | 204  | 
(c, map_atyps (K (Term.aT base_sort)) T)) raw_supparams;  | 
| 32206 | 205  | 
val supparam_names = map fst supparams;  | 
206  | 
fun mk_param ((c, _), _) = Free (c, (the o AList.lookup (op =) supparams) c);  | 
|
207  | 
    val supexpr = (map (fn sup => (sup, (("", false),
 | 
|
| 
67450
 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 
ballarin 
parents: 
66334 
diff
changeset
 | 
208  | 
(Expression.Positional (map (SOME o mk_param) (Locale.params_of thy sup)), [])))) sups,  | 
| 32206 | 209  | 
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
 | 
210  | 
|
| 32206 | 211  | 
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
 | 
212  | 
|
| 
29665
 
2b956243d123
explicit check for exactly one type variable in class specification elements
 
haftmann 
parents: 
29632 
diff
changeset
 | 
213  | 
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
 | 
214  | 
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
 | 
215  | 
|
| 
72536
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
216  | 
fun prep_class_spec prep_class prep_include prep_class_elems ctxt raw_supclasses raw_includes raw_elems =  | 
| 
29575
 
41d604e59e93
improved and corrected reading of class specs -- still draft version
 
haftmann 
parents: 
29558 
diff
changeset
 | 
217  | 
let  | 
| 
72536
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
218  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 
29665
 
2b956243d123
explicit check for exactly one type variable in class specification elements
 
haftmann 
parents: 
29632 
diff
changeset
 | 
219  | 
|
| 
29575
 
41d604e59e93
improved and corrected reading of class specs -- still draft version
 
haftmann 
parents: 
29558 
diff
changeset
 | 
220  | 
(* prepare import *)  | 
| 
 
41d604e59e93
improved and corrected reading of class specs -- still draft version
 
haftmann 
parents: 
29558 
diff
changeset
 | 
221  | 
val inter_sort = curry (Sorts.inter_sort (Sign.classes_of thy));  | 
| 
72536
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
222  | 
val sups = Sign.minimize_sort thy (map (prep_class ctxt) raw_supclasses);  | 
| 45431 | 223  | 
val _ =  | 
224  | 
(case filter_out (Class.is_class thy) sups of  | 
|
225  | 
[] => ()  | 
|
226  | 
      | 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
 | 
227  | 
val raw_supparams = (map o apsnd) (snd o snd) (Class.these_params thy sups);  | 
| 32206 | 228  | 
val raw_supparam_names = map fst raw_supparams;  | 
| 45431 | 229  | 
val _ =  | 
230  | 
if has_duplicates (op =) raw_supparam_names then  | 
|
231  | 
        error ("Duplicate parameter(s) in superclasses: " ^
 | 
|
232  | 
(commas_quote (duplicates (op =) raw_supparam_names)))  | 
|
| 
29575
 
41d604e59e93
improved and corrected reading of class specs -- still draft version
 
haftmann 
parents: 
29558 
diff
changeset
 | 
233  | 
else ();  | 
| 
 
41d604e59e93
improved and corrected reading of class specs -- still draft version
 
haftmann 
parents: 
29558 
diff
changeset
 | 
234  | 
|
| 
 
41d604e59e93
improved and corrected reading of class specs -- still draft version
 
haftmann 
parents: 
29558 
diff
changeset
 | 
235  | 
(* infer types and base sort *)  | 
| 
72536
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
236  | 
val includes = map (prep_include ctxt) raw_includes;  | 
| 
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
237  | 
val includes_ctxt = Bundle.includes includes ctxt;  | 
| 
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
238  | 
val (base_sort, supparam_names, supexpr, inferred_elems) = prep_class_elems includes_ctxt sups raw_elems;  | 
| 32206 | 239  | 
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
 | 
240  | 
|
| 
 
41d604e59e93
improved and corrected reading of class specs -- still draft version
 
haftmann 
parents: 
29558 
diff
changeset
 | 
241  | 
(* process elements as class specification *)  | 
| 
72536
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
242  | 
val class_ctxt = Class.begin sups base_sort includes_ctxt;  | 
| 
47311
 
1addbe2a7458
close context elements via Expression.cert/read_declaration;
 
wenzelm 
parents: 
46922 
diff
changeset
 | 
243  | 
val ((_, _, syntax_elems, _), _) = class_ctxt  | 
| 29702 | 244  | 
|> Expression.cert_declaration supexpr I inferred_elems;  | 
| 45431 | 245  | 
fun check_vars e vs =  | 
246  | 
if null vs then  | 
|
247  | 
        error ("No type variable in part of specification element " ^
 | 
|
248  | 
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
 | 
249  | 
else ();  | 
| 
 
2b956243d123
explicit check for exactly one type variable in class specification elements
 
haftmann 
parents: 
29632 
diff
changeset
 | 
250  | 
fun check_element (e as Element.Fixes fxs) =  | 
| 45431 | 251  | 
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
 | 
252  | 
| check_element (e as Element.Assumes assms) =  | 
| 45431 | 253  | 
List.app (fn (_, ts_pss) =>  | 
254  | 
List.app (fn (t, _) => check_vars e (Term.add_tfrees t [])) ts_pss) assms  | 
|
255  | 
| check_element _ = ();  | 
|
256  | 
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
 | 
257  | 
fun fork_syn (Element.Fixes xs) =  | 
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30335 
diff
changeset
 | 
258  | 
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
 | 
259  | 
#>> Element.Fixes  | 
| 
 
2b956243d123
explicit check for exactly one type variable in class specification elements
 
haftmann 
parents: 
29632 
diff
changeset
 | 
260  | 
| fork_syn x = pair x;  | 
| 
29575
 
41d604e59e93
improved and corrected reading of class specs -- still draft version
 
haftmann 
parents: 
29558 
diff
changeset
 | 
261  | 
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
 | 
262  | 
|
| 
72536
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
263  | 
in (((sups, supparam_names), (sup_sort, base_sort, supexpr)), (includes, elems, global_syntax)) end;  | 
| 
29575
 
41d604e59e93
improved and corrected reading of class specs -- still draft version
 
haftmann 
parents: 
29558 
diff
changeset
 | 
264  | 
|
| 
72536
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
265  | 
val cert_class_spec = prep_class_spec (K I) (K I) cert_class_elems;  | 
| 
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
266  | 
val read_class_spec = prep_class_spec Proof_Context.read_class Bundle.check read_class_elems;  | 
| 
29665
 
2b956243d123
explicit check for exactly one type variable in class specification elements
 
haftmann 
parents: 
29632 
diff
changeset
 | 
267  | 
|
| 
 
2b956243d123
explicit check for exactly one type variable in class specification elements
 
haftmann 
parents: 
29632 
diff
changeset
 | 
268  | 
|
| 
 
2b956243d123
explicit check for exactly one type variable in class specification elements
 
haftmann 
parents: 
29632 
diff
changeset
 | 
269  | 
(* class establishment *)  | 
| 
29575
 
41d604e59e93
improved and corrected reading of class specs -- still draft version
 
haftmann 
parents: 
29558 
diff
changeset
 | 
270  | 
|
| 32206 | 271  | 
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
 | 
272  | 
let  | 
| 29816 | 273  | 
(*FIXME simplify*)  | 
| 32206 | 274  | 
val supconsts = supparam_names  | 
| 
38379
 
67d71449e85b
more convenient split of class modules: class and class_declaration
 
haftmann 
parents: 
38376 
diff
changeset
 | 
275  | 
|> AList.make (snd o the o AList.lookup (op =) (Class.these_params thy sups))  | 
| 70387 | 276  | 
|> (map o apsnd o apsnd o map_atyps) (K (Term.aT [class]));  | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29439 
diff
changeset
 | 
277  | 
val all_params = Locale.params_of thy class;  | 
| 32206 | 278  | 
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
 | 
279  | 
fun add_const ((raw_c, raw_ty), _) thy =  | 
| 25683 | 280  | 
let  | 
| 
30755
 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
281  | 
val b = Binding.name raw_c;  | 
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30335 
diff
changeset
 | 
282  | 
val c = Sign.full_name thy b;  | 
| 70387 | 283  | 
val ty = map_atyps (K (Term.aT base_sort)) raw_ty;  | 
| 79411 | 284  | 
val ty0 = Term.strip_sortsT ty;  | 
| 70387 | 285  | 
val ty' = map_atyps (K (Term.aT [class])) ty0;  | 
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30335 
diff
changeset
 | 
286  | 
val syn = (the_default NoSyn o AList.lookup Binding.eq_name global_syntax) b;  | 
| 25683 | 287  | 
in  | 
288  | 
thy  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
289  | 
|> Sign.declare_const_global ((b, ty0), syn)  | 
| 25683 | 290  | 
|> snd  | 
| 42494 | 291  | 
|> pair ((Variable.check_name b, ty), (c, ty'))  | 
| 25683 | 292  | 
end;  | 
| 
28715
 
238f9966c80e
class morphism stemming from locale interpretation
 
haftmann 
parents: 
28674 
diff
changeset
 | 
293  | 
in  | 
| 
 
238f9966c80e
class morphism stemming from locale interpretation
 
haftmann 
parents: 
28674 
diff
changeset
 | 
294  | 
thy  | 
| 
38379
 
67d71449e85b
more convenient split of class modules: class and class_declaration
 
haftmann 
parents: 
38376 
diff
changeset
 | 
295  | 
|> Sign.add_path (Class.class_prefix class)  | 
| 
28715
 
238f9966c80e
class morphism stemming from locale interpretation
 
haftmann 
parents: 
28674 
diff
changeset
 | 
296  | 
|> fold_map add_const raw_params  | 
| 
 
238f9966c80e
class morphism stemming from locale interpretation
 
haftmann 
parents: 
28674 
diff
changeset
 | 
297  | 
||> Sign.restore_naming thy  | 
| 
 
238f9966c80e
class morphism stemming from locale interpretation
 
haftmann 
parents: 
28674 
diff
changeset
 | 
298  | 
|-> (fn params => pair (supconsts @ (map o apfst) fst params, params))  | 
| 
 
238f9966c80e
class morphism stemming from locale interpretation
 
haftmann 
parents: 
28674 
diff
changeset
 | 
299  | 
end;  | 
| 
 
238f9966c80e
class morphism stemming from locale interpretation
 
haftmann 
parents: 
28674 
diff
changeset
 | 
300  | 
|
| 32206 | 301  | 
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
 | 
302  | 
let  | 
| 29816 | 303  | 
(*FIXME simplify*)  | 
| 25683 | 304  | 
fun globalize param_map = map_aterms  | 
305  | 
(fn Free (v, ty) => Const ((fst o the o AList.lookup (op =) param_map) v, ty)  | 
|
| 79438 | 306  | 
| _ => raise Same.SAME);  | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29439 
diff
changeset
 | 
307  | 
val raw_pred = Locale.intros_of thy class  | 
| 25683 | 308  | 
|> 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
 | 
309  | 
|> Option.map (Logic.unvarify_global o Logic.strip_imp_concl o Thm.prop_of);  | 
| 45431 | 310  | 
fun get_axiom thy =  | 
| 
51685
 
385ef6706252
more standard module name Axclass (according to file name);
 
wenzelm 
parents: 
51551 
diff
changeset
 | 
311  | 
(case #axioms (Axclass.get_info thy class) of  | 
| 45431 | 312  | 
[] => NONE  | 
313  | 
| [thm] => SOME thm);  | 
|
| 
24968
 
f9bafc868847
replaced Sign.add_consts_authentic by Sign.declare_const;
 
wenzelm 
parents: 
24949 
diff
changeset
 | 
314  | 
in  | 
| 
 
f9bafc868847
replaced Sign.add_consts_authentic by Sign.declare_const;
 
wenzelm 
parents: 
24949 
diff
changeset
 | 
315  | 
thy  | 
| 63352 | 316  | 
|> add_consts class base_sort sups supparam_names global_syntax |-> (fn (param_map, params) =>  | 
317  | 
Axclass.define_class (bname, supsort)  | 
|
318  | 
(map (fst o snd) params)  | 
|
319  | 
[(Binding.empty_atts, Option.map (globalize param_map) raw_pred |> the_list)]  | 
|
320  | 
#> snd  | 
|
321  | 
#> `get_axiom  | 
|
322  | 
#-> (fn assm_axiom => fold (Sign.add_const_constraint o apsnd SOME o snd) params  | 
|
323  | 
#> pair (param_map, params, assm_axiom)))  | 
|
| 
24968
 
f9bafc868847
replaced Sign.add_consts_authentic by Sign.declare_const;
 
wenzelm 
parents: 
24949 
diff
changeset
 | 
324  | 
end;  | 
| 
 
f9bafc868847
replaced Sign.add_consts_authentic by Sign.declare_const;
 
wenzelm 
parents: 
24949 
diff
changeset
 | 
325  | 
|
| 
72536
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
326  | 
fun gen_class prep_class_spec b raw_includes raw_supclasses raw_elems thy =  | 
| 24748 | 327  | 
let  | 
| 
36635
 
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
 
haftmann 
parents: 
36464 
diff
changeset
 | 
328  | 
val class = Sign.full_name thy b;  | 
| 
72536
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
329  | 
val ctxt = Proof_Context.init_global thy;  | 
| 
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
330  | 
val (((sups, supparam_names), (supsort, base_sort, supexpr)), (includes, elems, global_syntax)) =  | 
| 
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
331  | 
prep_class_spec ctxt raw_supclasses raw_includes raw_elems;  | 
| 
77955
 
c4677a6aae2c
more standard name bindings (amending 5bf71b4da706): avoid odd full_name like "Orderings.class.Orderings.preorder.of_class.intro" with many redundant name space accesses;
 
wenzelm 
parents: 
77879 
diff
changeset
 | 
332  | 
val of_class_binding = Binding.qualify_name true b "intro_of_class";  | 
| 24218 | 333  | 
in  | 
334  | 
thy  | 
|
| 
77955
 
c4677a6aae2c
more standard name bindings (amending 5bf71b4da706): avoid odd full_name like "Orderings.class.Orderings.preorder.of_class.intro" with many redundant name space accesses;
 
wenzelm 
parents: 
77879 
diff
changeset
 | 
335  | 
|> Expression.add_locale b (Binding.qualify true "class" b) includes supexpr elems  | 
| 33671 | 336  | 
|> snd |> Local_Theory.exit_global  | 
| 
36635
 
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
 
haftmann 
parents: 
36464 
diff
changeset
 | 
337  | 
|> adjungate_axclass b class base_sort sups supsort supparam_names global_syntax  | 
| 29526 | 338  | 
|-> (fn (param_map, params, assm_axiom) =>  | 
| 29547 | 339  | 
`(fn thy => calculate thy class sups base_sort param_map assm_axiom)  | 
| 52636 | 340  | 
#-> (fn (base_morph, eq_morph, export_morph, some_axiom, some_assm_intro, of_class) =>  | 
| 73845 | 341  | 
Context.theory_map (Locale.add_registration  | 
| 69058 | 342  | 
         {inst = (class, base_morph),
 | 
| 
68851
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
343  | 
mixin = Option.map (rpair true) eq_morph,  | 
| 73845 | 344  | 
export = 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
 | 
345  | 
#> Class.register class sups params base_sort base_morph export_morph some_axiom some_assm_intro of_class  | 
| 
77955
 
c4677a6aae2c
more standard name bindings (amending 5bf71b4da706): avoid odd full_name like "Orderings.class.Orderings.preorder.of_class.intro" with many redundant name space accesses;
 
wenzelm 
parents: 
77879 
diff
changeset
 | 
346  | 
#> Global_Theory.store_thm (of_class_binding, of_class)))  | 
| 
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
 | 
347  | 
|> snd  | 
| 
72536
 
589645894305
bundle mixins for locale and class specifications
 
haftmann 
parents: 
72505 
diff
changeset
 | 
348  | 
|> Named_Target.init includes class  | 
| 25038 | 349  | 
|> pair class  | 
| 24218 | 350  | 
end;  | 
351  | 
||
352  | 
in  | 
|
353  | 
||
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29439 
diff
changeset
 | 
354  | 
val class = gen_class cert_class_spec;  | 
| 26518 | 355  | 
val class_cmd = gen_class read_class_spec;  | 
| 24218 | 356  | 
|
357  | 
end; (*local*)  | 
|
358  | 
||
359  | 
||
| 45421 | 360  | 
|
| 29358 | 361  | 
(** subclass relations **)  | 
| 25462 | 362  | 
|
| 29358 | 363  | 
local  | 
| 25462 | 364  | 
|
| 
57181
 
2d13bf9ea77b
dropped obscure and unused ad-hoc before_exit hook for named targets
 
haftmann 
parents: 
56921 
diff
changeset
 | 
365  | 
fun gen_subclass prep_class do_proof raw_sup lthy =  | 
| 25462 | 366  | 
let  | 
| 42360 | 367  | 
val thy = Proof_Context.theory_of lthy;  | 
| 29558 | 368  | 
val proto_sup = prep_class thy raw_sup;  | 
| 63268 | 369  | 
val proto_sub =  | 
370  | 
(case Named_Target.class_of lthy of  | 
|
| 57182 | 371  | 
SOME class => class  | 
| 63268 | 372  | 
| NONE => error "Not in a class target");  | 
| 
51685
 
385ef6706252
more standard module name Axclass (according to file name);
 
wenzelm 
parents: 
51551 
diff
changeset
 | 
373  | 
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
 | 
374  | 
|
| 
67450
 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 
ballarin 
parents: 
66334 
diff
changeset
 | 
375  | 
    val expr = ([(sup, (("", false), (Expression.Positional [], [])))], []);
 | 
| 
 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 
ballarin 
parents: 
66334 
diff
changeset
 | 
376  | 
val (([props], _, deps, _, export), goal_ctxt) =  | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29439 
diff
changeset
 | 
377  | 
Expression.cert_goal_expression expr lthy;  | 
| 29526 | 378  | 
val some_prop = try the_single props;  | 
| 29558 | 379  | 
val some_dep_morph = try the_single (map snd deps);  | 
380  | 
fun after_qed some_wit =  | 
|
| 
54866
 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 
haftmann 
parents: 
54742 
diff
changeset
 | 
381  | 
Class.register_subclass (sub, sup) some_dep_morph some_wit export;  | 
| 29558 | 382  | 
in do_proof after_qed some_prop goal_ctxt end;  | 
| 25485 | 383  | 
|
| 
29575
 
41d604e59e93
improved and corrected reading of class specs -- still draft version
 
haftmann 
parents: 
29558 
diff
changeset
 | 
384  | 
fun user_proof after_qed some_prop =  | 
| 
 
41d604e59e93
improved and corrected reading of class specs -- still draft version
 
haftmann 
parents: 
29558 
diff
changeset
 | 
385  | 
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
 | 
386  | 
[the_list some_prop];  | 
| 25485 | 387  | 
|
| 
29575
 
41d604e59e93
improved and corrected reading of class specs -- still draft version
 
haftmann 
parents: 
29558 
diff
changeset
 | 
388  | 
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
 | 
389  | 
after_qed (Option.map  | 
| 
 
41d604e59e93
improved and corrected reading of class specs -- still draft version
 
haftmann 
parents: 
29558 
diff
changeset
 | 
390  | 
(fn prop => Element.prove_witness ctxt prop tac) some_prop) ctxt;  | 
| 28666 | 391  | 
|
| 29358 | 392  | 
in  | 
| 28666 | 393  | 
|
| 
57181
 
2d13bf9ea77b
dropped obscure and unused ad-hoc before_exit hook for named targets
 
haftmann 
parents: 
56921 
diff
changeset
 | 
394  | 
fun prove_subclass tac = gen_subclass (K I) (tactic_proof tac);  | 
| 54882 | 395  | 
|
396  | 
fun subclass x = gen_subclass (K I) user_proof x;  | 
|
397  | 
fun subclass_cmd x =  | 
|
398  | 
gen_subclass (Proof_Context.read_class o Proof_Context.init_global) user_proof x;  | 
|
| 25462 | 399  | 
|
| 29358 | 400  | 
end; (*local*)  | 
401  | 
||
| 24218 | 402  | 
end;  |