src/Pure/Isar/class.ML
author haftmann
Wed, 28 Jan 2009 16:35:42 +0100
changeset 29665 2b956243d123
parent 29632 c3d576157244
child 29702 a7512f22e916
permissions -rw-r--r--
explicit check for exactly one type variable in class specification elements
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
     1
(*  Title:      Pure/Isar/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
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
     4
Type classes derived from primitive axclasses and locales - interfaces.
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     5
*)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     6
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     7
signature CLASS =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     8
sig
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
     9
  include CLASS_TARGET
29439
83601bdadae8 construct explicit class morphism
haftmann
parents: 29378
diff changeset
    10
    (*FIXME the split into class_target.ML, theory_target.ML and
83601bdadae8 construct explicit class morphism
haftmann
parents: 29378
diff changeset
    11
      class.ML is artificial*)
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
    12
26247
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
    13
  val class: bstring -> class list -> Element.context_i list
29378
2821c2c5270d proper local_theory after Class.class
haftmann
parents: 29360
diff changeset
    14
    -> theory -> string * local_theory
26247
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
    15
  val class_cmd: bstring -> xstring list -> Element.context list
29378
2821c2c5270d proper local_theory after Class.class
haftmann
parents: 29360
diff changeset
    16
    -> theory -> string * local_theory
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
    17
  val prove_subclass: tactic -> class -> local_theory -> local_theory
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
    18
  val subclass: class -> local_theory -> Proof.state
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
    19
  val subclass_cmd: xstring -> local_theory -> Proof.state
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    20
end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    21
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    22
structure Class : CLASS =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    23
struct
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    24
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
    25
open Class_Target;
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
    26
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
    27
(** class definitions **)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    28
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    29
local
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    30
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
    31
(* calculating class-related rules including canonical interpretation *)
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
    32
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    33
fun calculate thy class sups base_sort param_map assm_axiom =
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    34
  let
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    35
    val empty_ctxt = ProofContext.init thy;
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    36
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    37
    (* instantiation of canonical interpretation *)
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
    38
    val aT = TFree (Name.aT, base_sort);
29627
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    39
    val param_map_const = (map o apsnd) Const param_map;
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    40
    val param_map_inst = (map o apsnd)
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    41
      (Const o apsnd (map_atyps (K aT))) param_map;
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    42
    val const_morph = Element.inst_morphism thy
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    43
      (Symtab.empty, Symtab.make param_map_inst);
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    44
    val (([props], [(_, inst_morph)], export_morph), _) = empty_ctxt
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    45
      |> Expression.cert_goal_expression ([(class, (("", false),
29627
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    46
           Expression.Named param_map_const))], []);
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    47
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    48
    (* witness for canonical interpretation *)
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    49
    val prop = try the_single props;
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    50
    val wit = Option.map (fn prop => let
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    51
        val sup_axioms = map_filter (fst o rules thy) sups;
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    52
        val loc_intro_tac = case Locale.intros_of thy class
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    53
          of (_, NONE) => all_tac
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    54
           | (_, SOME intro) => ALLGOALS (Tactic.rtac intro);
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    55
        val tac = loc_intro_tac
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    56
          THEN ALLGOALS (ProofContext.fact_tac (sup_axioms @ the_list assm_axiom))
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    57
      in Element.prove_witness empty_ctxt prop tac end) prop;
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    58
    val axiom = Option.map Element.conclude_witness wit;
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    59
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    60
    (* canonical interpretation *)
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    61
    val base_morph = inst_morph
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    62
      $> Morphism.binding_morphism
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    63
           (Binding.add_prefix false (class_prefix class))
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    64
      $> Element.satisfy_morphism (the_list wit);
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    65
    val defs = these_defs thy sups;
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    66
    val eq_morph = Element.eq_morphism thy defs;
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    67
    val morph = base_morph $> eq_morph;
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
    (* assm_intro *)
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    70
    fun prove_assm_intro thm = 
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    71
      let
