src/Pure/Isar/class_declaration.ML
author haftmann
Sun, 15 Nov 2020 07:17:05 +0000
changeset 72605 a4cb880e873a
parent 72536 589645894305
child 73845 bfce186331be
permissions -rw-r--r--
type alias for mixin bundles
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     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
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     5
*)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     6
38379
67d71449e85b more convenient split of class modules: class and class_declaration
haftmann
parents: 38376
diff changeset
     7
signature CLASS_DECLARATION =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     8
sig
72605
a4cb880e873a type alias for mixin bundles
haftmann
parents: 72536
diff changeset
     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
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    17
end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    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
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    20
struct
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    21
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
    22
(** class definitions **)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    23
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    24
local
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    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
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    28
fun calculate thy class sups base_sort param_map assm_axiom =
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    29
  let
70309
wenzelm
parents: 69058
diff changeset
    30
    val thy_ctxt = Proof_Context.init_global thy;
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    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
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    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
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    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 =
70387
wenzelm
parents: 70309
diff changeset
    44
      Element.instantiate_normalize_morphism ([(a_tfree, certT (Term.aT [class]))], [])
70309
wenzelm
parents: 69058
diff changeset
    45
    val (([raw_props], _, [(_, raw_inst_morph)], _, export_morph), _) = thy_ctxt
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    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
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
    48
    val (props, inst_morph) =
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
    49
      if null param_map
29797
08ef36ed2f8a handling type classes without parameters
haftmann
parents: 29702
diff changeset
    50
      then (raw_props |> map (Morphism.term typ_morph),
08ef36ed2f8a handling type classes without parameters
haftmann
parents: 29702
diff changeset
    51
        raw_inst_morph $> typ_morph)
08ef36ed2f8a handling type classes without parameters
haftmann
parents: 29702
diff changeset
    52
      else (raw_props, raw_inst_morph); (*FIXME proper handling in
08ef36ed2f8a handling type classes without parameters
haftmann
parents: 29702
diff changeset
    53
        locale.ML / expression.ML would be desirable*)
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    54
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    55
    (* witness for canonical interpretation *)
52636
238cec044ebf tuned variable names
haftmann
parents: 51685
diff changeset
    56
    val some_prop = try the_single props;
238cec044ebf tuned variable names
haftmann
parents: 51685
diff changeset
    57
    val some_witn = Option.map (fn prop =>
45431
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
    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
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
    60
        val loc_intro_tac =
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
    61
          (case Locale.intros_of thy class of
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
    62
            (_, NONE) => all_tac
70309
wenzelm
parents: 69058
diff changeset
    63
          | (_, SOME intro) => ALLGOALS (resolve_tac thy_ctxt [intro]));
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    64
        val tac = loc_intro_tac
70309
wenzelm
parents: 69058
diff changeset
    65
          THEN ALLGOALS (Proof_Context.fact_tac thy_ctxt (sup_axioms @ the_list assm_axiom));
wenzelm
parents: 69058
diff changeset
    66
      in Element.prove_witness thy_ctxt prop tac end) some_prop;
wenzelm
parents: 69058
diff changeset
    67
    val some_axiom = Option.map (Element.conclude_witness thy_ctxt) some_witn;
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    68
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    69
    (* canonical interpretation *)
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    70
    val base_morph = inst_morph
54740
91f54d386680 maintain morphism names for diagnostic purposes;
wenzelm
parents: 52788
diff changeset
    71
      $> Morphism.binding_morphism "class_binding" (Binding.prefix false (Class.class_prefix class))
52636
238cec044ebf tuned variable names
haftmann
parents: 51685
diff changeset
    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
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    74
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    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
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    77
      let
