haftmann@24218: (* Title: Pure/Isar/class.ML haftmann@24218: ID: $Id$ haftmann@24218: Author: Florian Haftmann, TU Muenchen haftmann@24218: haftmann@24218: Type classes derived from primitive axclasses and locales. haftmann@24218: *) haftmann@24218: haftmann@24218: signature CLASS = haftmann@24218: sig haftmann@24218: val fork_mixfix: bool -> string option -> mixfix -> mixfix * mixfix haftmann@24218: haftmann@24218: val axclass_cmd: bstring * xstring list haftmann@24218: -> ((bstring * Attrib.src list) * string list) list -> theory -> class * theory haftmann@24218: val class: bstring -> class list -> Element.context_i Locale.element list haftmann@24218: -> string list -> theory -> string * Proof.context haftmann@24218: val class_cmd: bstring -> string list -> Element.context Locale.element list haftmann@24218: -> string list -> theory -> string * Proof.context haftmann@24218: val instance_arity: arity list -> ((bstring * Attrib.src list) * term) list haftmann@24218: -> theory -> Proof.state haftmann@24218: val instance_arity_cmd: (bstring * string list * string) list haftmann@24218: -> ((bstring * Attrib.src list) * string) list haftmann@24218: -> theory -> Proof.state haftmann@24218: val prove_instance_arity: tactic -> arity list haftmann@24218: -> ((bstring * Attrib.src list) * term) list haftmann@24218: -> theory -> theory haftmann@24218: val instance_class: class * class -> theory -> Proof.state haftmann@24218: val instance_class_cmd: string * string -> theory -> Proof.state haftmann@24218: val instance_sort: class * sort -> theory -> Proof.state haftmann@24218: val instance_sort_cmd: string * string -> theory -> Proof.state haftmann@24218: val prove_instance_sort: tactic -> class * sort -> theory -> theory haftmann@24218: haftmann@24218: val class_of_locale: theory -> string -> class option haftmann@24218: val add_const_in_class: string -> (string * term) * Syntax.mixfix haftmann@24218: -> theory -> theory haftmann@24218: haftmann@24304: val unoverload: theory -> thm -> thm haftmann@24304: val overload: theory -> thm -> thm haftmann@24304: val inst_const: theory -> string * string -> string haftmann@24304: haftmann@24218: val print_classes: theory -> unit haftmann@24218: val intro_classes_tac: thm list -> tactic haftmann@24218: val default_intro_classes_tac: thm list -> tactic haftmann@24218: end; haftmann@24218: haftmann@24218: structure Class : CLASS = haftmann@24218: struct haftmann@24218: haftmann@24218: (** auxiliary **) haftmann@24218: haftmann@24218: fun fork_mixfix is_loc some_class mx = haftmann@24218: let haftmann@24218: val mx' = Syntax.unlocalize_mixfix mx; haftmann@24218: val mx_global = if is_some some_class orelse (is_loc andalso mx = mx') haftmann@24218: then NoSyn else mx'; haftmann@24218: val mx_local = if is_loc then mx else NoSyn; haftmann@24218: in (mx_global, mx_local) end; haftmann@24218: haftmann@24218: fun axclass_cmd (class, raw_superclasses) raw_specs thy = haftmann@24218: let haftmann@24218: val ctxt = ProofContext.init thy; haftmann@24218: val superclasses = map (Sign.read_class thy) raw_superclasses; haftmann@24218: val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst) raw_specs; haftmann@24218: val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd) raw_specs) haftmann@24218: |> snd haftmann@24218: |> (map o map) fst; haftmann@24218: in AxClass.define_class (class, superclasses) [] (name_atts ~~ axiomss) thy end; haftmann@24218: haftmann@24218: haftmann@24218: (** axclasses with implicit parameter handling **) haftmann@24218: haftmann@24218: (* axclass instances *) haftmann@24218: haftmann@24218: local haftmann@24218: haftmann@24218: fun gen_instance mk_prop add_thm after_qed insts thy = haftmann@24218: let haftmann@24218: fun after_qed' results = haftmann@24218: ProofContext.theory ((fold o fold) add_thm results #> after_qed); haftmann@24218: in haftmann@24218: thy haftmann@24218: |> ProofContext.init haftmann@24218: |> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])]) o maps (mk_prop thy)) insts) haftmann@24218: end; haftmann@24218: haftmann@24218: in haftmann@24218: haftmann@24218: val axclass_instance_arity = haftmann@24218: gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity; haftmann@24218: val axclass_instance_sort = haftmann@24218: gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel)) haftmann@24218: AxClass.add_classrel I o single; haftmann@24218: haftmann@24218: end; (*local*) haftmann@24218: haftmann@24218: haftmann@24218: (* introducing axclasses with implicit parameter handling *) haftmann@24218: haftmann@24218: fun axclass_params (name, raw_superclasses) raw_consts raw_dep_axioms other_consts thy = haftmann@24218: let haftmann@24218: val superclasses = map (Sign.certify_class thy) raw_superclasses; haftmann@24218: val consts = (map o apfst o apsnd) (Sign.certify_typ thy) raw_consts; haftmann@24218: val prefix = Logic.const_of_class name; haftmann@24218: fun mk_const_name c = NameSpace.map_base (NameSpace.append prefix) haftmann@24218: (Sign.full_name thy c); haftmann@24218: fun add_const ((c, ty), syn) = haftmann@24218: Sign.add_consts_authentic [(c, ty, syn)] haftmann@24218: #> pair (mk_const_name c, ty); haftmann@24218: fun mk_axioms cs thy = haftmann@24218: raw_dep_axioms thy cs haftmann@24218: |> (map o apsnd o map) (Sign.cert_prop thy) haftmann@24218: |> rpair thy; haftmann@24218: fun add_constraint class (c, ty) = haftmann@24218: Sign.add_const_constraint_i (c, SOME haftmann@24218: (Term.map_type_tfree (fn (v, _) => TFree (v, [class])) ty)); haftmann@24218: in haftmann@24218: thy haftmann@24218: |> Theory.add_path prefix haftmann@24218: |> fold_map add_const consts haftmann@24218: ||> Theory.restore_naming thy haftmann@24218: |-> (fn cs => mk_axioms cs haftmann@24218: #-> (fn axioms_prop => AxClass.define_class (name, superclasses) haftmann@24218: (map fst cs @ other_consts) axioms_prop haftmann@24218: #-> (fn class => `(fn thy => AxClass.get_definition thy class) haftmann@24218: #-> (fn {intro, axioms, ...} => fold (add_constraint class) cs haftmann@24218: #> pair (class, ((intro, (map Thm.prop_of axioms, axioms)), cs)))))) haftmann@24218: end; haftmann@24218: haftmann@24218: haftmann@24304: (* explicit constants for overloaded definitions *) haftmann@24304: haftmann@24304: structure InstData = TheoryDataFun haftmann@24304: ( haftmann@24304: type T = (string * thm) Symtab.table Symtab.table; haftmann@24304: (*constant name ~> type constructor ~> (constant name, equation)*) haftmann@24304: val empty = Symtab.empty; haftmann@24304: val copy = I; haftmann@24304: val extend = I; haftmann@24304: fun merge _ = Symtab.join (K (Symtab.merge (K true))); haftmann@24304: ); haftmann@24304: haftmann@24304: fun inst_thms f thy = haftmann@24304: Symtab.fold (Symtab.fold (cons o f o snd o snd) o snd) (InstData.get thy) []; haftmann@24304: fun add_inst (c, tyco) inst = (InstData.map o Symtab.map_default (c, Symtab.empty)) haftmann@24304: (Symtab.update_new (tyco, inst)); haftmann@24304: haftmann@24304: fun unoverload thy thm = MetaSimplifier.rewrite_rule (inst_thms I thy) thm; haftmann@24304: fun overload thy thm = MetaSimplifier.rewrite_rule (inst_thms symmetric thy) thm; haftmann@24304: haftmann@24304: fun inst_const thy (c, tyco) = haftmann@24304: (fst o the o Symtab.lookup ((the o Symtab.lookup (InstData.get thy)) c)) tyco; haftmann@24304: haftmann@24304: fun add_inst_def (class, tyco) (c, ty) thy = haftmann@24304: let haftmann@24304: val tyco_base = NameSpace.base tyco; haftmann@24304: val name_inst = NameSpace.base class ^ "_" ^ tyco_base ^ "_inst"; haftmann@24304: val c_inst_base = NameSpace.base c ^ "_" ^ tyco_base; haftmann@24304: in haftmann@24304: thy haftmann@24304: |> Sign.sticky_prefix name_inst haftmann@24304: |> Sign.add_consts_i [(c_inst_base, ty, Syntax.NoSyn)] haftmann@24304: |> `(fn thy => Sign.full_name thy c_inst_base) haftmann@24304: |-> (fn c_inst => PureThy.add_defs_i true haftmann@24304: [((Thm.def_name c_inst_base, Logic.mk_equals (Const (c_inst, ty), Const (c, ty))), [])] haftmann@24304: #-> (fn [def] => add_inst (c, tyco) (c_inst, symmetric def)) haftmann@24304: #> Sign.restore_naming thy) haftmann@24304: end; haftmann@24304: haftmann@24304: fun add_inst_def' (class, tyco) (c, ty) thy = haftmann@24304: if case Symtab.lookup (InstData.get thy) c haftmann@24304: of NONE => true haftmann@24304: | SOME tab => is_none (Symtab.lookup tab tyco) haftmann@24304: then add_inst_def (class, tyco) (c, Logic.unvarifyT ty) thy haftmann@24304: else thy; haftmann@24304: haftmann@24304: fun add_def ((class, tyco), ((name, prop), atts)) thy = haftmann@24304: let haftmann@24304: val ((lhs as Const (c, ty), args), rhs) = (apfst Term.strip_comb o Logic.dest_equals) prop; haftmann@24304: fun add_inst' def ([], (Const (c_inst, ty))) = haftmann@24304: if forall (fn TFree_ => true | _ => false) (Sign.const_typargs thy (c_inst, ty)) haftmann@24304: then add_inst (c, tyco) (c_inst, def) haftmann@24304: else add_inst_def (class, tyco) (c, ty) haftmann@24304: | add_inst' _ t = add_inst_def (class, tyco) (c, ty); haftmann@24304: in haftmann@24304: thy haftmann@24304: |> PureThy.add_defs_i true [((name, prop), map (Attrib.attribute thy) atts)] haftmann@24304: |-> (fn [def] => add_inst' def (args, rhs) #> pair def) haftmann@24304: end; haftmann@24304: haftmann@24304: haftmann@24218: (* instances with implicit parameter handling *) haftmann@24218: haftmann@24218: local haftmann@24218: haftmann@24218: fun gen_read_def thy prep_att read_def ((raw_name, raw_atts), raw_t) = haftmann@24218: let haftmann@24218: val (_, t) = read_def thy (raw_name, raw_t); haftmann@24218: val ((c, ty), _) = Sign.cert_def (Sign.pp thy) t; haftmann@24218: val atts = map (prep_att thy) raw_atts; haftmann@24218: val insts = Consts.typargs (Sign.consts_of thy) (c, ty); haftmann@24218: val name = case raw_name haftmann@24218: of "" => NONE haftmann@24218: | _ => SOME raw_name; haftmann@24218: in (c, (insts, ((name, t), atts))) end; haftmann@24218: haftmann@24218: fun read_def_cmd thy = gen_read_def thy Attrib.intern_src Theory.read_axm; haftmann@24218: fun read_def thy = gen_read_def thy (K I) (K I); haftmann@24218: haftmann@24218: fun gen_instance_arity prep_arity read_def do_proof raw_arities raw_defs theory = haftmann@24218: let haftmann@24218: val arities = map (prep_arity theory) raw_arities; haftmann@24218: val _ = if null arities then error "at least one arity must be given" else (); haftmann@24218: val _ = case (duplicates (op =) o map #1) arities haftmann@24218: of [] => () haftmann@24218: | dupl_tycos => error ("type constructors occur more than once in arities: " haftmann@24218: ^ (commas o map quote) dupl_tycos); haftmann@24218: val super_sort = (Graph.all_succs o #classes o Sorts.rep_algebra o Sign.classes_of) theory haftmann@24218: fun get_consts_class tyco ty class = haftmann@24218: let haftmann@24218: val cs = (these o Option.map snd o try (AxClass.params_of_class theory)) class; haftmann@24218: val subst_ty = map_type_tfree (K ty); haftmann@24218: in haftmann@24304: map (fn (c, ty) => (c, ((class, tyco), subst_ty ty))) cs haftmann@24218: end; haftmann@24218: fun get_consts_sort (tyco, asorts, sort) = haftmann@24218: let haftmann@24218: val ty = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) (Name.names Name.context "'a" asorts)) haftmann@24218: in maps (get_consts_class tyco ty) (super_sort sort) end; haftmann@24218: val cs = maps get_consts_sort arities; haftmann@24218: fun mk_typnorm thy (ty, ty_sc) = haftmann@24218: case try (Sign.typ_match thy (Logic.varifyT ty_sc, ty)) Vartab.empty haftmann@24218: of SOME env => SOME (Logic.varifyT #> Envir.typ_subst_TVars env #> Logic.unvarifyT) haftmann@24218: | NONE => NONE; haftmann@24218: fun read_defs defs cs thy_read = haftmann@24218: let haftmann@24218: fun read raw_def cs = haftmann@24218: let haftmann@24218: val (c, (inst, ((name_opt, t), atts))) = read_def thy_read raw_def; haftmann@24218: val ty = Consts.instance (Sign.consts_of thy_read) (c, inst); haftmann@24304: val ((class, tyco), ty') = case AList.lookup (op =) cs c haftmann@24218: of NONE => error ("illegal definition for constant " ^ quote c) haftmann@24218: | SOME class_ty => class_ty; haftmann@24218: val name = case name_opt haftmann@24218: of NONE => Thm.def_name (Logic.name_arity (tyco, [], c)) haftmann@24218: | SOME name => name; haftmann@24218: val t' = case mk_typnorm thy_read (ty', ty) haftmann@24218: of NONE => error ("illegal definition for constant " ^ haftmann@24218: quote (c ^ "::" ^ setmp show_sorts true haftmann@24218: (Sign.string_of_typ thy_read) ty)) haftmann@24218: | SOME norm => map_types norm t haftmann@24218: in (((class, tyco), ((name, t'), atts)), AList.delete (op =) c cs) end; haftmann@24218: in fold_map read defs cs end; haftmann@24304: val (defs, other_cs) = read_defs raw_defs cs haftmann@24218: (fold Sign.primitive_arity arities (Theory.copy theory)); haftmann@24218: fun get_remove_contraint c thy = haftmann@24218: let haftmann@24218: val ty = Sign.the_const_constraint thy c; haftmann@24218: in haftmann@24218: thy haftmann@24218: |> Sign.add_const_constraint_i (c, NONE) haftmann@24218: |> pair (c, Logic.unvarifyT ty) haftmann@24218: end; haftmann@24304: fun after_qed cs defs = haftmann@24304: fold Sign.add_const_constraint_i (map (apsnd SOME) cs) haftmann@24304: #> fold (Code.add_func false) defs; haftmann@24218: in haftmann@24218: theory haftmann@24218: |> fold_map get_remove_contraint (map fst cs |> distinct (op =)) haftmann@24304: ||>> fold_map add_def defs haftmann@24304: ||> fold (fn (c, ((class, tyco), ty)) => add_inst_def' (class, tyco) (c, ty)) other_cs haftmann@24218: |-> (fn (cs, defs) => do_proof (after_qed cs defs) arities) haftmann@24218: end; haftmann@24218: haftmann@24218: fun instance_arity_cmd' do_proof = gen_instance_arity Sign.read_arity read_def_cmd do_proof; haftmann@24218: fun instance_arity' do_proof = gen_instance_arity Sign.cert_arity read_def do_proof; haftmann@24218: fun tactic_proof tac after_qed arities = haftmann@24218: fold (fn arity => AxClass.prove_arity arity tac) arities haftmann@24218: #> after_qed; haftmann@24218: haftmann@24218: in haftmann@24218: haftmann@24218: val instance_arity_cmd = instance_arity_cmd' axclass_instance_arity; haftmann@24218: val instance_arity = instance_arity' axclass_instance_arity; haftmann@24218: val prove_instance_arity = instance_arity' o tactic_proof; haftmann@24218: haftmann@24218: end; (*local*) haftmann@24218: haftmann@24218: haftmann@24218: haftmann@24218: (** combining locales and axclasses **) haftmann@24218: haftmann@24218: (* theory data *) haftmann@24218: haftmann@24218: datatype class_data = ClassData of { haftmann@24218: locale: string, haftmann@24218: consts: (string * string) list haftmann@24218: (*locale parameter ~> toplevel theory constant*), haftmann@24218: v: string option, haftmann@24218: intro: thm haftmann@24218: } * thm list (*derived defs*); haftmann@24218: haftmann@24218: fun rep_classdata (ClassData c) = c; haftmann@24218: haftmann@24218: fun merge_pair f1 f2 ((x1, y1), (x2, y2)) = (f1 (x1, x2), f2 (y1, y2)); haftmann@24218: haftmann@24218: structure ClassData = TheoryDataFun haftmann@24218: ( haftmann@24218: type T = class_data Graph.T * class Symtab.table (*locale name ~> class name*); haftmann@24218: val empty = (Graph.empty, Symtab.empty); haftmann@24218: val copy = I; haftmann@24218: val extend = I; haftmann@24218: fun merge _ = merge_pair (Graph.merge (K true)) (Symtab.merge (K true)); haftmann@24218: ); haftmann@24218: haftmann@24218: haftmann@24218: (* queries *) haftmann@24218: haftmann@24218: val lookup_class_data = Option.map rep_classdata oo try o Graph.get_node o fst o ClassData.get; haftmann@24218: fun class_of_locale thy = Symtab.lookup ((snd o ClassData.get) thy); haftmann@24218: haftmann@24218: fun the_class_data thy class = haftmann@24218: case lookup_class_data thy class haftmann@24218: of NONE => error ("undeclared class " ^ quote class) haftmann@24218: | SOME data => data; haftmann@24218: haftmann@24218: val ancestry = Graph.all_succs o fst o ClassData.get; haftmann@24218: haftmann@24218: fun param_map thy = haftmann@24218: let haftmann@24218: fun params class = haftmann@24218: let haftmann@24218: val const_typs = (#params o AxClass.get_definition thy) class; haftmann@24218: val const_names = (#consts o fst o the_class_data thy) class; haftmann@24218: in haftmann@24218: (map o apsnd) (fn c => (c, (the o AList.lookup (op =) const_typs) c)) const_names haftmann@24218: end; haftmann@24218: in maps params o ancestry thy end; haftmann@24218: haftmann@24218: fun these_defs thy = maps (these o Option.map snd o lookup_class_data thy) o ancestry thy; haftmann@24218: haftmann@24218: fun these_intros thy = haftmann@24218: Graph.fold (fn (_, (data, _)) => insert Thm.eq_thm ((#intro o fst o rep_classdata) data)) haftmann@24218: ((fst o ClassData.get) thy) []; haftmann@24218: haftmann@24218: fun print_classes thy = haftmann@24218: let haftmann@24218: val algebra = Sign.classes_of thy; haftmann@24218: val arities = haftmann@24218: Symtab.empty haftmann@24218: |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) => haftmann@24218: Symtab.map_default (class, []) (insert (op =) tyco)) arities) haftmann@24218: ((#arities o Sorts.rep_algebra) algebra); haftmann@24218: val the_arities = these o Symtab.lookup arities; haftmann@24218: fun mk_arity class tyco = haftmann@24218: let haftmann@24218: val Ss = Sorts.mg_domain algebra tyco [class]; haftmann@24218: in Sign.pretty_arity thy (tyco, Ss, [class]) end; haftmann@24218: fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: " haftmann@24218: ^ setmp show_sorts false (Sign.string_of_typ thy o Type.strip_sorts) ty); haftmann@24218: fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [ haftmann@24218: (SOME o Pretty.str) ("class " ^ class ^ ":"), haftmann@24218: (SOME o Pretty.block) [Pretty.str "supersort: ", haftmann@24218: (Sign.pretty_sort thy o Sign.certify_sort thy o Sign.super_classes thy) class], haftmann@24218: Option.map (Pretty.str o prefix "locale: " o #locale o fst) (lookup_class_data thy class), haftmann@24218: ((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks) (Pretty.str "parameters:" :: ps)) o map mk_param haftmann@24218: o these o Option.map #params o try (AxClass.get_definition thy)) class, haftmann@24218: (SOME o Pretty.block o Pretty.breaks) [ haftmann@24218: Pretty.str "instances:", haftmann@24218: Pretty.list "" "" (map (mk_arity class) (the_arities class)) haftmann@24218: ] haftmann@24218: ] haftmann@24218: in haftmann@24218: (Pretty.writeln o Pretty.chunks o separate (Pretty.str "") o map mk_entry o Sorts.all_classes) haftmann@24218: algebra haftmann@24218: end; haftmann@24218: haftmann@24218: haftmann@24218: (* updaters *) haftmann@24218: haftmann@24218: fun add_class_data ((class, superclasses), (locale, consts, v, intro)) = haftmann@24218: ClassData.map (fn (gr, tab) => ( haftmann@24218: gr haftmann@24218: |> Graph.new_node (class, ClassData ({ locale = locale, consts = consts, haftmann@24218: v = v, intro = intro }, [])) haftmann@24218: |> fold (curry Graph.add_edge class) superclasses, haftmann@24218: tab haftmann@24218: |> Symtab.update (locale, class) haftmann@24218: )); haftmann@24218: haftmann@24218: fun add_class_const_thm (class, thm) = (ClassData.map o apfst o Graph.map_node class) haftmann@24218: (fn ClassData (data, thms) => ClassData (data, thm :: thms)); haftmann@24218: haftmann@24218: (* tactics and methods *) haftmann@24218: haftmann@24218: fun intro_classes_tac facts st = haftmann@24218: let haftmann@24218: val thy = Thm.theory_of_thm st; haftmann@24218: val classes = Sign.all_classes thy; haftmann@24218: val class_trivs = map (Thm.class_triv thy) classes; haftmann@24218: val class_intros = these_intros thy; haftmann@24218: fun add_axclass_intro class = haftmann@24218: case try (AxClass.get_definition thy) class of SOME {intro, ...} => cons intro | _ => I; haftmann@24218: val axclass_intros = fold add_axclass_intro classes []; haftmann@24218: in haftmann@24218: st haftmann@24218: |> ((ALLGOALS (Method.insert_tac facts THEN' haftmann@24218: REPEAT_ALL_NEW (resolve_tac (class_trivs @ class_intros @ axclass_intros)))) haftmann@24218: THEN Tactic.distinct_subgoals_tac) haftmann@24218: end; haftmann@24218: haftmann@24218: fun default_intro_classes_tac [] = intro_classes_tac [] haftmann@24218: | default_intro_classes_tac _ = Tactical.no_tac; (*no error message!*) haftmann@24218: haftmann@24218: fun default_tac rules ctxt facts = haftmann@24218: HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE haftmann@24218: default_intro_classes_tac facts; haftmann@24218: haftmann@24218: val _ = Context.add_setup (Method.add_methods haftmann@24218: [("intro_classes", Method.no_args (Method.METHOD intro_classes_tac), haftmann@24218: "back-chain introduction rules of classes"), haftmann@24218: ("default", Method.thms_ctxt_args (Method.METHOD oo default_tac), haftmann@24218: "apply some intro/elim rule")]); haftmann@24218: haftmann@24218: haftmann@24218: (* tactical interfaces to locale commands *) haftmann@24218: haftmann@24218: fun prove_interpretation tac prfx_atts expr insts thy = haftmann@24218: thy haftmann@24218: |> Locale.interpretation_i I prfx_atts expr insts haftmann@24218: |> Proof.global_terminal_proof (Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE) haftmann@24218: |> ProofContext.theory_of; haftmann@24218: haftmann@24218: fun prove_interpretation_in tac after_qed (name, expr) thy = haftmann@24218: thy haftmann@24218: |> Locale.interpretation_in_locale (ProofContext.theory after_qed) (name, expr) haftmann@24218: |> Proof.global_terminal_proof (Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE) haftmann@24218: |> ProofContext.theory_of; haftmann@24218: haftmann@24218: haftmann@24218: (* constructing class introduction and other rules from axclass and locale rules *) haftmann@24218: haftmann@24218: fun mk_instT class = Symtab.empty haftmann@24218: |> Symtab.update (AxClass.param_tyvarname, TFree (AxClass.param_tyvarname, [class])); haftmann@24218: haftmann@24218: fun mk_inst class param_names cs = haftmann@24218: Symtab.empty haftmann@24218: |> fold2 (fn v => fn (c, ty) => Symtab.update (v, Const haftmann@24218: (c, Term.map_type_tfree (fn (v, _) => TFree (v, [class])) ty))) param_names cs; haftmann@24218: haftmann@24218: fun OF_LAST thm1 thm2 = haftmann@24218: let haftmann@24218: val n = (length o Logic.strip_imp_prems o prop_of) thm2; haftmann@24218: in (thm1 RSN (n, thm2)) end; haftmann@24218: haftmann@24218: fun strip_all_ofclass thy sort = haftmann@24218: let haftmann@24218: val typ = TVar ((AxClass.param_tyvarname, 0), sort); haftmann@24218: fun prem_inclass t = haftmann@24218: case Logic.strip_imp_prems t haftmann@24218: of ofcls :: _ => try Logic.dest_inclass ofcls haftmann@24218: | [] => NONE; haftmann@24218: fun strip_ofclass class thm = haftmann@24218: thm OF (fst o AxClass.of_sort thy (typ, [class])) AxClass.cache; haftmann@24218: fun strip thm = case (prem_inclass o Thm.prop_of) thm haftmann@24218: of SOME (_, class) => thm |> strip_ofclass class |> strip haftmann@24218: | NONE => thm; haftmann@24218: in strip end; haftmann@24218: haftmann@24218: fun class_intro thy locale class sups = haftmann@24218: let haftmann@24218: fun class_elim class = haftmann@24218: case (map Drule.unconstrainTs o #axioms o AxClass.get_definition thy) class haftmann@24218: of [thm] => SOME thm haftmann@24218: | [] => NONE; haftmann@24218: val pred_intro = case Locale.intros thy locale haftmann@24218: of ([ax_intro], [intro]) => intro |> OF_LAST ax_intro |> SOME haftmann@24218: | ([intro], []) => SOME intro haftmann@24218: | ([], [intro]) => SOME intro haftmann@24218: | _ => NONE; haftmann@24218: val pred_intro' = pred_intro haftmann@24218: |> Option.map (fn intro => intro OF map_filter class_elim sups); haftmann@24218: val class_intro = (#intro o AxClass.get_definition thy) class; haftmann@24218: val raw_intro = case pred_intro' haftmann@24218: of SOME pred_intro => class_intro |> OF_LAST pred_intro haftmann@24218: | NONE => class_intro; haftmann@24218: val sort = Sign.super_classes thy class; haftmann@24218: val typ = TVar ((AxClass.param_tyvarname, 0), sort); haftmann@24218: val defs = these_defs thy sups; haftmann@24218: in haftmann@24218: raw_intro haftmann@24218: |> Drule.instantiate' [SOME (Thm.ctyp_of thy typ)] [] haftmann@24218: |> strip_all_ofclass thy sort haftmann@24218: |> Thm.strip_shyps haftmann@24218: |> MetaSimplifier.rewrite_rule defs haftmann@24218: |> Drule.unconstrainTs haftmann@24218: end; haftmann@24218: haftmann@24218: fun interpretation_in_rule thy (class1, class2) = haftmann@24218: let haftmann@24218: val (params, consts) = split_list (param_map thy [class1]); haftmann@24218: (*FIXME also remember this at add_class*) haftmann@24218: fun mk_axioms class = haftmann@24218: let haftmann@24218: val name_locale = (#locale o fst o the_class_data thy) class; haftmann@24218: val inst = mk_inst class params consts; haftmann@24218: in haftmann@24218: Locale.global_asms_of thy name_locale haftmann@24218: |> maps snd haftmann@24218: |> (map o map_aterms) (fn Free (s, _) => (the o Symtab.lookup inst) s | t => t) haftmann@24218: |> (map o map_types o map_atyps) (fn TFree _ => TFree ("'a", [class1]) | T => T) haftmann@24218: |> map (ObjectLogic.ensure_propT thy) haftmann@24218: end; haftmann@24218: val (prems, concls) = pairself mk_axioms (class1, class2); haftmann@24218: in haftmann@24218: Goal.prove_global thy [] prems (Logic.mk_conjunction_list concls) haftmann@24218: (Locale.intro_locales_tac true (ProofContext.init thy)) haftmann@24218: end; haftmann@24218: haftmann@24218: haftmann@24218: (* classes *) haftmann@24218: haftmann@24218: local haftmann@24218: haftmann@24218: fun read_param thy raw_t = haftmann@24218: let haftmann@24218: val t = Sign.read_term thy raw_t haftmann@24218: in case try dest_Const t haftmann@24218: of SOME (c, _) => c haftmann@24218: | NONE => error ("Not a constant: " ^ Sign.string_of_term thy t) haftmann@24218: end; haftmann@24218: haftmann@24218: fun gen_class add_locale prep_class prep_param bname haftmann@24218: raw_supclasses raw_elems raw_other_consts thy = haftmann@24218: let haftmann@24218: (*FIXME need proper concept for reading locale statements*) haftmann@24218: fun subst_classtyvar (_, _) = haftmann@24218: TFree (AxClass.param_tyvarname, []) haftmann@24218: | subst_classtyvar (v, sort) = haftmann@24218: error ("Sort constraint illegal in type class, for type variable " ^ v ^ "::" ^ Sign.string_of_sort thy sort); haftmann@24218: (*val subst_classtyvars = Element.map_ctxt {name = I, var = I, term = I, haftmann@24218: typ = Term.map_type_tfree subst_classtyvar, fact = I, attrib = I};*) haftmann@24218: val other_consts = map (prep_param thy) raw_other_consts; haftmann@24218: val (elems, includes) = fold_rev (fn Locale.Elem e => apfst (cons e) haftmann@24218: | Locale.Expr i => apsnd (cons i)) raw_elems ([], []); haftmann@24218: val supclasses = map (prep_class thy) raw_supclasses; haftmann@24218: val sups = filter (is_some o lookup_class_data thy) supclasses haftmann@24218: |> Sign.certify_sort thy; haftmann@24218: val supsort = Sign.certify_sort thy supclasses; haftmann@24218: val suplocales = map (Locale.Locale o #locale o fst o the_class_data thy) sups; haftmann@24218: val supexpr = Locale.Merge (suplocales @ includes); haftmann@24218: val supparams = (map fst o Locale.parameters_of_expr thy) haftmann@24218: (Locale.Merge suplocales); haftmann@24218: val supconsts = AList.make (the o AList.lookup (op =) (param_map thy sups)) haftmann@24218: (map fst supparams); haftmann@24218: (*val elems_constrains = map haftmann@24218: (Element.Constrains o apsnd (Term.map_type_tfree subst_classtyvar)) supparams;*) haftmann@24218: fun mk_tyvar (_, sort) = TFree (AxClass.param_tyvarname, haftmann@24218: if Sign.subsort thy (supsort, sort) then sort else error haftmann@24218: ("Sort " ^ Sign.string_of_sort thy sort haftmann@24218: ^ " is less general than permitted least general sort " haftmann@24218: ^ Sign.string_of_sort thy supsort)); haftmann@24218: fun extract_params thy name_locale = haftmann@24218: let haftmann@24218: val params = Locale.parameters_of thy name_locale; haftmann@24218: val v = case (maps typ_tfrees o map (snd o fst)) params haftmann@24218: of (v, _) :: _ => SOME v haftmann@24218: | _ => NONE; haftmann@24218: in haftmann@24218: (v, (map (fst o fst) params, params haftmann@24218: |> (map o apfst o apsnd o Term.map_type_tfree) mk_tyvar haftmann@24218: |> (map o apsnd) (fork_mixfix true NONE #> fst) haftmann@24218: |> chop (length supconsts) haftmann@24218: |> snd)) haftmann@24218: end; haftmann@24218: fun extract_assumes name_locale params thy cs = haftmann@24218: let haftmann@24218: val consts = supconsts @ (map (fst o fst) params ~~ cs); haftmann@24218: fun subst (Free (c, ty)) = haftmann@24218: Const ((fst o the o AList.lookup (op =) consts) c, ty) haftmann@24218: | subst t = t; haftmann@24218: val super_defs = these_defs thy sups; haftmann@24218: fun prep_asm ((name, atts), ts) = haftmann@24218: ((NameSpace.base name, map (Attrib.attribute thy) atts), haftmann@24218: (map o map_aterms) ((*MetaSimplifier.rewrite_term thy super_defs [] o *)subst) ts); haftmann@24218: in haftmann@24218: Locale.global_asms_of thy name_locale haftmann@24218: |> map prep_asm haftmann@24218: end; haftmann@24218: fun note_intro name_axclass class_intro = haftmann@24218: PureThy.note_thmss_qualified "" ((Logic.const_of_class o NameSpace.base) name_axclass) haftmann@24218: [(("intro", []), [([class_intro], [])])] haftmann@24218: #> snd haftmann@24218: in haftmann@24218: thy haftmann@24218: |> add_locale (SOME "") bname supexpr ((*elems_constrains @*) elems) haftmann@24218: |-> (fn name_locale => ProofContext.theory_result ( haftmann@24218: `(fn thy => extract_params thy name_locale) haftmann@24218: #-> (fn (v, (param_names, params)) => haftmann@24218: axclass_params (bname, supsort) params (extract_assumes name_locale params) other_consts haftmann@24218: #-> (fn (name_axclass, ((_, (ax_terms, ax_axioms)), consts)) => haftmann@24218: `(fn thy => class_intro thy name_locale name_axclass sups) haftmann@24218: #-> (fn class_intro => haftmann@24218: add_class_data ((name_axclass, sups), haftmann@24218: (name_locale, map (fst o fst) params ~~ map fst consts, v, haftmann@24218: class_intro)) haftmann@24218: (*FIXME: class_data should already contain data relevant haftmann@24218: for interpretation; use this later for class target*) haftmann@24218: (*FIXME: general export_fixes which may be parametrized haftmann@24218: with pieces of an emerging class*) haftmann@24218: #> note_intro name_axclass class_intro haftmann@24218: #> prove_interpretation ((ALLGOALS o ProofContext.fact_tac) ax_axioms) haftmann@24218: ((false, Logic.const_of_class bname), []) (Locale.Locale name_locale) haftmann@24218: ((mk_instT name_axclass, mk_inst name_axclass param_names (map snd supconsts @ consts)), []) haftmann@24218: #> pair name_axclass haftmann@24218: ))))) haftmann@24218: end; haftmann@24218: haftmann@24218: in haftmann@24218: haftmann@24218: val class_cmd = gen_class Locale.add_locale Sign.intern_class read_param; haftmann@24218: val class = gen_class Locale.add_locale_i Sign.certify_class (K I); haftmann@24218: haftmann@24218: end; (*local*) haftmann@24218: haftmann@24218: local haftmann@24218: haftmann@24218: fun instance_subclass (class1, class2) thy = haftmann@24218: let haftmann@24218: val interp = interpretation_in_rule thy (class1, class2); haftmann@24218: val ax = #axioms (AxClass.get_definition thy class1); haftmann@24218: val intro = #intro (AxClass.get_definition thy class2) haftmann@24218: |> Drule.instantiate' [SOME (Thm.ctyp_of thy haftmann@24218: (TVar ((AxClass.param_tyvarname, 0), [class1])))] []; haftmann@24218: val thm = haftmann@24218: intro haftmann@24218: |> OF_LAST (interp OF ax) haftmann@24218: |> strip_all_ofclass thy (Sign.super_classes thy class2); haftmann@24218: in haftmann@24218: thy |> AxClass.add_classrel thm haftmann@24218: end; haftmann@24218: haftmann@24218: fun instance_subsort (class, sort) thy = haftmann@24218: let haftmann@24218: val super_sort = (Graph.all_succs o #classes o Sorts.rep_algebra haftmann@24218: o Sign.classes_of) thy sort; haftmann@24218: val classes = filter_out (fn class' => Sign.subsort thy ([class], [class'])) haftmann@24218: (rev super_sort); haftmann@24218: in haftmann@24218: thy |> fold (curry instance_subclass class) classes haftmann@24218: end; haftmann@24218: haftmann@24218: fun instance_sort' do_proof (class, sort) theory = haftmann@24218: let haftmann@24218: val loc_name = (#locale o fst o the_class_data theory) class; haftmann@24218: val loc_expr = haftmann@24218: (Locale.Merge o map (Locale.Locale o #locale o fst o the_class_data theory)) sort; haftmann@24218: in haftmann@24218: theory haftmann@24218: |> do_proof (instance_subsort (class, sort)) (loc_name, loc_expr) haftmann@24218: end; haftmann@24218: haftmann@24218: fun gen_instance_sort prep_class prep_sort (raw_class, raw_sort) theory = haftmann@24218: let haftmann@24218: val class = prep_class theory raw_class; haftmann@24218: val sort = prep_sort theory raw_sort; haftmann@24218: in haftmann@24218: theory haftmann@24218: |> instance_sort' (Locale.interpretation_in_locale o ProofContext.theory) (class, sort) haftmann@24218: end; haftmann@24218: haftmann@24218: fun gen_instance_class prep_class (raw_class, raw_superclass) theory = haftmann@24218: let haftmann@24218: val class = prep_class theory raw_class; haftmann@24218: val superclass = prep_class theory raw_superclass; haftmann@24218: in haftmann@24218: theory haftmann@24218: |> axclass_instance_sort (class, superclass) haftmann@24218: end; haftmann@24218: haftmann@24218: in haftmann@24218: wenzelm@24276: val instance_sort_cmd = gen_instance_sort Sign.read_class Syntax.global_read_sort; haftmann@24218: val instance_sort = gen_instance_sort Sign.certify_class Sign.certify_sort; haftmann@24218: val prove_instance_sort = instance_sort' o prove_interpretation_in; haftmann@24218: val instance_class_cmd = gen_instance_class Sign.read_class; haftmann@24218: val instance_class = gen_instance_class Sign.certify_class; haftmann@24218: haftmann@24218: end; (*local*) haftmann@24218: haftmann@24218: haftmann@24218: (** class target **) haftmann@24218: haftmann@24218: fun export_fixes thy class = haftmann@24218: let haftmann@24218: val v = (#v o fst o the_class_data thy) class; haftmann@24218: val constrain_sort = curry (Sorts.inter_sort (Sign.classes_of thy)) [class]; haftmann@24218: val subst_typ = Term.map_type_tfree (fn var as (w, sort) => haftmann@24218: if SOME w = v then TFree (w, constrain_sort sort) else TFree var); haftmann@24218: val consts = param_map thy [class]; haftmann@24218: fun subst_aterm (t as Free (v, ty)) = (case AList.lookup (op =) consts v haftmann@24218: of SOME (c, _) => Const (c, ty) haftmann@24218: | NONE => t) haftmann@24218: | subst_aterm t = t; haftmann@24218: in map_types subst_typ #> Term.map_aterms subst_aterm end; haftmann@24218: haftmann@24218: fun add_const_in_class class ((c, rhs), syn) thy = haftmann@24218: let haftmann@24218: val prfx = (Logic.const_of_class o NameSpace.base) class; haftmann@24218: fun mk_name inject c = haftmann@24218: let haftmann@24218: val n1 = Sign.full_name thy c; haftmann@24218: val n2 = NameSpace.qualifier n1; haftmann@24218: val n3 = NameSpace.base n1; haftmann@24218: in NameSpace.implode (n2 :: inject @ [n3]) end; haftmann@24218: val abbr' = mk_name [prfx, prfx] c; haftmann@24218: val rhs' = export_fixes thy class rhs; haftmann@24218: val ty' = Term.fastype_of rhs'; haftmann@24218: val def = (c, Logic.mk_equals (Const (mk_name [prfx] c, ty'), rhs')); haftmann@24218: val (syn', _) = fork_mixfix true NONE syn; haftmann@24218: fun interpret def = haftmann@24218: let haftmann@24218: val def' = symmetric def haftmann@24218: val tac = (ALLGOALS o ProofContext.fact_tac) [def']; haftmann@24218: val name_locale = (#locale o fst o the_class_data thy) class; haftmann@24218: val def_eq = Thm.prop_of def'; haftmann@24218: val (params, consts) = split_list (param_map thy [class]); haftmann@24218: in haftmann@24218: prove_interpretation tac ((false, prfx), []) (Locale.Locale name_locale) haftmann@24218: ((mk_instT class, mk_inst class params consts), [def_eq]) haftmann@24218: #> add_class_const_thm (class, def') haftmann@24218: end; haftmann@24218: in haftmann@24218: thy haftmann@24218: |> Sign.hide_consts_i true [abbr'] haftmann@24218: |> Sign.add_path prfx haftmann@24218: |> Sign.add_consts_authentic [(c, ty', syn')] haftmann@24218: |> Sign.parent_path haftmann@24218: |> Sign.sticky_prefix prfx haftmann@24218: |> PureThy.add_defs_i false [(def, [])] haftmann@24218: |-> (fn [def] => interpret def) haftmann@24218: |> Sign.restore_naming thy haftmann@24218: end; haftmann@24218: haftmann@24218: end;