29627
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    72
        val ((_, [thm']), _) = Variable.import_thms true [thm] empty_ctxt;
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    73
        val thm'' = Morphism.thm (const_morph $> eq_morph) thm';
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    74
        val tac = ALLGOALS (ProofContext.fact_tac [thm'']);
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    75
      in Goal.prove_global thy [] [] (Thm.prop_of thm'') (K tac) end;
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    76
    val assm_intro = Option.map prove_assm_intro
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    77
      (fst (Locale.intros_of thy class));
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    78
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    79
    (* of_class *)
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    80
    val of_class_prop_concl = Logic.mk_inclass (aT, class);
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    81
    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
    82
      | SOME prop => Logic.mk_implies (Morphism.term const_morph
152ace41f3fb correct proof of assm_intro rule
haftmann
parents: 29624
diff changeset
    83
          ((map_types o map_atyps) (K aT) prop), of_class_prop_concl);
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    84
    val sup_of_classes = map (snd o rules thy) sups;
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    85
    val loc_axiom_intros = map Drule.standard' (Locale.axioms_of thy class);
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    86
    val axclass_intro = #intro (AxClass.get_info thy class);
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    87
    val base_sort_trivs = Drule.sort_triv thy (aT, base_sort);
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    88
    val tac = REPEAT (SOMEGOAL
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    89
      (Tactic.match_tac (axclass_intro :: sup_of_classes
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    90
         @ loc_axiom_intros @ base_sort_trivs)
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    91
           ORELSE' Tactic.assume_tac));
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    92
    val of_class = Goal.prove_global thy [] [] of_class_prop (K tac);
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    93
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    94
  in (base_morph, morph, export_morph, axiom, assm_intro, of_class) end;
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
    95
29632
c3d576157244 fixed reading of class specs: declare class operations in context
haftmann
parents: 29627
diff changeset
    96
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
    97
(* reading and processing class specifications *)
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
    98
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
    99
fun prep_class_elems prep_decl thy supexpr sups proto_base_sort raw_elems =
29632
c3d576157244 fixed reading of class specs: declare class operations in context
haftmann
parents: 29627
diff changeset
   100
  let
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   101
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   102
    (* user space type system: only permits 'a type variable, improves towards 'a *)
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   103
    val base_constraints = (map o apsnd)
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   104
      (map_type_tfree (K (TVar ((Name.aT, 0), proto_base_sort))) o fst o snd)
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   105
        (these_operations thy sups);
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   106
    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
   107
          if v = Name.aT then T
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   108
          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
   109
      | T => T);
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   110
    fun singleton_fixate thy algebra Ts =
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   111
      let
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   112
        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
   113
        val tfrees = extract
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   114
          (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
   115
        val inferred_sort = extract
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   116
          (fn TVar (_, sort) => curry (Sorts.inter_sort algebra) sort | _ => I);
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   117
        val fixate_sort = if null tfrees then inferred_sort
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   118
          else let val a_sort = (snd o the_single) tfrees
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   119
          in if Sorts.sort_le algebra (a_sort, inferred_sort)
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   120
            then Sorts.inter_sort algebra (a_sort, inferred_sort)
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   121
            else error ("Type inference imposes additional sort constraint "
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   122
              ^ Syntax.string_of_sort_global thy inferred_sort
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   123
              ^ " of type parameter " ^ Name.aT ^ " of sort "
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   124
              ^ Syntax.string_of_sort_global thy a_sort)
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   125
          end
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   126
      in (map o map_atyps) (K (TFree (Name.aT, fixate_sort))) Ts end;
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   127
    fun add_typ_check level name f = Context.proof_map (Syntax.add_typ_check level name (fn xs => fn ctxt =>
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   128
      let val xs' = f xs in if eq_list (op =) (xs, xs') then NONE else SOME (xs', ctxt) end));
29632
c3d576157244 fixed reading of class specs: declare class operations in context
haftmann
parents: 29627
diff changeset
   129
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   130
    (* preproceesing elements, retrieving base sort from type-checked elements *)
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   131
    val ((_, _, inferred_elems), _) = ProofContext.init thy
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   132
      |> fold (ProofContext.add_const_constraint o apsnd SOME) base_constraints
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   133
      |> redeclare_operations thy sups
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   134
      |> add_typ_check 10 "reject_bcd_etc" reject_bcd_etc
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   135
      |> add_typ_check ~10 "singleton_fixate" (singleton_fixate thy (Sign.classes_of thy))
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   136
      |> prep_decl supexpr raw_elems;
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   137
    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
   138
      | 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
   139
      | 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
   140
          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
   141
      | fold_element_types f (Element.Defines _) =
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   142
          error ("\"defines\" element not allowed in class specification.")
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   143
      | fold_element_types f (Element.Notes _) =
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   144
          error ("\"notes\" element not allowed in class specification.");
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   145
    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
   146
      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
   147
       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
   148
        | [(_, sort)] => sort
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   149
        | _ => error "Multiple type variables in class specification"
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   150
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   151
  in (base_sort, inferred_elems) end;
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   152
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   153
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
   154
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
   155
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   156
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
   157
  let
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   158
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   159
    (* prepare import *)
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   160
    val inter_sort = curry (Sorts.inter_sort (Sign.classes_of thy));
29608
564ea783ace8 no base sort in class import
haftmann
parents: 29575
diff changeset
   161
    val sups = map (prep_class thy) raw_supclasses
564ea783ace8 no base sort in class import
haftmann
parents: 29575
diff changeset
   162
      |> Sign.minimize_sort thy;
564ea783ace8 no base sort in class import
haftmann
parents: 29575
diff changeset
   163
    val _ = case filter_out (is_class thy) sups
564ea783ace8 no base sort in class import
haftmann
parents: 29575
diff changeset
   164
     of [] => ()
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   165
      | no_classes => error ("No proper classes: " ^ commas (map quote no_classes));
29608
564ea783ace8 no base sort in class import
haftmann
parents: 29575
diff changeset
   166
          val supparams = (map o apsnd) (snd o snd) (these_params thy sups);
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   167
    val supparam_names = map fst supparams;
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   168
    val _ = if has_duplicates (op =) supparam_names
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   169
      then error ("Duplicate parameter(s) in superclasses: "
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   170
        ^ (commas o map quote o duplicates (op =)) supparam_names)
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   171
      else ();
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   172
    val supexpr = (map (fn sup => (sup, (("", false), Expression.Positional [])))
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   173
      sups, []);
29608
564ea783ace8 no base sort in class import
haftmann
parents: 29575
diff changeset
   174
    val given_basesort = fold inter_sort (map (base_sort thy) sups) [];
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   175
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   176
    (* infer types and base sort *)
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   177
    val (base_sort, inferred_elems) = prep_class_elems thy supexpr sups
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   178
      given_basesort raw_elems;
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   179
    val sup_sort = inter_sort base_sort sups
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   180
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   181
    (* process elements as class specification *)
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   182
    val class_ctxt = begin sups base_sort (ProofContext.init thy)
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   183
    val ((_, _, syntax_elems), _) = class_ctxt
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   184
      (*#> fold (Variable.declare_constraints o Free) ((map o apsnd o map_atyps)
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   185
          (K (TFree (Name.aT, base_sort))) supparams) FIXME really necessary?*)
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   186
            (*FIXME should constraints be issued in begin?*)
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   187
      |> Expression.cert_declaration supexpr inferred_elems;
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   188
    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
   189
      then error ("No type variable in part of specification element "
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   190
        ^ (Pretty.output o Pretty.chunks) (Element.pretty_ctxt class_ctxt e))
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   191
      else ();
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   192
    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
   193
          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
   194
      | check_element (e as Element.Assumes assms) =
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   195
          maps (fn (_, ts_pss) => map
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   196
            (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
   197
      | check_element e = [()];
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   198
    val _ = map check_element syntax_elems;
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   199
    fun fork_syn (Element.Fixes xs) =
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   200
          fold_map (fn (c, ty, syn) => cons (Binding.base_name c, syn) #> pair (c, ty, NoSyn)) xs
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   201
          #>> Element.Fixes
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   202
      | fork_syn x = pair x;
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   203
    val (elems, global_syntax) = fold_map fork_syn syntax_elems [];
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   204
    val constrain = Element.Constrains ((map o apsnd o map_atyps)
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   205
      (K (TFree (Name.aT, base_sort))) supparams);
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   206
      (*FIXME 2009 perhaps better: control type variable by explicit
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   207
      parameter instantiation of import expression*)
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   208
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   209
  in (((sups, supparam_names), (sup_sort, base_sort, supexpr)), (constrain :: elems, global_syntax)) end;
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   210
29665
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   211
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
   212
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
   213
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   214
2b956243d123 explicit check for exactly one type variable in class specification elements
haftmann
parents: 29632
diff changeset
   215
(* class establishment *)
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   216
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   217
fun add_consts bname class base_sort sups supparams global_syntax thy =
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   218
  let
29526
0b32c8b84d3e code cleanup
haftmann
parents: 29509
diff changeset
   219
    (*FIXME 2009 simplify*)
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   220
    val supconsts = supparams
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   221
      |> AList.make (snd o the o AList.lookup (op =) (these_params thy sups))
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   222
      |> (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
   223
    val all_params = Locale.params_of thy class;
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   224
    val raw_params = (snd o chop (length supparams)) all_params;
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   225
    fun add_const (b, SOME raw_ty, _) thy =
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   226
      let
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   227
        val v = Binding.base_name b;
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28941
diff changeset
   228
        val c = Sign.full_bname thy v;
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   229
        val ty = map_atyps (K (TFree (Name.aT, base_sort))) raw_ty;
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   230
        val ty0 = Type.strip_sorts ty;
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   231
        val ty' = map_atyps (K (TFree (Name.aT, [class]))) ty0;
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   232
        val syn = (the_default NoSyn o AList.lookup (op =) global_syntax) v;
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   233
      in
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   234
        thy
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28941
diff changeset
   235
        |> Sign.declare_const [] ((Binding.name v, ty0), syn)
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   236
        |> snd
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   237
        |> pair ((v, ty), (c, ty'))
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   238
      end;
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   239
  in
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   240
    thy
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
   241
    |> Sign.add_path (class_prefix class)
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   242
    |> fold_map add_const raw_params
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   243
    ||> Sign.restore_naming thy
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   244
    |-> (fn params => pair (supconsts @ (map o apfst) fst params, params))
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   245
  end;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   246
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   247
fun adjungate_axclass bname class base_sort sups supsort supparams global_syntax thy =
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   248
  let
29526
0b32c8b84d3e code cleanup
haftmann
parents: 29509
diff changeset
   249
    (*FIXME 2009 simplify*)
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   250
    fun globalize param_map = map_aterms
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   251
      (fn Free (v, ty) => Const ((fst o the o AList.lookup (op =) param_map) v, ty)
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   252
        | t => t);
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   253
    val raw_pred = Locale.intros_of thy class
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   254
      |> fst
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   255
      |> Option.map (Logic.unvarify o Logic.strip_imp_concl o Thm.prop_of);
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   256
    fun get_axiom thy = case (#axioms o AxClass.get_info thy) class
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   257
     of [] => NONE
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   258
      | [thm] => SOME thm;
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   259
  in
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   260
    thy
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   261
    |> add_consts bname class base_sort sups supparams global_syntax
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   262
    |-> (fn (param_map, params) => AxClass.define_class (bname, supsort)
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   263
          (map (fst o snd) params)
29526
0b32c8b84d3e code cleanup
haftmann
parents: 29509
diff changeset
   264
          [((Binding.empty, []),
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   265
            Option.map (globalize param_map) raw_pred |> the_list)]
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   266
    #> snd
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   267
    #> `get_axiom
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   268
    #-> (fn assm_axiom => fold (Sign.add_const_constraint o apsnd SOME o snd) params
29526
0b32c8b84d3e code cleanup
haftmann
parents: 29509
diff changeset
   269
    #> pair (param_map, params, assm_axiom)))
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   270
  end;
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   271
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   272
fun gen_class prep_spec bname raw_supclasses raw_elems thy =
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   273
  let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28941
diff changeset
   274
    val class = Sign.full_bname thy bname;
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   275
    val (((sups, supparams), (supsort, base_sort, supexpr)), (elems, global_syntax)) =
26247
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
   276
      prep_spec thy raw_supclasses raw_elems;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   277
  in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   278
    thy
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   279
    |> Expression.add_locale bname "" supexpr elems
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   280
    |> snd |> LocalTheory.exit_global
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   281
    |> adjungate_axclass bname class base_sort sups supsort supparams global_syntax
29526
0b32c8b84d3e code cleanup
haftmann
parents: 29509
diff changeset
   282
    |-> (fn (param_map, params, assm_axiom) =>
29547
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
   283
       `(fn thy => calculate thy class sups base_sort param_map assm_axiom)
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
   284
    #-> (fn (base_morph, morph, export_morph, axiom, assm_intro, of_class) =>
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
   285
       Locale.add_registration (class, (morph, export_morph))
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
   286
    #> Locale.activate_global_facts (class, morph $> export_morph)
f2587922591e improved calculation of morphisms and rules
haftmann
parents: 29526
diff changeset
   287
    #> register class sups params base_sort base_morph axiom assm_intro of_class))
29378
2821c2c5270d proper local_theory after Class.class
haftmann
parents: 29360
diff changeset
   288
    |> TheoryTarget.init (SOME class)
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   289
    |> pair class
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   290
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   291
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   292
in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   293
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   294
val class = gen_class cert_class_spec;
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   295
val class_cmd = gen_class read_class_spec;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   296
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   297
end; (*local*)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   298
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   299
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   300
(** subclass relations **)
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   301
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   302
local
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   303
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   304
fun gen_subclass prep_class do_proof raw_sup lthy =
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   305
  let
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   306
    val thy = ProofContext.theory_of lthy;
29558
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   307
    val proto_sup = prep_class thy raw_sup;
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   308
    val proto_sub = case TheoryTarget.peek lthy
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   309
     of {is_class = false, ...} => error "Not in a class context"
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   310
      | {target, ...} => target;
29558
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   311
    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
   312
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   313
    val expr = ([(sup, (("", false), Expression.Positional []))], []);
29558
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   314
    val (([props], deps, export), goal_ctxt) =
29509
1ff0f3f08a7b migrated class package to new locale implementation
haftmann
parents: 29439
diff changeset
   315
      Expression.cert_goal_expression expr lthy;
29526
0b32c8b84d3e code cleanup
haftmann
parents: 29509
diff changeset
   316
    val some_prop = try the_single props;
29558
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   317
    val some_dep_morph = try the_single (map snd deps);
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   318
    fun after_qed some_wit =
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   319
      ProofContext.theory (register_subclass (sub, sup)
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   320
        some_dep_morph some_wit export)
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   321
      #> ProofContext.theory_of #> TheoryTarget.init (SOME sub);
9846af6c6d6a improved tackling of subclasses
haftmann
parents: 29547
diff changeset
   322
  in do_proof after_qed some_prop goal_ctxt end;
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   323
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   324
fun user_proof after_qed some_prop =
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   325
  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
   326
    [the_list some_prop];
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   327
29575
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   328
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
   329
  after_qed (Option.map
41d604e59e93 improved and corrected reading of class specs -- still draft version
haftmann
parents: 29558
diff changeset
   330
    (fn prop => Element.prove_witness ctxt prop tac) some_prop) ctxt;
28666
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   331
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   332
in
28666
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   333
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   334
val subclass = gen_subclass (K I) user_proof;
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   335
fun prove_subclass tac = gen_subclass (K I) (tactic_proof tac);
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   336
val subclass_cmd = gen_subclass Sign.read_class user_proof;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   337
29358
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   338
end; (*local*)
efdfe5dfe008 rearranged target theories
haftmann
parents: 29333
diff changeset
   339
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   340
end;