70309
wenzelm
parents: 69058
diff changeset
    78
        val ((_, [thm']), _) = Variable.import true [thm] thy_ctxt;
45431
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
    79
        val const_eq_morph =
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
    80
          (case eq_morph of
46856
wenzelm
parents: 45433
diff changeset
    81
            SOME eq_morph => const_morph $> eq_morph
45431
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
    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
238cec044ebf tuned variable names
haftmann
parents: 51685
diff changeset
    88
    val some_assm_intro = Option.map prove_assm_intro (fst (Locale.intros_of thy class));
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    89
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    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
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
    92
    val of_class_prop =
52636
238cec044ebf tuned variable names
haftmann
parents: 51685
diff changeset
    93
      (case some_prop of
45431
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
    94
        NONE => of_class_prop_concl
29627
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    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
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58837
diff changeset
   101
    fun tac ctxt =
45431
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   102
      REPEAT (SOMEGOAL
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58837
diff changeset
   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));
71018
d32ed8927a42 more robust expose_proofs corresponding to register_proofs/consolidate_theory;
wenzelm
parents: 71017
diff changeset
   105
    val of_class = Goal.prove_sorry_global thy [] [] of_class_prop (tac o #context);
52636
238cec044ebf tuned variable names
haftmann
parents: 51685
diff changeset
   106
  in (base_morph, eq_morph, export_morph, some_axiom, some_assm_intro, of_class) end;
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
   107
29632
c3d576157244 fixed reading of class specs: declare class operations in context
haftmann
parents: 29627
diff changeset
   108
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   109
(* reading and processing class specifications *)
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   110
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   111
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
   112
  let
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   113
    (* 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
   114
    val thy = Proof_Context.theory_of ctxt;
36460
c643b23e8592 empty class specifcations observe default sort
haftmann
parents: 35845
diff changeset
   115
    val algebra = Sign.classes_of thy;
c643b23e8592 empty class specifcations observe default sort
haftmann
parents: 35845
diff changeset
   116
    val inter_sort = curry (Sorts.inter_sort algebra);
45421
2bef6da4a6a6 tuned layout;
wenzelm
parents: 42494
diff changeset
   117
    val proto_base_sort =
2bef6da4a6a6 tuned layout;
wenzelm
parents: 42494
diff changeset
   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
90d0103af558 explicit input marker for operations
haftmann
parents: 59621
diff changeset
   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
wenzelm
parents: 60346
diff changeset
   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
haftmann
parents: 57182
diff changeset
   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
haftmann
parents: 57182
diff changeset
   128
          (fold o fold_atyps) (fn TVar (_, S) => inter_sort S | _ => I) Ts [];
45421
2bef6da4a6a6 tuned layout;
wenzelm
parents: 42494
diff changeset
   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");
70387
wenzelm
parents: 70309
diff changeset
   142
        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
   143
      in
58293
haftmann
parents: 57182
diff changeset
   144
        (map o map_atyps)
haftmann
parents: 57182
diff changeset
   145
          (fn T as TVar (xi, _) => if Type_Infer.is_param xi then fixateT else T | T => T) Ts
62952
wenzelm
parents: 60346
diff changeset
   146
      end);
wenzelm
parents: 60346
diff changeset
   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
9228350683d0 use proto_base_sort uniformly
haftmann
parents: 58294
diff changeset
   151
        val param_namesT = map_filter (try (fst o dest_TVar)) param_Ts;
9228350683d0 use proto_base_sort uniformly
haftmann
parents: 58294
diff changeset
   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
9228350683d0 use proto_base_sort uniformly
haftmann
parents: 58294
diff changeset
   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
9228350683d0 use proto_base_sort uniformly
haftmann
parents: 58294
diff changeset
   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
0a3adceb9c67 tuned comments
haftmann
parents: 35021
diff changeset
   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
wenzelm
parents: 60346
diff changeset
   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, _), _) =
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   169
      ctxt
62952
wenzelm
parents: 60346
diff changeset
   170
      |> Context.proof_map (Syntax_Phases.term_check 0 "unify_params" (K unify_params))
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   171
      |> prep_decl raw_supexpr init_class_body raw_elems;
38435
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   172
    fun filter_element (Element.Fixes []) = NONE
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   173
      | filter_element (e as Element.Fixes _) = SOME e
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   174
      | filter_element (Element.Constrains []) = NONE
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   175
      | filter_element (e as Element.Constrains _) = SOME e
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   176
      | filter_element (Element.Assumes []) = NONE
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   177
      | filter_element (e as Element.Assumes _) = SOME e
45421
2bef6da4a6a6 tuned layout;
wenzelm
parents: 42494
diff changeset
   178
      | filter_element (Element.Defines _) =
2bef6da4a6a6 tuned layout;
wenzelm
parents: 42494
diff changeset
   179
          error ("\"defines\" element not allowed in class specification.")
2bef6da4a6a6 tuned layout;
wenzelm
parents: 42494
diff changeset
   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
2bef6da4a6a6 tuned layout;
wenzelm
parents: 42494
diff changeset
   183
          error ("\"notes\" element not allowed in class specification.");
38435
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   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
2bef6da4a6a6 tuned layout;
wenzelm
parents: 42494
diff changeset
   189
    val base_sort =
2bef6da4a6a6 tuned layout;
wenzelm
parents: 42494
diff changeset
   190
      if null inferred_elems then proto_base_sort
2bef6da4a6a6 tuned layout;
wenzelm
parents: 42494
diff changeset
   191
      else
2bef6da4a6a6 tuned layout;
wenzelm
parents: 42494
diff changeset
   192
        (case (fold o fold_element_types) Term.add_tfreesT inferred_elems [] of
2bef6da4a6a6 tuned layout;
wenzelm
parents: 42494
diff changeset
   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
2bef6da4a6a6 tuned layout;
wenzelm
parents: 42494
diff changeset
   195
        | _ => error "Multiple type variables in class specification");
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   196
    val supparams = map (fn ((c, T), _) =>
70387
wenzelm
parents: 70309
diff changeset
   197
      (c, map_atyps (K (Term.aT base_sort)) T)) raw_supparams;
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   198
    val supparam_names = map fst supparams;
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   199
    fun mk_param ((c, _), _) = Free (c, (the o AList.lookup (op =) supparams) c);
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   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
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   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
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   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
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   209
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
   210
  let
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   211
    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
   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));
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   215
    val sups = Sign.minimize_sort thy (map (prep_class ctxt) raw_supclasses);
45431
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   216
    val _ =
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   217
      (case filter_out (Class.is_class thy) sups of
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   218
        [] => ()
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   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
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   221
    val raw_supparam_names = map fst raw_supparams;
45431
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   222
    val _ =
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   223
      if has_duplicates (op =) raw_supparam_names then
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   224
        error ("Duplicate parameter(s) in superclasses: " ^
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   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 *)
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   229
    val includes = map (prep_include ctxt) raw_includes;
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   230
    val includes_ctxt = Bundle.includes includes ctxt;
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   231
    val (base_sort, supparam_names, supexpr, inferred_elems) = prep_class_elems includes_ctxt sups raw_elems;
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   232
    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
   233
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   234
    (* process elements as class specification *)
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   235
    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
   236
    val ((_, _, syntax_elems, _), _) = class_ctxt
29702
a7512f22e916 proper declared constants in class expressions
haftmann
parents: 29665
diff changeset
   237
      |> Expression.cert_declaration supexpr I inferred_elems;
45431
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   238
    fun check_vars e vs =
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   239
      if null vs then
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   240
        error ("No type variable in part of specification element " ^
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   241
          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
   242
      else ();
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   243
    fun check_element (e as Element.Fixes fxs) =
45431
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   244
          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
   245
      | check_element (e as Element.Assumes assms) =
45431
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   246
          List.app (fn (_, ts_pss) =>
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   247
            List.app (fn (t, _) => check_vars e (Term.add_tfrees t [])) ts_pss) assms
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   248
      | check_element _ = ();
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   249
    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
   250
    fun fork_syn (Element.Fixes xs) =
30344
10a67c5ddddb more uniform handling of binding in targets and derived elements;
wenzelm
parents: 30335
diff changeset
   251
          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
   252
          #>> Element.Fixes
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   253
      | fork_syn x = pair x;
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   254
    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
   255
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   256
  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
   257
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   258
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
   259
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
   260
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   261
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   262
(* class establishment *)
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   263
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   264
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
   265
  let
29816
78e0a90694dd more robust failure in error situations
haftmann
parents: 29797
diff changeset
   266
    (*FIXME simplify*)
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   267
    val supconsts = supparam_names
38379
67d71449e85b more convenient split of class modules: class and class_declaration
haftmann
parents: 38376
diff changeset
   268
      |> AList.make (snd o the o AList.lookup (op =) (Class.these_params thy sups))
70387
wenzelm
parents: 70309
diff changeset
   269
      |> (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
   270
    val all_params = Locale.params_of thy class;
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   271
    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
   272
    fun add_const ((raw_c, raw_ty), _) thy =
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   273
      let
30755
7ef503d216c2 simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
wenzelm
parents: 30585
diff changeset
   274
        val b = Binding.name raw_c;
30344
10a67c5ddddb more uniform handling of binding in targets and derived elements;
wenzelm
parents: 30335
diff changeset
   275
        val c = Sign.full_name thy b;
70387
wenzelm
parents: 70309
diff changeset
   276
        val ty = map_atyps (K (Term.aT base_sort)) raw_ty;
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   277
        val ty0 = Type.strip_sorts ty;
70387
wenzelm
parents: 70309
diff changeset
   278
        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
   279
        val syn = (the_default NoSyn o AList.lookup Binding.eq_name global_syntax) b;
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   280
      in
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   281
        thy
42375
774df7c59508 report Name_Space.declare/define, relatively to context;
wenzelm
parents: 42360
diff changeset
   282
        |> Sign.declare_const_global ((b, ty0), syn)
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   283
        |> snd
42494
eef1a23c9077 tuned signature -- eliminated odd comment;
wenzelm
parents: 42402
diff changeset
   284
        |> pair ((Variable.check_name b, ty), (c, ty'))
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   285
      end;
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   286
  in
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   287
    thy
38379
67d71449e85b more convenient split of class modules: class and class_declaration
haftmann
parents: 38376
diff changeset
   288
    |> Sign.add_path (Class.class_prefix class)
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   289
    |> fold_map add_const raw_params
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   290
    ||> Sign.restore_naming thy
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   291
    |-> (fn params => pair (supconsts @ (map o apfst) fst params, params))
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   292
  end;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   293
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   294
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
   295
  let
29816
78e0a90694dd more robust failure in error situations
haftmann
parents: 29797
diff changeset
   296
    (*FIXME simplify*)
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   297
    fun globalize param_map = map_aterms
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   298
      (fn Free (v, ty) => Const ((fst o the o AList.lookup (op =) param_map) v, ty)
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   299
        | t => t);
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   300
    val raw_pred = Locale.intros_of thy class
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   301
      |> 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
   302
      |> Option.map (Logic.unvarify_global o Logic.strip_imp_concl o Thm.prop_of);
45431
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   303
    fun get_axiom thy =
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51551
diff changeset
   304
      (case #axioms (Axclass.get_info thy class) of
45431
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   305
         [] => NONE
924c2f6dcd05 misc tuning and simplification;
wenzelm
parents: 45429
diff changeset
   306
      | [thm] => SOME thm);
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   307
  in
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   308
    thy
63352
4eaf35781b23 tuned signature;
wenzelm
parents: 63268
diff changeset
   309
    |> add_consts class base_sort sups supparam_names global_syntax |-> (fn (param_map, params) =>
4eaf35781b23 tuned signature;
wenzelm
parents: 63268
diff changeset
   310
      Axclass.define_class (bname, supsort)
4eaf35781b23 tuned signature;
wenzelm
parents: 63268
diff changeset
   311
        (map (fst o snd) params)
4eaf35781b23 tuned signature;
wenzelm
parents: 63268
diff changeset
   312
          [(Binding.empty_atts, Option.map (globalize param_map) raw_pred |> the_list)]
4eaf35781b23 tuned signature;
wenzelm
parents: 63268
diff changeset
   313
      #> snd
4eaf35781b23 tuned signature;
wenzelm
parents: 63268
diff changeset
   314
      #> `get_axiom
4eaf35781b23 tuned signature;
wenzelm
parents: 63268
diff changeset
   315
      #-> (fn assm_axiom => fold (Sign.add_const_constraint o apsnd SOME o snd) params
4eaf35781b23 tuned signature;
wenzelm
parents: 63268
diff changeset
   316
      #> pair (param_map, params, assm_axiom)))
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   317
  end;
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   318
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   319
fun gen_class prep_class_spec b raw_includes raw_supclasses raw_elems thy =
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   320
  let
36635
080b755377c0 locale predicates of classes carry a mandatory "class" prefix
haftmann
parents: 36464
diff changeset
   321
    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
   322
    val prefix = Binding.qualify true "class";
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   323
    val ctxt = Proof_Context.init_global thy;
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   324
    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
   325
      prep_class_spec ctxt raw_supclasses raw_includes raw_elems;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   326
  in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   327
    thy
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   328
    |> Expression.add_locale b (prefix b) includes supexpr elems
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33553
diff changeset
   329
    |> snd |> Local_Theory.exit_global
36635
080b755377c0 locale predicates of classes carry a mandatory "class" prefix
haftmann
parents: 36464
diff changeset
   330
    |> adjungate_axclass b class base_sort sups supsort supparam_names global_syntax
29526
0b32c8b84d3e code cleanup
haftmann
parents: 29509
diff changeset
   331
    |-> (fn (param_map, params, assm_axiom) =>
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
   332
       `(fn thy => calculate thy class sups base_sort param_map assm_axiom)
52636
238cec044ebf tuned variable names
haftmann
parents: 51685
diff changeset
   333
    #-> (fn (base_morph, eq_morph, export_morph, some_axiom, some_assm_intro, of_class) =>
72505
974071d873ba tuned interfaces
haftmann
parents: 71018
diff changeset
   334
       Locale.add_registration_theory'
69058
f4fb93197670 tuned signature;
wenzelm
parents: 68853
diff changeset
   335
         {inst = (class, base_morph),
68851
6c9825c1e26b clarified signature: explicit type Locale.registration;
wenzelm
parents: 67699
diff changeset
   336
           mixin = Option.map (rpair true) eq_morph,
68853
d36f00510e40 tuned signature;
wenzelm
parents: 68851
diff changeset
   337
           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
   338
    #> 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
   339
    #> 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
   340
    |> snd
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 72505
diff changeset
   341
    |> Named_Target.init includes class
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   342
    |> pair class
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   343
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   344
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   345
in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   346
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   347
val class = gen_class cert_class_spec;
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   348
val class_cmd = gen_class read_class_spec;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   349
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   350
end; (*local*)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   351
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   352
45421
2bef6da4a6a6 tuned layout;
wenzelm
parents: 42494
diff changeset
   353
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   354
(** subclass relations **)
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   355
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   356
local
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   357
57181
2d13bf9ea77b dropped obscure and unused ad-hoc before_exit hook for named targets
haftmann
parents: 56921
diff changeset
   358
fun gen_subclass prep_class do_proof raw_sup lthy =
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   359
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42357
diff changeset
   360
    val thy = Proof_Context.theory_of lthy;
29558
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   361
    val proto_sup = prep_class thy raw_sup;
63268
wenzelm
parents: 62952
diff changeset
   362
    val proto_sub =
wenzelm
parents: 62952
diff changeset
   363
      (case Named_Target.class_of lthy of
57182
79d43c510b84 less baroque interface
haftmann
parents: 57181
diff changeset
   364
        SOME class => class
63268
wenzelm
parents: 62952
diff changeset
   365
      | NONE => error "Not in a class target");
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51551
diff changeset
   366
    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
   367
67450
b0ae74b86ef3 Experimental support for rewrite morphisms in locale instances.
ballarin
parents: 66334
diff changeset
   368
    val expr = ([(sup, (("", false), (Expression.Positional [], [])))], []);
b0ae74b86ef3 Experimental support for rewrite morphisms in locale instances.
ballarin
parents: 66334
diff changeset
   369
    val (([props], _, deps, _, export), goal_ctxt) =
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   370
      Expression.cert_goal_expression expr lthy;
29526
0b32c8b84d3e code cleanup
haftmann
parents: 29509
diff changeset
   371
    val some_prop = try the_single props;
29558
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   372
    val some_dep_morph = try the_single (map snd deps);
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   373
    fun after_qed some_wit =
54866
7b9a67cbd48f self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents: 54742
diff changeset
   374
      Class.register_subclass (sub, sup) some_dep_morph some_wit export;
29558
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   375
  in do_proof after_qed some_prop goal_ctxt end;
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   376
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   377
fun user_proof after_qed some_prop =
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   378
  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
   379
    [the_list some_prop];
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   380
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   381
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
   382
  after_qed (Option.map
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   383
    (fn prop => Element.prove_witness ctxt prop tac) some_prop) ctxt;
28666
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   384
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   385
in
28666
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   386
57181
2d13bf9ea77b dropped obscure and unused ad-hoc before_exit hook for named targets
haftmann
parents: 56921
diff changeset
   387
fun prove_subclass tac = gen_subclass (K I) (tactic_proof tac);
54882
61276a7fc369 made SML/NJ happy;
wenzelm
parents: 54866
diff changeset
   388
61276a7fc369 made SML/NJ happy;
wenzelm
parents: 54866
diff changeset
   389
fun subclass x = gen_subclass (K I) user_proof x;
61276a7fc369 made SML/NJ happy;
wenzelm
parents: 54866
diff changeset
   390
fun subclass_cmd x =
61276a7fc369 made SML/NJ happy;
wenzelm
parents: 54866
diff changeset
   391
  gen_subclass (Proof_Context.read_class o Proof_Context.init_global) user_proof x;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   392
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   393
end; (*local*)
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   394
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   395
end;