src/Pure/Isar/class_declaration.ML
author wenzelm
Thu, 24 Mar 2011 16:56:19 +0100
changeset 42083 e1209fc7ecdc
parent 41585 45d7da4e4ccf
child 42357 3305f573294e
permissions -rw-r--r--
added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
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
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
     9
  val class: (local_theory -> local_theory) -> binding -> class list ->
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
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
    11
  val class_cmd: (local_theory -> local_theory) -> binding -> xstring list ->
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
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
    13
  val prove_subclass: (local_theory -> local_theory) -> tactic -> class ->
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
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
    15
  val subclass: (local_theory -> local_theory) -> class -> local_theory -> Proof.state
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
    16
  val subclass_cmd: (local_theory -> local_theory) -> 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
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36464
diff changeset
    30
    val empty_ctxt = ProofContext.init_global thy;
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    31
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    32
    (* instantiation of canonical interpretation *)
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
    33
    val aT = TFree (Name.aT, base_sort);
29627
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    34
    val param_map_const = (map o apsnd) Const param_map;
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    35
    val param_map_inst = (map o apsnd)
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    36
      (Const o apsnd (map_atyps (K aT))) param_map;
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    37
    val const_morph = Element.inst_morphism thy
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    38
      (Symtab.empty, Symtab.make param_map_inst);
29797
08ef36ed2f8a handling type classes without parameters
haftmann
parents: 29702
diff changeset
    39
    val typ_morph = Element.inst_morphism thy
08ef36ed2f8a handling type classes without parameters
haftmann
parents: 29702
diff changeset
    40
      (Symtab.empty |> Symtab.update (Name.aT, TFree (Name.aT, [class])), Symtab.empty);
08ef36ed2f8a handling type classes without parameters
haftmann
parents: 29702
diff changeset
    41
    val (([raw_props], [(_, raw_inst_morph)], export_morph), _) = empty_ctxt
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    42
      |> Expression.cert_goal_expression ([(class, (("", false),
29627
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    43
           Expression.Named param_map_const))], []);
29797
08ef36ed2f8a handling type classes without parameters
haftmann
parents: 29702
diff changeset
    44
    val (props, inst_morph) = if null param_map
08ef36ed2f8a handling type classes without parameters
haftmann
parents: 29702
diff changeset
    45
      then (raw_props |> map (Morphism.term typ_morph),
08ef36ed2f8a handling type classes without parameters
haftmann
parents: 29702
diff changeset
    46
        raw_inst_morph $> typ_morph)
08ef36ed2f8a handling type classes without parameters
haftmann
parents: 29702
diff changeset
    47
      else (raw_props, raw_inst_morph); (*FIXME proper handling in
08ef36ed2f8a handling type classes without parameters
haftmann
parents: 29702
diff changeset
    48
        locale.ML / expression.ML would be desirable*)
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    49
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    50
    (* witness for canonical interpretation *)
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    51
    val prop = try the_single props;
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    52
    val wit = Option.map (fn prop => let
38379
67d71449e85b more convenient split of class modules: class and class_declaration
haftmann
parents: 38376
diff changeset
    53
        val sup_axioms = map_filter (fst o Class.rules thy) sups;
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    54
        val loc_intro_tac = case Locale.intros_of thy class
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    55
          of (_, NONE) => all_tac
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    56
           | (_, SOME intro) => ALLGOALS (Tactic.rtac intro);
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    57
        val tac = loc_intro_tac
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    58
          THEN ALLGOALS (ProofContext.fact_tac (sup_axioms @ the_list assm_axiom))
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    59
      in Element.prove_witness empty_ctxt prop tac end) prop;
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    60
    val axiom = Option.map Element.conclude_witness wit;
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    61
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    62
    (* canonical interpretation *)
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    63
    val base_morph = inst_morph
38379
67d71449e85b more convenient split of class modules: class and class_declaration
haftmann
parents: 38376
diff changeset
    64
      $> Morphism.binding_morphism (Binding.prefix false (Class.class_prefix class))
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    65
      $> Element.satisfy_morphism (the_list wit);
38379
67d71449e85b more convenient split of class modules: class and class_declaration
haftmann
parents: 38376
diff changeset
    66
    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
    67
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    68
    (* assm_intro *)
