src/Pure/Isar/class.ML
author wenzelm
Sat, 08 May 2010 19:14:13 +0200
changeset 36754 403585a89772
parent 36674 d95f39448121
child 36755 6e7704471eaa
permissions -rw-r--r--
unified/simplified Pretty.margin_default;
discontinued special Pretty.setmargin etc;
explicit margin argument for Pretty.string_of_margin etc.;
     1 (*  Title:      Pure/Isar/class.ML
     2     Author:     Florian Haftmann, TU Muenchen
     3 
     4 Type classes derived from primitive axclasses and locales -- interfaces.
     5 *)
     6 
     7 signature CLASS =
     8 sig
     9   include CLASS_TARGET
    10     (*FIXME the split into class_target.ML, theory_target.ML and
    11       class.ML is artificial*)
    12 
    13   val class: binding -> class list -> Element.context_i list
    14     -> theory -> string * local_theory
    15   val class_cmd: binding -> xstring list -> Element.context list
    16     -> theory -> string * local_theory
    17   val prove_subclass: tactic -> class -> local_theory -> local_theory
    18   val subclass: class -> local_theory -> Proof.state
    19   val subclass_cmd: xstring -> local_theory -> Proof.state
    20 end;
    21 
    22 structure Class : CLASS =
    23 struct
    24 
    25 open Class_Target;
    26 
    27 (** class definitions **)
    28 
    29 local
    30 
    31 (* calculating class-related rules including canonical interpretation *)
    32 
    33 fun calculate thy class sups base_sort param_map assm_axiom =
    34   let
    35     val empty_ctxt = ProofContext.init_global thy;
    36 
    37     (* instantiation of canonical interpretation *)
    38     val aT = TFree (Name.aT, base_sort);
    39     val param_map_const = (map o apsnd) Const param_map;
    40     val param_map_inst = (map o apsnd)
    41       (Const o apsnd (map_atyps (K aT))) param_map;
    42     val const_morph = Element.inst_morphism thy
    43       (Symtab.empty, Symtab.make param_map_inst);
    44     val typ_morph = Element.inst_morphism thy
    45       (Symtab.empty |> Symtab.update (Name.aT, TFree (Name.aT, [class])), Symtab.empty);
    46     val (([raw_props], [(_, raw_inst_morph)], export_morph), _) = empty_ctxt
    47       |> Expression.cert_goal_expression ([(class, (("", false),
    48            Expression.Named param_map_const))], []);
    49     val (props, inst_morph) = if null param_map
    50       then (raw_props |> map (Morphism.term typ_morph),
    51         raw_inst_morph $> typ_morph)
    52       else (raw_props, raw_inst_morph); (*FIXME proper handling in
    53         locale.ML / expression.ML would be desirable*)
    54 
    55     (* witness for canonical interpretation *)
    56     val prop = try the_single props;
    57     val wit = Option.map (fn prop => let
    58         val sup_axioms = map_filter (fst o rules thy) sups;
    59         val loc_intro_tac = case Locale.intros_of thy class
    60           of (_, NONE) => all_tac
    61            | (_, SOME intro) => ALLGOALS (Tactic.rtac intro);
    62         val tac = loc_intro_tac
    63           THEN ALLGOALS (ProofContext.fact_tac (sup_axioms @ the_list assm_axiom))
    64       in Element.prove_witness empty_ctxt prop tac end) prop;
    65     val axiom = Option.map Element.conclude_witness wit;
    66 
    67     (* canonical interpretation *)
    68     val base_morph = inst_morph
    69       $> Morphism.binding_morphism (Binding.prefix false (class_prefix class))
    70       $> Element.satisfy_morphism (the_list wit);
    71     val eq_morph = Element.eq_morphism thy (these_defs thy sups);
    72 
    73     (* assm_intro *)
    74     fun prove_assm_intro thm =
    75       let
    76         val ((_, [thm']), _) = Variable.import true [thm] empty_ctxt;
    77         val const_eq_morph = case eq_morph
    78          of SOME eq_morph => const_morph $> eq_morph
    79           | NONE => const_morph
    80         val thm'' = Morphism.thm const_eq_morph thm';
    81         val tac = ALLGOALS (ProofContext.fact_tac [thm'']);
    82       in Skip_Proof.prove_global thy [] [] (Thm.prop_of thm'') (K tac) end;
    83     val assm_intro = Option.map prove_assm_intro
    84       (fst (Locale.intros_of thy class));
    85 
    86     (* of_class *)
    87     val of_class_prop_concl = Logic.mk_of_class (aT, class);
    88     val of_class_prop = case prop of NONE => of_class_prop_concl
    89       | SOME prop => Logic.mk_implies (Morphism.term const_morph
    90           ((map_types o map_atyps) (K aT) prop), of_class_prop_concl);
    91     val sup_of_classes = map (snd o rules thy) sups;
    92     val loc_axiom_intros = map Drule.export_without_context_open (Locale.axioms_of thy class);
    93     val axclass_intro = #intro (AxClass.get_info thy class);
    94     val base_sort_trivs = Thm.of_sort (Thm.ctyp_of thy aT, base_sort);
    95     val tac = REPEAT (SOMEGOAL
    96       (Tactic.match_tac (axclass_intro :: sup_of_classes
    97          @ loc_axiom_intros @ base_sort_trivs)
    98            ORELSE' Tactic.assume_tac));
    99     val of_class = Skip_Proof.prove_global thy [] [] of_class_prop (K tac);
   100 
   101   in (base_morph, eq_morph, export_morph, axiom, assm_intro, of_class) end;
   102 
   103 
   104 (* reading and processing class specifications *)
   105 
   106 fun prep_class_elems prep_decl thy sups raw_elems =
   107   let
   108 
   109     (* user space type system: only permits 'a type variable, improves towards 'a *)
   110     val algebra = Sign.classes_of thy;
   111     val inter_sort = curry (Sorts.inter_sort algebra);
   112     val proto_base_sort = if null sups then Sign.defaultS thy
   113       else fold inter_sort (map (base_sort thy) sups) [];
   114     val base_constraints = (map o apsnd)
   115       (map_type_tfree (K (TVar ((Name.aT, 0), proto_base_sort))) o fst o snd)
   116         (these_operations thy sups);
   117     val reject_bcd_etc = (map o map_atyps) (fn T as TFree (v, sort) =>
   118           if v = Name.aT then T
   119           else error ("No type variable other than " ^ Name.aT ^ " allowed in class specification")
   120       | T => T);
   121     fun singleton_fixate Ts =
   122       let
   123         fun extract f = (fold o fold_atyps) f Ts [];
   124         val tfrees = extract
   125           (fn TFree (v, sort) => insert (op =) (v, sort) | _ => I);
   126         val inferred_sort = extract
   127           (fn TVar (_, sort) => inter_sort sort | _ => I);
   128         val fixate_sort = if null tfrees then inferred_sort
   129           else case tfrees
   130            of [(_, a_sort)] => if Sorts.sort_le algebra (a_sort, inferred_sort)
   131                 then inter_sort a_sort inferred_sort
   132                 else error ("Type inference imposes additional sort constraint "
   133                   ^ Syntax.string_of_sort_global thy inferred_sort
   134                   ^ " of type parameter " ^ Name.aT ^ " of sort "
   135                   ^ Syntax.string_of_sort_global thy a_sort ^ ".")
   136             | _ => error "Multiple type variables in class specification.";
   137       in (map o map_atyps) (K (TFree (Name.aT, fixate_sort))) Ts end;
   138     fun add_typ_check level name f = Context.proof_map
   139       (Syntax.add_typ_check level name (fn xs => fn ctxt =>
   140         let val xs' = f xs in if eq_list (op =) (xs, xs') then NONE else SOME (xs', ctxt) end));
   141 
   142     (* preprocessing elements, retrieving base sort from type-checked elements *)
   143     val init_class_body = fold (ProofContext.add_const_constraint o apsnd SOME) base_constraints
   144       #> redeclare_operations thy sups
   145       #> add_typ_check 10 "reject_bcd_etc" reject_bcd_etc
   146       #> add_typ_check ~10 "singleton_fixate" singleton_fixate;
   147     val raw_supexpr = (map (fn sup => (sup, (("", false),
   148       Expression.Positional []))) sups, []);
   149     val ((raw_supparams, _, inferred_elems), _) = ProofContext.init_global thy
   150       |> prep_decl raw_supexpr init_class_body raw_elems;
   151     fun fold_element_types f (Element.Fixes fxs) = fold (fn (_, SOME T, _) => f T) fxs
   152       | fold_element_types f (Element.Constrains cnstrs) = fold (f o snd) cnstrs
   153       | fold_element_types f (Element.Assumes assms) = fold (fold (fn (t, ts) =>
   154           fold_types f t #> (fold o fold_types) f ts) o snd) assms
   155       | fold_element_types f (Element.Defines _) =
   156           error ("\"defines\" element not allowed in class specification.")
   157       | fold_element_types f (Element.Notes _) =
   158           error ("\"notes\" element not allowed in class specification.");
   159     val base_sort = if null inferred_elems then proto_base_sort else
   160       case (fold o fold_element_types) Term.add_tfreesT inferred_elems []
   161        of [] => error "No type variable in class specification"
   162         | [(_, sort)] => sort
   163         | _ => error "Multiple type variables in class specification";
   164     val supparams = map (fn ((c, T), _) =>
   165       (c, map_atyps (K (TFree (Name.aT, base_sort))) T)) raw_supparams;
   166     val supparam_names = map fst supparams;
   167     fun mk_param ((c, _), _) = Free (c, (the o AList.lookup (op =) supparams) c);
   168     val supexpr = (map (fn sup => (sup, (("", false),
   169       Expression.Positional (map (SOME o mk_param) (Locale.params_of thy sup))))) sups,
   170         map (fn (c, T) => (Binding.name c, SOME T, NoSyn)) supparams);
   171 
   172   in (base_sort, supparam_names, supexpr, inferred_elems) end;
   173 
   174 val cert_class_elems = prep_class_elems Expression.cert_declaration;
   175 val read_class_elems = prep_class_elems Expression.cert_read_declaration;
   176 
   177 fun prep_class_spec prep_class prep_class_elems thy raw_supclasses raw_elems =
   178   let
   179 
   180     (* prepare import *)
   181     val inter_sort = curry (Sorts.inter_sort (Sign.classes_of thy));
   182     val sups = map (prep_class thy) raw_supclasses
   183       |> Sign.minimize_sort thy;
   184     val _ = case filter_out (is_class thy) sups
   185      of [] => ()
   186       | no_classes => error ("No (proper) classes: " ^ commas (map quote no_classes));
   187     val raw_supparams = (map o apsnd) (snd o snd) (these_params thy sups);
   188     val raw_supparam_names = map fst raw_supparams;
   189     val _ = if has_duplicates (op =) raw_supparam_names
   190       then error ("Duplicate parameter(s) in superclasses: "
   191         ^ (commas o map quote o duplicates (op =)) raw_supparam_names)
   192       else ();
   193 
   194     (* infer types and base sort *)
   195     val (base_sort, supparam_names, supexpr, inferred_elems) =
   196       prep_class_elems thy sups raw_elems;
   197     val sup_sort = inter_sort base_sort sups;
   198 
   199     (* process elements as class specification *)
   200     val class_ctxt = begin sups base_sort (ProofContext.init_global thy);
   201     val ((_, _, syntax_elems), _) = class_ctxt
   202       |> Expression.cert_declaration supexpr I inferred_elems;
   203     fun check_vars e vs = if null vs
   204       then error ("No type variable in part of specification element "
   205         ^ (Pretty.output NONE o Pretty.chunks) (Element.pretty_ctxt class_ctxt e))
   206       else ();
   207     fun check_element (e as Element.Fixes fxs) =
   208           map (fn (_, SOME T, _) => check_vars e (Term.add_tfreesT T [])) fxs
   209       | check_element (e as Element.Assumes assms) =
   210           maps (fn (_, ts_pss) => map
   211             (fn (t, _) => check_vars e (Term.add_tfrees t [])) ts_pss) assms
   212       | check_element e = [()];
   213     val _ = map check_element syntax_elems;
   214     fun fork_syn (Element.Fixes xs) =
   215           fold_map (fn (c, ty, syn) => cons (c, syn) #> pair (c, ty, NoSyn)) xs
   216           #>> Element.Fixes
   217       | fork_syn x = pair x;
   218     val (elems, global_syntax) = fold_map fork_syn syntax_elems [];
   219 
   220   in (((sups, supparam_names), (sup_sort, base_sort, supexpr)), (elems, global_syntax)) end;
   221 
   222 val cert_class_spec = prep_class_spec (K I) cert_class_elems;
   223 val read_class_spec = prep_class_spec Sign.intern_class read_class_elems;
   224 
   225 
   226 (* class establishment *)
   227 
   228 fun add_consts class base_sort sups supparam_names global_syntax thy =
   229   let
   230     (*FIXME simplify*)
   231     val supconsts = supparam_names
   232       |> AList.make (snd o the o AList.lookup (op =) (these_params thy sups))
   233       |> (map o apsnd o apsnd o map_atyps o K o TFree) (Name.aT, [class]);
   234     val all_params = Locale.params_of thy class;
   235     val raw_params = (snd o chop (length supparam_names)) all_params;
   236     fun add_const ((raw_c, raw_ty), _) thy =
   237       let
   238         val b = Binding.name raw_c;
   239         val c = Sign.full_name thy b;
   240         val ty = map_atyps (K (TFree (Name.aT, base_sort))) raw_ty;
   241         val ty0 = Type.strip_sorts ty;
   242         val ty' = map_atyps (K (TFree (Name.aT, [class]))) ty0;
   243         val syn = (the_default NoSyn o AList.lookup Binding.eq_name global_syntax) b;
   244       in
   245         thy
   246         |> Sign.declare_const ((b, ty0), syn)
   247         |> snd
   248         |> pair ((Name.of_binding b, ty), (c, ty'))
   249       end;
   250   in
   251     thy
   252     |> Sign.add_path (class_prefix class)
   253     |> fold_map add_const raw_params
   254     ||> Sign.restore_naming thy
   255     |-> (fn params => pair (supconsts @ (map o apfst) fst params, params))
   256   end;
   257 
   258 fun adjungate_axclass bname class base_sort sups supsort supparam_names global_syntax thy =
   259   let
   260     (*FIXME simplify*)
   261     fun globalize param_map = map_aterms
   262       (fn Free (v, ty) => Const ((fst o the o AList.lookup (op =) param_map) v, ty)
   263         | t => t);
   264     val raw_pred = Locale.intros_of thy class
   265       |> fst
   266       |> Option.map (Logic.unvarify_global o Logic.strip_imp_concl o Thm.prop_of);
   267     fun get_axiom thy = case (#axioms o AxClass.get_info thy) class
   268      of [] => NONE
   269       | [thm] => SOME thm;
   270   in
   271     thy
   272     |> add_consts class base_sort sups supparam_names global_syntax
   273     |-> (fn (param_map, params) => AxClass.define_class (bname, supsort)
   274           (map (fst o snd) params)
   275           [(Thm.empty_binding, Option.map (globalize param_map) raw_pred |> the_list)]
   276     #> snd
   277     #> `get_axiom
   278     #-> (fn assm_axiom => fold (Sign.add_const_constraint o apsnd SOME o snd) params
   279     #> pair (param_map, params, assm_axiom)))
   280   end;
   281 
   282 fun gen_class prep_class_spec b raw_supclasses raw_elems thy =
   283   let
   284     val class = Sign.full_name thy b;
   285     val (((sups, supparam_names), (supsort, base_sort, supexpr)), (elems, global_syntax)) =
   286       prep_class_spec thy raw_supclasses raw_elems;
   287   in
   288     thy
   289     |> Expression.add_locale b (Binding.qualify true "class" b) supexpr elems
   290     |> snd |> Local_Theory.exit_global
   291     |> adjungate_axclass b class base_sort sups supsort supparam_names global_syntax
   292     ||> Theory.checkpoint
   293     |-> (fn (param_map, params, assm_axiom) =>
   294        `(fn thy => calculate thy class sups base_sort param_map assm_axiom)
   295     #-> (fn (base_morph, eq_morph, export_morph, axiom, assm_intro, of_class) =>
   296        Locale.add_registration (class, case eq_morph of SOME eq_morph => base_morph $> eq_morph | NONE => base_morph)
   297          (*FIXME should avoid modification of base morphism, but then problem with sublocale linorder < distrib_lattice*)
   298          (Option.map (rpair true) eq_morph) export_morph
   299     #> register class sups params base_sort base_morph export_morph axiom assm_intro of_class))
   300     |> Theory_Target.init (SOME class)
   301     |> pair class
   302   end;
   303 
   304 in
   305 
   306 val class = gen_class cert_class_spec;
   307 val class_cmd = gen_class read_class_spec;
   308 
   309 end; (*local*)
   310 
   311 
   312 (** subclass relations **)
   313 
   314 local
   315 
   316 fun gen_subclass prep_class do_proof raw_sup lthy =
   317   let
   318     val thy = ProofContext.theory_of lthy;
   319     val proto_sup = prep_class thy raw_sup;
   320     val proto_sub = case Theory_Target.peek lthy
   321      of {is_class = false, ...} => error "Not in a class context"
   322       | {target, ...} => target;
   323     val (sub, sup) = AxClass.cert_classrel thy (proto_sub, proto_sup);
   324 
   325     val expr = ([(sup, (("", false), Expression.Positional []))], []);
   326     val (([props], deps, export), goal_ctxt) =
   327       Expression.cert_goal_expression expr lthy;
   328     val some_prop = try the_single props;
   329     val some_dep_morph = try the_single (map snd deps);
   330     fun after_qed some_wit =
   331       ProofContext.theory (register_subclass (sub, sup)
   332         some_dep_morph some_wit export)
   333       #> ProofContext.theory_of #> Theory_Target.init (SOME sub);
   334   in do_proof after_qed some_prop goal_ctxt end;
   335 
   336 fun user_proof after_qed some_prop =
   337   Element.witness_proof (after_qed o try the_single o the_single)
   338     [the_list some_prop];
   339 
   340 fun tactic_proof tac after_qed some_prop ctxt =
   341   after_qed (Option.map
   342     (fn prop => Element.prove_witness ctxt prop tac) some_prop) ctxt;
   343 
   344 in
   345 
   346 val subclass = gen_subclass (K I) user_proof;
   347 fun prove_subclass tac = gen_subclass (K I) (tactic_proof tac);
   348 val subclass_cmd = gen_subclass (ProofContext.read_class o ProofContext.init_global) user_proof;
   349 
   350 end; (*local*)
   351 
   352 end;