30344
10a67c5ddddb more uniform handling of binding in targets and derived elements;
wenzelm
parents: 30335
diff changeset
    69
    fun prove_assm_intro thm =
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    70
      let
31794
71af1fd6a5e4 renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents: 31696
diff changeset
    71
        val ((_, [thm']), _) = Variable.import true [thm] empty_ctxt;
36674
d95f39448121 eq_morphism is always optional: avoid trivial morphism for empty list of equations
haftmann
parents: 36672
diff changeset
    72
        val const_eq_morph = case eq_morph
d95f39448121 eq_morphism is always optional: avoid trivial morphism for empty list of equations
haftmann
parents: 36672
diff changeset
    73
         of SOME eq_morph => const_morph $> eq_morph
d95f39448121 eq_morphism is always optional: avoid trivial morphism for empty list of equations
haftmann
parents: 36672
diff changeset
    74
          | NONE => const_morph
d95f39448121 eq_morphism is always optional: avoid trivial morphism for empty list of equations
haftmann
parents: 36672
diff changeset
    75
        val thm'' = Morphism.thm const_eq_morph thm';
29627
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    76
        val tac = ALLGOALS (ProofContext.fact_tac [thm'']);
32970
fbd2bb2489a8 operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents: 32886
diff changeset
    77
      in Skip_Proof.prove_global thy [] [] (Thm.prop_of thm'') (K tac) end;
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    78
    val assm_intro = Option.map prove_assm_intro
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    79
      (fst (Locale.intros_of thy class));
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    80
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    81
    (* of_class *)
31943
5e960a0780a2 renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
wenzelm
parents: 31904
diff changeset
    82
    val of_class_prop_concl = Logic.mk_of_class (aT, class);
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    83
    val of_class_prop = case prop of NONE => of_class_prop_concl
29627
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    84
      | SOME prop => Logic.mk_implies (Morphism.term const_morph
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    85
          ((map_types o map_atyps) (K aT) prop), of_class_prop_concl);
38379
67d71449e85b more convenient split of class modules: class and class_declaration
haftmann
parents: 38376
diff changeset
    86
    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
    87
    val loc_axiom_intros = map Drule.export_without_context_open (Locale.axioms_of thy class);
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    88
    val axclass_intro = #intro (AxClass.get_info thy class);
31944
c8a35979a5bc clarified Thm.of_class/of_sort/class_triv;
wenzelm
parents: 31943
diff changeset
    89
    val base_sort_trivs = Thm.of_sort (Thm.ctyp_of thy aT, base_sort);
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    90
    val tac = REPEAT (SOMEGOAL
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    91
      (Tactic.match_tac (axclass_intro :: sup_of_classes
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    92
         @ loc_axiom_intros @ base_sort_trivs)
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    93
           ORELSE' Tactic.assume_tac));
32970
fbd2bb2489a8 operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents: 32886
diff changeset
    94
    val of_class = Skip_Proof.prove_global thy [] [] of_class_prop (K tac);
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    95
32886
aba29da80c1b do not use Locale.add_registration_eqs any longer
haftmann
parents: 32850
diff changeset
    96
  in (base_morph, eq_morph, export_morph, axiom, assm_intro, of_class) end;
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    97
29632
c3d576157244 fixed reading of class specs: declare class operations in context
haftmann
parents: 29627
diff changeset
    98
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
    99
(* reading and processing class specifications *)
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   100
36460
c643b23e8592 empty class specifcations observe default sort
haftmann
parents: 35845
diff changeset
   101
fun prep_class_elems prep_decl thy sups raw_elems =
29632
c3d576157244 fixed reading of class specs: declare class operations in context
haftmann
parents: 29627
diff changeset
   102
  let
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   103
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   104
    (* user space type system: only permits 'a type variable, improves towards 'a *)
36460
c643b23e8592 empty class specifcations observe default sort
haftmann
parents: 35845
diff changeset
   105
    val algebra = Sign.classes_of thy;
c643b23e8592 empty class specifcations observe default sort
haftmann
parents: 35845
diff changeset
   106
    val inter_sort = curry (Sorts.inter_sort algebra);
c643b23e8592 empty class specifcations observe default sort
haftmann
parents: 35845
diff changeset
   107
    val proto_base_sort = if null sups then Sign.defaultS thy
38379
67d71449e85b more convenient split of class modules: class and class_declaration
haftmann
parents: 38376
diff changeset
   108
      else fold inter_sort (map (Class.base_sort thy) sups) [];
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   109
    val base_constraints = (map o apsnd)
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   110
      (map_type_tfree (K (TVar ((Name.aT, 0), proto_base_sort))) o fst o snd)
38379
67d71449e85b more convenient split of class modules: class and class_declaration
haftmann
parents: 38376
diff changeset
   111
        (Class.these_operations thy sups);
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   112
    val reject_bcd_etc = (map o map_atyps) (fn T as TFree (v, sort) =>
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   113
          if v = Name.aT then T
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   114
          else error ("No type variable other than " ^ Name.aT ^ " allowed in class specification")
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   115
      | T => T);
36460
c643b23e8592 empty class specifcations observe default sort
haftmann
parents: 35845
diff changeset
   116
    fun singleton_fixate Ts =
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   117
      let
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   118
        fun extract f = (fold o fold_atyps) f Ts [];
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   119
        val tfrees = extract
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   120
          (fn TFree (v, sort) => insert (op =) (v, sort) | _ => I);
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   121
        val inferred_sort = extract
36460
c643b23e8592 empty class specifcations observe default sort
haftmann
parents: 35845
diff changeset
   122
          (fn TVar (_, sort) => inter_sort sort | _ => I);
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   123
        val fixate_sort = if null tfrees then inferred_sort
29816
78e0a90694dd more robust failure in error situations
haftmann
parents: 29797
diff changeset
   124
          else case tfrees
78e0a90694dd more robust failure in error situations
haftmann
parents: 29797
diff changeset
   125
           of [(_, a_sort)] => if Sorts.sort_le algebra (a_sort, inferred_sort)
36460
c643b23e8592 empty class specifcations observe default sort
haftmann
parents: 35845
diff changeset
   126
                then inter_sort a_sort inferred_sort
29816
78e0a90694dd more robust failure in error situations
haftmann
parents: 29797
diff changeset
   127
                else error ("Type inference imposes additional sort constraint "
78e0a90694dd more robust failure in error situations
haftmann
parents: 29797
diff changeset
   128
                  ^ Syntax.string_of_sort_global thy inferred_sort
78e0a90694dd more robust failure in error situations
haftmann
parents: 29797
diff changeset
   129
                  ^ " of type parameter " ^ Name.aT ^ " of sort "
38875
c7a66b584147 tuned messages: discontinued spurious full-stops (messages are occasionally composed unexpectedly);
wenzelm
parents: 38756
diff changeset
   130
                  ^ Syntax.string_of_sort_global thy a_sort)
c7a66b584147 tuned messages: discontinued spurious full-stops (messages are occasionally composed unexpectedly);
wenzelm
parents: 38756
diff changeset
   131
            | _ => error "Multiple type variables in class specification";
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   132
      in (map o map_atyps) (K (TFree (Name.aT, fixate_sort))) Ts end;
40188
eddda8e38360 consider base sorts of superclasses simultaneously during processing of class specification -- avoids clash of different base sorts
haftmann
parents: 38875
diff changeset
   133
    fun after_infer_fixate Ts =
eddda8e38360 consider base sorts of superclasses simultaneously during processing of class specification -- avoids clash of different base sorts
haftmann
parents: 38875
diff changeset
   134
      let
eddda8e38360 consider base sorts of superclasses simultaneously during processing of class specification -- avoids clash of different base sorts
haftmann
parents: 38875
diff changeset
   135
        val sort' = (fold o fold_atyps) (fn T as TFree _ => I | T as TVar (vi, sort) =>
eddda8e38360 consider base sorts of superclasses simultaneously during processing of class specification -- avoids clash of different base sorts
haftmann
parents: 38875
diff changeset
   136
          if Type_Infer.is_param vi then inter_sort sort else I) Ts [];
eddda8e38360 consider base sorts of superclasses simultaneously during processing of class specification -- avoids clash of different base sorts
haftmann
parents: 38875
diff changeset
   137
      in
eddda8e38360 consider base sorts of superclasses simultaneously during processing of class specification -- avoids clash of different base sorts
haftmann
parents: 38875
diff changeset
   138
        (map o map_atyps)
eddda8e38360 consider base sorts of superclasses simultaneously during processing of class specification -- avoids clash of different base sorts
haftmann
parents: 38875
diff changeset
   139
          (fn T as TFree _ => T | T as TVar (vi, _) =>
eddda8e38360 consider base sorts of superclasses simultaneously during processing of class specification -- avoids clash of different base sorts
haftmann
parents: 38875
diff changeset
   140
            if Type_Infer.is_param vi then Type_Infer.param 0 (Name.aT, sort') else T) Ts
eddda8e38360 consider base sorts of superclasses simultaneously during processing of class specification -- avoids clash of different base sorts
haftmann
parents: 38875
diff changeset
   141
      end;
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   142
    fun add_typ_check level name f = Context.proof_map
38435
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   143
      (Syntax.add_typ_check level name (fn Ts => fn ctxt =>
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   144
        let val Ts' = f Ts in if eq_list (op =) (Ts, Ts') then NONE else SOME (Ts', ctxt) end));
29632
c3d576157244 fixed reading of class specs: declare class operations in context
haftmann
parents: 29627
diff changeset
   145
35120
0a3adceb9c67 tuned comments
haftmann
parents: 35021
diff changeset
   146
    (* preprocessing elements, retrieving base sort from type-checked elements *)
38493
95a41e6ef5a8 enforcing a singleton type inference parameter after type inference and before fixation prevents multiple type variables in import during class declaration
haftmann
parents: 38435
diff changeset
   147
    val raw_supexpr = (map (fn sup => (sup, (("", false),
95a41e6ef5a8 enforcing a singleton type inference parameter after type inference and before fixation prevents multiple type variables in import during class declaration
haftmann
parents: 38435
diff changeset
   148
      Expression.Positional []))) sups, []);
29702
a7512f22e916 proper declared constants in class expressions
haftmann
parents: 29665
diff changeset
   149
    val init_class_body = fold (ProofContext.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
   150
      #> Class.redeclare_operations thy sups
29702
a7512f22e916 proper declared constants in class expressions
haftmann
parents: 29665
diff changeset
   151
      #> add_typ_check 10 "reject_bcd_etc" reject_bcd_etc
36460
c643b23e8592 empty class specifcations observe default sort
haftmann
parents: 35845
diff changeset
   152
      #> add_typ_check ~10 "singleton_fixate" singleton_fixate;
38435
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   153
    val ((raw_supparams, _, raw_inferred_elems), _) = ProofContext.init_global thy
38493
95a41e6ef5a8 enforcing a singleton type inference parameter after type inference and before fixation prevents multiple type variables in import during class declaration
haftmann
parents: 38435
diff changeset
   154
      |> add_typ_check 5 "after_infer_fixate" after_infer_fixate
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   155
      |> prep_decl raw_supexpr init_class_body raw_elems;
38435
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   156
    fun filter_element (Element.Fixes []) = NONE
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   157
      | filter_element (e as Element.Fixes _) = SOME e
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   158
      | filter_element (Element.Constrains []) = NONE
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   159
      | filter_element (e as Element.Constrains _) = SOME e
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   160
      | filter_element (Element.Assumes []) = NONE
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   161
      | filter_element (e as Element.Assumes _) = SOME e
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   162
      | filter_element (Element.Defines _) = error ("\"defines\" element not allowed in class specification.")
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   163
      | filter_element (Element.Notes _) = error ("\"notes\" element not allowed in class specification.");
1e1ef69ec0de corrected handling of `constrains` elements
haftmann
parents: 38390
diff changeset
   164
    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
   165
    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
   166
      | 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
   167
      | fold_element_types f (Element.Assumes assms) = fold (fold (fn (t, ts) =>
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   168
          fold_types f t #> (fold o fold_types) f ts) o snd) assms
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   169
    val base_sort = if null inferred_elems then proto_base_sort else
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   170
      case (fold o fold_element_types) Term.add_tfreesT inferred_elems []
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   171
       of [] => error "No type variable in class specification"
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   172
        | [(_, sort)] => sort
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   173
        | _ => error "Multiple type variables in class specification";
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   174
    val supparams = map (fn ((c, T), _) =>
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   175
      (c, map_atyps (K (TFree (Name.aT, base_sort))) T)) raw_supparams;
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   176
    val supparam_names = map fst supparams;
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   177
    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
   178
    val supexpr = (map (fn sup => (sup, (("", false),
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   179
      Expression.Positional (map (SOME o mk_param) (Locale.params_of thy sup))))) sups,
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   180
        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
   181
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   182
  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
   183
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   184
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
   185
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
   186
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   187
fun prep_class_spec prep_class prep_class_elems thy raw_supclasses raw_elems =
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   188
  let
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   189
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   190
    (* prepare import *)
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   191
    val inter_sort = curry (Sorts.inter_sort (Sign.classes_of thy));
29608
564ea783ace8 no base sort in class import
haftmann
parents: 29575
diff changeset
   192
    val sups = map (prep_class thy) raw_supclasses
564ea783ace8 no base sort in class import
haftmann
parents: 29575
diff changeset
   193
      |> Sign.minimize_sort thy;
38379
67d71449e85b more convenient split of class modules: class and class_declaration
haftmann
parents: 38376
diff changeset
   194
    val _ = case filter_out (Class.is_class thy) sups
29608
564ea783ace8 no base sort in class import
haftmann
parents: 29575
diff changeset
   195
     of [] => ()
29797
08ef36ed2f8a handling type classes without parameters
haftmann
parents: 29702
diff changeset
   196
      | no_classes => error ("No (proper) classes: " ^ commas (map quote no_classes));
38379
67d71449e85b more convenient split of class modules: class and class_declaration
haftmann
parents: 38376
diff changeset
   197
    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
   198
    val raw_supparam_names = map fst raw_supparams;
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   199
    val _ = if has_duplicates (op =) raw_supparam_names
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   200
      then error ("Duplicate parameter(s) in superclasses: "
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   201
        ^ (commas o map quote o duplicates (op =)) raw_supparam_names)
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   202
      else ();
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   203
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   204
    (* infer types and base sort *)
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   205
    val (base_sort, supparam_names, supexpr, inferred_elems) =
36460
c643b23e8592 empty class specifcations observe default sort
haftmann
parents: 35845
diff changeset
   206
      prep_class_elems thy sups raw_elems;
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   207
    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
   208
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   209
    (* process elements as class specification *)
38379
67d71449e85b more convenient split of class modules: class and class_declaration
haftmann
parents: 38376
diff changeset
   210
    val class_ctxt = Class.begin sups base_sort (ProofContext.init_global thy);
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   211
    val ((_, _, syntax_elems), _) = class_ctxt
29702
a7512f22e916 proper declared constants in class expressions
haftmann
parents: 29665
diff changeset
   212
      |> Expression.cert_declaration supexpr I inferred_elems;
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   213
    fun check_vars e vs = if null vs
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   214
      then error ("No type variable in part of specification element "
36746
6e7704471eaa tuned error message: regular Pretty.string_of instead of raw Pretty.output;
wenzelm
parents: 36745
diff changeset
   215
        ^ (Pretty.string_of o 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
   216
      else ();
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   217
    fun check_element (e as Element.Fixes fxs) =
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   218
          map (fn (_, SOME T, _) => check_vars e (Term.add_tfreesT T [])) fxs
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   219
      | check_element (e as Element.Assumes assms) =
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   220
          maps (fn (_, ts_pss) => map
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   221
            (fn (t, _) => check_vars e (Term.add_tfrees t [])) ts_pss) assms
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   222
      | check_element e = [()];
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   223
    val _ = map check_element syntax_elems;
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   224
    fun fork_syn (Element.Fixes xs) =
30344
10a67c5ddddb more uniform handling of binding in targets and derived elements;
wenzelm
parents: 30335
diff changeset
   225
          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
   226
          #>> Element.Fixes
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   227
      | fork_syn x = pair x;
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   228
    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
   229
32713
b8381161adb1 dropped dead code
haftmann
parents: 32206
diff changeset
   230
  in (((sups, supparam_names), (sup_sort, base_sort, supexpr)), (elems, global_syntax)) end;
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   231
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   232
val cert_class_spec = prep_class_spec (K I) cert_class_elems;
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   233
val read_class_spec = prep_class_spec Sign.intern_class read_class_elems;
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   234
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   235
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   236
(* class establishment *)
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   237
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   238
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
   239
  let
29816
78e0a90694dd more robust failure in error situations
haftmann
parents: 29797
diff changeset
   240
    (*FIXME simplify*)
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   241
    val supconsts = supparam_names
38379
67d71449e85b more convenient split of class modules: class and class_declaration
haftmann
parents: 38376
diff changeset
   242
      |> AList.make (snd o the o AList.lookup (op =) (Class.these_params thy sups))
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   243
      |> (map o apsnd o apsnd o map_atyps o K o TFree) (Name.aT, [class]);
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   244
    val all_params = Locale.params_of thy class;
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   245
    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
   246
    fun add_const ((raw_c, raw_ty), _) thy =
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   247
      let
30755
7ef503d216c2 simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
wenzelm
parents: 30585
diff changeset
   248
        val b = Binding.name raw_c;
30344
10a67c5ddddb more uniform handling of binding in targets and derived elements;
wenzelm
parents: 30335
diff changeset
   249
        val c = Sign.full_name thy b;
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   250
        val ty = map_atyps (K (TFree (Name.aT, base_sort))) raw_ty;
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   251
        val ty0 = Type.strip_sorts ty;
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   252
        val ty' = map_atyps (K (TFree (Name.aT, [class]))) ty0;
30344
10a67c5ddddb more uniform handling of binding in targets and derived elements;
wenzelm
parents: 30335
diff changeset
   253
        val syn = (the_default NoSyn o AList.lookup Binding.eq_name global_syntax) b;
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   254
      in
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   255
        thy
33173
b8ca12f6681a eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
wenzelm
parents: 32970
diff changeset
   256
        |> Sign.declare_const ((b, ty0), syn)
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   257
        |> snd
30585
6b2ba4666336 use Name.of_binding for basic logical entities without name space (fixes, case names etc.);
wenzelm
parents: 30344
diff changeset
   258
        |> pair ((Name.of_binding b, ty), (c, ty'))
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   259
      end;
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   260
  in
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   261
    thy
38379
67d71449e85b more convenient split of class modules: class and class_declaration
haftmann
parents: 38376
diff changeset
   262
    |> Sign.add_path (Class.class_prefix class)
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   263
    |> fold_map add_const raw_params
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   264
    ||> Sign.restore_naming thy
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   265
    |-> (fn params => pair (supconsts @ (map o apfst) fst params, params))
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   266
  end;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   267
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   268
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
   269
  let
29816
78e0a90694dd more robust failure in error situations
haftmann
parents: 29797
diff changeset
   270
    (*FIXME simplify*)
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   271
    fun globalize param_map = map_aterms
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   272
      (fn Free (v, ty) => Const ((fst o the o AList.lookup (op =) param_map) v, ty)
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   273
        | t => t);
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   274
    val raw_pred = Locale.intros_of thy class
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   275
      |> 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
   276
      |> Option.map (Logic.unvarify_global o Logic.strip_imp_concl o Thm.prop_of);
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   277
    fun get_axiom thy = case (#axioms o AxClass.get_info thy) class
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   278
     of [] => NONE
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   279
      | [thm] => SOME thm;
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   280
  in
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   281
    thy
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   282
    |> add_consts class base_sort sups supparam_names global_syntax
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   283
    |-> (fn (param_map, params) => AxClass.define_class (bname, supsort)
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   284
          (map (fst o snd) params)
30211
556d1810cdad Thm.binding;
wenzelm
parents: 29816
diff changeset
   285
          [(Thm.empty_binding, Option.map (globalize param_map) raw_pred |> the_list)]
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   286
    #> snd
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   287
    #> `get_axiom
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   288
    #-> (fn assm_axiom => fold (Sign.add_const_constraint o apsnd SOME o snd) params
29526
0b32c8b84d3e code cleanup
haftmann
parents: 29509
diff changeset
   289
    #> pair (param_map, params, assm_axiom)))
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   290
  end;
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   291
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
   292
fun gen_class prep_class_spec before_exit b raw_supclasses raw_elems thy =
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   293
  let
36635
080b755377c0 locale predicates of classes carry a mandatory "class" prefix
haftmann
parents: 36464
diff changeset
   294
    val class = Sign.full_name thy b;
32206
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   295
    val (((sups, supparam_names), (supsort, base_sort, supexpr)), (elems, global_syntax)) =
b2e93cda0be8 improved handling of parameter import; tuned
haftmann
parents: 32113
diff changeset
   296
      prep_class_spec thy raw_supclasses raw_elems;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   297
  in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   298
    thy
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
   299
    |> Expression.add_locale I b (Binding.qualify true "class" b) supexpr elems
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33553
diff changeset
   300
    |> snd |> Local_Theory.exit_global
36635
080b755377c0 locale predicates of classes carry a mandatory "class" prefix
haftmann
parents: 36464
diff changeset
   301
    |> adjungate_axclass b class base_sort sups supsort supparam_names global_syntax
31696
8b3dac635907 skip_proof where appropriate
haftmann
parents: 31637
diff changeset
   302
    ||> Theory.checkpoint
29526
0b32c8b84d3e code cleanup
haftmann
parents: 29509
diff changeset
   303
    |-> (fn (param_map, params, assm_axiom) =>
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
   304
       `(fn thy => calculate thy class sups base_sort param_map assm_axiom)
32886
aba29da80c1b do not use Locale.add_registration_eqs any longer
haftmann
parents: 32850
diff changeset
   305
    #-> (fn (base_morph, eq_morph, export_morph, axiom, assm_intro, of_class) =>
38107
3a46cebd7983 Make registrations generic data.
ballarin
parents: 37101
diff changeset
   306
       Context.theory_map (Locale.add_registration (class, base_morph)
3a46cebd7983 Make registrations generic data.
ballarin
parents: 37101
diff changeset
   307
         (Option.map (rpair true) eq_morph) export_morph)
38379
67d71449e85b more convenient split of class modules: class and class_declaration
haftmann
parents: 38376
diff changeset
   308
    #> Class.register class sups params base_sort base_morph export_morph axiom assm_intro of_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
   309
    |> Named_Target.init before_exit class
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   310
    |> pair class
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   311
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   312
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   313
in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   314
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   315
val class = gen_class cert_class_spec;
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   316
val class_cmd = gen_class read_class_spec;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   317
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   318
end; (*local*)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   319
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   320
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   321
(** subclass relations **)
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   322
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   323
local
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   324
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
   325
fun gen_subclass prep_class do_proof before_exit raw_sup lthy =
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   326
  let
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   327
    val thy = ProofContext.theory_of lthy;
29558
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   328
    val proto_sup = prep_class thy raw_sup;
38350
480b2de9927c renamed Theory_Target to the more appropriate Named_Target
haftmann
parents: 38107
diff changeset
   329
    val proto_sub = case Named_Target.peek lthy
38390
cb72d89bb444 named target is optional
haftmann
parents: 38389
diff changeset
   330
     of SOME {target, is_class = true, ...} => target
cb72d89bb444 named target is optional
haftmann
parents: 38389
diff changeset
   331
      | _ => error "Not in a class target";
31987
haftmann
parents: 31944
diff changeset
   332
    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
   333
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   334
    val expr = ([(sup, (("", false), Expression.Positional []))], []);
29558
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   335
    val (([props], deps, export), goal_ctxt) =
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   336
      Expression.cert_goal_expression expr lthy;
29526
0b32c8b84d3e code cleanup
haftmann
parents: 29509
diff changeset
   337
    val some_prop = try the_single props;
29558
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   338
    val some_dep_morph = try the_single (map snd deps);
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   339
    fun after_qed some_wit =
38756
d07959fabde6 renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 38493
diff changeset
   340
      ProofContext.background_theory (Class.register_subclass (sub, sup)
29558
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   341
        some_dep_morph some_wit export)
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
   342
      #> ProofContext.theory_of #> Named_Target.init before_exit sub;
29558
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   343
  in do_proof after_qed some_prop goal_ctxt end;
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   344
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   345
fun user_proof after_qed some_prop =
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   346
  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
   347
    [the_list some_prop];
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   348
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   349
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
   350
  after_qed (Option.map
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   351
    (fn prop => Element.prove_witness ctxt prop tac) some_prop) ctxt;
28666
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   352
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   353
in
28666
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   354
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   355
val subclass = gen_subclass (K I) user_proof;
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
   356
fun prove_subclass before_exit tac = gen_subclass (K I) (tactic_proof tac) before_exit;
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36464
diff changeset
   357
val subclass_cmd = gen_subclass (ProofContext.read_class o ProofContext.init_global) user_proof;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   358
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   359
end; (*local*)
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   360
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   361
end;