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@25062: val fork_mixfix: bool -> bool -> mixfix -> (mixfix * mixfix) * mixfix haftmann@24218: haftmann@24218: val axclass_cmd: bstring * xstring list haftmann@24589: -> ((bstring * Attrib.src list) * string list) list haftmann@24589: -> theory -> class * theory haftmann@24589: val classrel_cmd: xstring * xstring -> theory -> Proof.state haftmann@24914: haftmann@25002: val class: bstring -> class list -> Element.context_i Locale.element list haftmann@24218: -> string list -> theory -> string * Proof.context haftmann@25002: val class_cmd: bstring -> xstring list -> Element.context Locale.element list haftmann@24589: -> xstring list -> theory -> string * Proof.context wenzelm@25060: val init: class -> Proof.context -> Proof.context haftmann@25062: val add_const_in_class: string -> (string * mixfix) * term haftmann@25038: -> theory -> string * theory haftmann@25062: val add_abbrev_in_class: string -> Syntax.mode -> (string * mixfix) * term haftmann@25038: -> theory -> string * theory haftmann@25002: val remove_constraint: class -> string -> Proof.context -> Proof.context haftmann@25038: val is_class: theory -> class -> bool haftmann@25002: val these_params: theory -> sort -> (string * (string * typ)) list haftmann@24589: val intro_classes_tac: thm list -> tactic haftmann@24589: val default_intro_classes_tac: thm list -> tactic haftmann@24589: val print_classes: theory -> unit haftmann@25002: val uncheck: bool ref haftmann@24423: haftmann@24589: val instance_arity: (theory -> theory) -> arity list -> theory -> Proof.state haftmann@24589: val instance: arity list -> ((bstring * Attrib.src list) * term) list haftmann@24423: -> (thm list -> theory -> theory) haftmann@24218: -> theory -> Proof.state haftmann@24589: val instance_cmd: (bstring * xstring list * xstring) list haftmann@24589: -> ((bstring * Attrib.src list) * xstring) list haftmann@24423: -> (thm list -> theory -> theory) haftmann@24218: -> theory -> Proof.state haftmann@24589: val prove_instance: tactic -> arity list haftmann@24218: -> ((bstring * Attrib.src list) * term) list haftmann@24423: -> theory -> thm list * theory haftmann@24423: val unoverload: theory -> conv haftmann@24423: val overload: theory -> conv haftmann@24423: val unoverload_const: theory -> string * typ -> string haftmann@24423: val inst_const: theory -> string * string -> string haftmann@24423: val param_const: theory -> string -> (string * string) option haftmann@24218: end; haftmann@24218: haftmann@24218: structure Class : CLASS = haftmann@24218: struct haftmann@24218: haftmann@24218: (** auxiliary **) haftmann@24218: haftmann@25062: val classN = "class"; haftmann@25062: val introN = "intro"; haftmann@25062: haftmann@25062: fun fork_mixfix false _ mx = ((NoSyn, NoSyn), mx) haftmann@25062: | fork_mixfix true false mx = ((NoSyn, mx), NoSyn) haftmann@25062: | fork_mixfix true true mx = ((mx, NoSyn), NoSyn); haftmann@25062: (*let haftmann@25062: val mx' = Syntax.unlocalize_mixfix mx; haftmann@25062: in if mx' = mx then ((NoSyn, mx), NoSyn) else ((mx', mx), NoSyn) end;*) haftmann@24218: haftmann@25002: fun prove_interpretation tac prfx_atts expr inst = haftmann@25002: Locale.interpretation_i I prfx_atts expr inst haftmann@24589: #> Proof.global_terminal_proof haftmann@24589: (Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE) haftmann@24589: #> ProofContext.theory_of; haftmann@24589: wenzelm@25020: fun OF_LAST thm1 thm2 = thm1 RSN (Thm.nprems_of thm2, thm2); haftmann@24589: haftmann@24589: fun strip_all_ofclass thy sort = haftmann@24589: let wenzelm@24847: val typ = TVar ((Name.aT, 0), sort); haftmann@24589: fun prem_inclass t = haftmann@24589: case Logic.strip_imp_prems t haftmann@24589: of ofcls :: _ => try Logic.dest_inclass ofcls haftmann@24589: | [] => NONE; haftmann@24589: fun strip_ofclass class thm = haftmann@24589: thm OF (fst o AxClass.of_sort thy (typ, [class])) AxClass.cache; haftmann@24589: fun strip thm = case (prem_inclass o Thm.prop_of) thm haftmann@24589: of SOME (_, class) => thm |> strip_ofclass class |> strip haftmann@24589: | NONE => thm; haftmann@24589: in strip end; haftmann@24589: haftmann@25038: fun get_remove_global_constraint c thy = haftmann@25038: let haftmann@25038: val ty = Sign.the_const_constraint thy c; haftmann@25038: in haftmann@25038: thy haftmann@25038: |> Sign.add_const_constraint (c, NONE) haftmann@25038: |> pair (c, Logic.unvarifyT ty) haftmann@25038: end; haftmann@25038: haftmann@24589: haftmann@24589: (** axclass command **) haftmann@24589: 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@24589: val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst) haftmann@24589: raw_specs; haftmann@24589: val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd) haftmann@24589: raw_specs) haftmann@24218: |> snd haftmann@24218: |> (map o map) fst; haftmann@24589: in haftmann@24589: AxClass.define_class (class, superclasses) [] haftmann@24589: (name_atts ~~ axiomss) thy haftmann@24589: end; 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@24589: |> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])]) haftmann@24589: o maps (mk_prop thy)) insts) haftmann@24218: end; haftmann@24218: haftmann@24218: in haftmann@24218: haftmann@24589: val instance_arity = haftmann@24218: gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity; haftmann@24589: val classrel = haftmann@24218: gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel)) haftmann@24218: AxClass.add_classrel I o single; haftmann@24589: val classrel_cmd = haftmann@24589: gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) haftmann@24589: AxClass.add_classrel I o single; haftmann@24218: haftmann@24218: end; (*local*) haftmann@24218: haftmann@24218: haftmann@24589: (** explicit constants for overloaded definitions **) haftmann@24304: haftmann@24304: structure InstData = TheoryDataFun haftmann@24304: ( haftmann@24423: type T = (string * thm) Symtab.table Symtab.table * (string * string) Symtab.table; haftmann@24423: (*constant name ~> type constructor ~> (constant name, equation), haftmann@24423: constant name ~> (constant name, type constructor)*) haftmann@24423: val empty = (Symtab.empty, Symtab.empty); haftmann@24304: val copy = I; haftmann@24304: val extend = I; haftmann@24423: fun merge _ ((taba1, tabb1), (taba2, tabb2)) = haftmann@24423: (Symtab.join (K (Symtab.merge (K true))) (taba1, taba2), haftmann@24423: Symtab.merge (K true) (tabb1, tabb2)); haftmann@24304: ); haftmann@24304: wenzelm@25020: fun inst_thms thy = (Symtab.fold (Symtab.fold (cons o snd o snd) o snd) o fst) haftmann@24589: (InstData.get thy) []; haftmann@24589: fun add_inst (c, tyco) inst = (InstData.map o apfst haftmann@24589: o Symtab.map_default (c, Symtab.empty)) (Symtab.update_new (tyco, inst)) haftmann@24423: #> (InstData.map o apsnd) (Symtab.update_new (fst inst, (c, tyco))); haftmann@24304: wenzelm@25020: fun unoverload thy = MetaSimplifier.rewrite true (inst_thms thy); wenzelm@25020: fun overload thy = MetaSimplifier.rewrite true (map Thm.symmetric (inst_thms thy)); haftmann@24304: haftmann@24304: fun inst_const thy (c, tyco) = haftmann@24423: (fst o the o Symtab.lookup ((the o Symtab.lookup (fst (InstData.get thy))) c)) tyco; haftmann@24423: fun unoverload_const thy (c_ty as (c, _)) = haftmann@24423: case AxClass.class_of_param thy c haftmann@24423: of SOME class => (case Sign.const_typargs thy c_ty haftmann@24589: of [Type (tyco, _)] => (case Symtab.lookup haftmann@24589: ((the o Symtab.lookup (fst (InstData.get thy))) c) tyco haftmann@24423: of SOME (c, _) => c haftmann@24423: | NONE => c) haftmann@24423: | [_] => c) haftmann@24423: | NONE => c; haftmann@24423: haftmann@24423: val param_const = Symtab.lookup o snd o InstData.get; haftmann@24304: haftmann@24304: fun add_inst_def (class, tyco) (c, ty) thy = haftmann@24304: let wenzelm@25024: val tyco_base = Sign.base_name tyco; wenzelm@25024: val name_inst = Sign.base_name class ^ "_" ^ tyco_base ^ "_inst"; wenzelm@25024: val c_inst_base = Sign.base_name c ^ "_" ^ tyco_base; haftmann@24304: in haftmann@24304: thy haftmann@24304: |> Sign.sticky_prefix name_inst wenzelm@25020: |> Sign.declare_const [] (c_inst_base, ty, NoSyn) wenzelm@24968: |-> (fn const as Const (c_inst, _) => wenzelm@24968: PureThy.add_defs_i false wenzelm@24968: [((Thm.def_name c_inst_base, Logic.mk_equals (const, Const (c, ty))), [])] wenzelm@24968: #-> (fn [thm] => add_inst (c, tyco) (c_inst, Thm.symmetric thm))) haftmann@24435: |> Sign.restore_naming thy haftmann@24304: end; haftmann@24304: haftmann@24304: fun add_inst_def' (class, tyco) (c, ty) thy = haftmann@24423: if case Symtab.lookup (fst (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@24589: val ((lhs as Const (c, ty), args), rhs) = haftmann@24589: (apfst Term.strip_comb o Logic.dest_equals) prop; haftmann@24701: fun (*add_inst' def ([], (Const (c_inst, ty))) = haftmann@24423: 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@24701: |*) add_inst' _ t = add_inst_def (class, tyco) (c, ty); haftmann@24304: in haftmann@24304: thy wenzelm@25020: |> PureThy.add_defs_i true [((name, prop), map (Attrib.attribute_i thy) atts)] haftmann@24304: |-> (fn [def] => add_inst' def (args, rhs) #> pair def) haftmann@24304: end; haftmann@24304: haftmann@24304: haftmann@24589: (** instances with implicit parameter handling **) haftmann@24218: haftmann@24218: local haftmann@24218: wenzelm@24707: fun gen_read_def thy prep_att parse_prop ((raw_name, raw_atts), raw_t) = haftmann@24218: let wenzelm@24707: val ctxt = ProofContext.init thy; wenzelm@24707: val t = parse_prop ctxt raw_t |> Syntax.check_prop ctxt; wenzelm@24981: val ((c, ty), _) = Sign.cert_def ctxt 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: wenzelm@24707: fun read_def_cmd thy = gen_read_def thy Attrib.intern_src Syntax.parse_prop; haftmann@24218: fun read_def thy = gen_read_def thy (K I) (K I); haftmann@24218: haftmann@24589: fun gen_instance prep_arity read_def do_proof raw_arities raw_defs after_qed theory = haftmann@24218: let haftmann@24218: val arities = map (prep_arity theory) raw_arities; wenzelm@25020: 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 [] => () wenzelm@25020: | dupl_tycos => error ("Type constructors occur more than once in arities: " wenzelm@25020: ^ commas_quote dupl_tycos); haftmann@24218: fun get_consts_class tyco ty class = haftmann@24218: let wenzelm@24968: val cs = (these o try (#params o AxClass.get_info 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@24589: val ty = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) wenzelm@24847: (Name.names Name.context Name.aT asorts)) wenzelm@24731: in maps (get_consts_class tyco ty) (Sign.complete_sort theory 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 wenzelm@25020: 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) wenzelm@25020: 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@24423: fun after_qed' cs defs = wenzelm@24766: fold Sign.add_const_constraint (map (apsnd SOME) cs) haftmann@24423: #> after_qed defs; haftmann@24218: in haftmann@24218: theory haftmann@25038: |> fold_map get_remove_global_constraint (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@24423: |-> (fn (cs, defs) => do_proof (after_qed' cs defs) arities defs) haftmann@24218: end; haftmann@24218: haftmann@24423: fun tactic_proof tac f arities defs = haftmann@24218: fold (fn arity => AxClass.prove_arity arity tac) arities haftmann@24423: #> f haftmann@24423: #> pair defs; haftmann@24218: haftmann@24218: in haftmann@24218: haftmann@24589: val instance = gen_instance Sign.cert_arity read_def haftmann@24589: (fn f => fn arities => fn defs => instance_arity f arities); haftmann@24589: val instance_cmd = gen_instance Sign.read_arity read_def_cmd haftmann@24589: (fn f => fn arities => fn defs => instance_arity f arities); haftmann@24589: fun prove_instance tac arities defs = haftmann@24589: gen_instance Sign.cert_arity read_def haftmann@24589: (tactic_proof tac) arities defs (K I); haftmann@24218: haftmann@24218: end; (*local*) haftmann@24218: haftmann@24218: haftmann@24218: haftmann@24589: (** class data **) haftmann@24218: haftmann@24218: datatype class_data = ClassData of { haftmann@24218: consts: (string * string) list haftmann@24836: (*locale parameter ~> constant name*), haftmann@25062: base_sort: sort, haftmann@25038: inst: (typ option list * term option list) * term Symtab.table haftmann@25038: (*canonical interpretation FIXME*), haftmann@25062: morphism: morphism, haftmann@25062: (*partial morphism of canonical interpretation*) haftmann@24657: intro: thm, haftmann@24657: defs: thm list, haftmann@25062: operations: (string * ((term * typ) * (typ * int))) list haftmann@25062: (*constant name ~> (locale term & typ, haftmann@25062: (constant constraint, instantiaton index of class typ))*) haftmann@24657: }; haftmann@24218: haftmann@24657: fun rep_class_data (ClassData d) = d; haftmann@25062: fun mk_class_data ((consts, base_sort, inst, morphism, intro), haftmann@25062: (defs, operations)) = haftmann@25062: ClassData { consts = consts, base_sort = base_sort, inst = inst, haftmann@25062: morphism = morphism, intro = intro, defs = defs, haftmann@25062: operations = operations }; haftmann@25062: fun map_class_data f (ClassData { consts, base_sort, inst, morphism, intro, haftmann@25062: defs, operations }) = haftmann@25062: mk_class_data (f ((consts, base_sort, inst, morphism, intro), haftmann@25062: (defs, operations))); haftmann@25038: fun merge_class_data _ (ClassData { consts = consts, haftmann@25062: base_sort = base_sort, inst = inst, morphism = morphism, intro = intro, haftmann@25062: defs = defs1, operations = operations1 }, haftmann@25062: ClassData { consts = _, base_sort = _, inst = _, morphism = _, intro = _, haftmann@25062: defs = defs2, operations = operations2 }) = haftmann@25062: mk_class_data ((consts, base_sort, inst, morphism, intro), haftmann@24914: (Thm.merge_thms (defs1, defs2), haftmann@25062: AList.merge (op =) (K true) (operations1, operations2))); haftmann@24218: haftmann@24218: structure ClassData = TheoryDataFun haftmann@24218: ( haftmann@25038: type T = class_data Graph.T haftmann@25038: val empty = Graph.empty; haftmann@24218: val copy = I; haftmann@24218: val extend = I; haftmann@25038: fun merge _ = Graph.join merge_class_data; haftmann@24218: ); haftmann@24218: haftmann@24218: haftmann@24218: (* queries *) haftmann@24218: haftmann@25038: val lookup_class_data = Option.map rep_class_data oo try o Graph.get_node o ClassData.get; haftmann@24218: haftmann@24589: fun the_class_data thy class = case lookup_class_data thy class wenzelm@25020: of NONE => error ("Undeclared class " ^ quote class) haftmann@24589: | SOME data => data; haftmann@24218: haftmann@25038: val is_class = is_some oo lookup_class_data; haftmann@25038: haftmann@25038: val ancestry = Graph.all_succs o ClassData.get; haftmann@24218: haftmann@25002: fun these_params thy = haftmann@24218: let haftmann@24218: fun params class = haftmann@24218: let wenzelm@24930: val const_typs = (#params o AxClass.get_info thy) class; haftmann@24657: val const_names = (#consts 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@24657: fun these_defs thy = maps (these o Option.map #defs o lookup_class_data thy) o ancestry thy; haftmann@24218: haftmann@25062: fun morphism thy = #morphism o the_class_data thy; haftmann@25062: haftmann@24218: fun these_intros thy = haftmann@24657: Graph.fold (fn (_, (data, _)) => insert Thm.eq_thm ((#intro o rep_class_data) data)) haftmann@25038: (ClassData.get thy) []; haftmann@24218: haftmann@24836: fun these_operations thy = haftmann@24836: maps (#operations o the_class_data thy) o ancestry thy; haftmann@24657: haftmann@25002: fun local_operation thy = AList.lookup (op =) o these_operations thy; haftmann@25002: haftmann@25062: fun sups_base_sort thy sort = haftmann@24901: let haftmann@25062: val sups = filter (is_class thy) sort haftmann@24901: |> Sign.minimize_sort thy; haftmann@25062: val base_sort = case sups haftmann@25062: of _ :: _ => maps (#base_sort o the_class_data thy) sups haftmann@25062: |> Sign.minimize_sort thy haftmann@24901: | [] => sort; haftmann@25062: in (sups, base_sort) end; haftmann@24901: haftmann@24218: fun print_classes thy = haftmann@24218: let wenzelm@24920: val ctxt = ProofContext.init thy; 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]; wenzelm@24920: in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end; haftmann@24218: fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: " wenzelm@24920: ^ setmp show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty); haftmann@24218: fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [ haftmann@25062: (SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"), haftmann@24218: (SOME o Pretty.block) [Pretty.str "supersort: ", wenzelm@24920: (Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class], haftmann@25062: if is_class thy class then (SOME o Pretty.str) haftmann@25062: ("locale: " ^ Locale.extern thy class) else NONE, haftmann@25062: ((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks) haftmann@25062: (Pretty.str "parameters:" :: ps)) o map mk_param wenzelm@24930: o these o Option.map #params o try (AxClass.get_info 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@24589: (Pretty.writeln o Pretty.chunks o separate (Pretty.str "") haftmann@24589: o map mk_entry o Sorts.all_classes) algebra haftmann@24218: end; haftmann@24218: haftmann@24218: haftmann@24218: (* updaters *) haftmann@24218: haftmann@25062: fun add_class_data ((class, superclasses), (cs, base_sort, inst, phi, intro)) thy = haftmann@25002: let haftmann@25038: (*FIXME*) haftmann@25038: val is_empty = null (fold (fn ((_, ty), _) => fold_atyps cons ty) cs []) haftmann@25038: andalso null ((fold o fold_types o fold_atyps) cons haftmann@25038: (maps snd (Locale.global_asms_of thy class)) []); haftmann@25038: (*FIXME*) haftmann@25038: val inst_params = map haftmann@25038: (SOME o the o Symtab.lookup inst o fst o fst) haftmann@25038: (Locale.parameters_of_expr thy (Locale.Locale class)); haftmann@25038: val instT = if is_empty then [] else [SOME (TFree (Name.aT, [class]))]; haftmann@25038: val inst' = ((instT, inst_params), inst); haftmann@25062: val operations = map (fn (v_ty as (_, ty'), (c, ty)) => haftmann@25062: (c, ((Free v_ty, ty'), (Logic.varifyT ty, 0)))) cs; haftmann@25002: val cs = (map o pairself) fst cs; haftmann@25038: val add_class = Graph.new_node (class, haftmann@25062: mk_class_data ((cs, base_sort, inst', phi, intro), ([], operations))) haftmann@25002: #> fold (curry Graph.add_edge class) superclasses; haftmann@25002: in haftmann@25038: ClassData.map add_class thy haftmann@25002: end; haftmann@24218: haftmann@25062: fun register_operation class ((c, rhs), some_def) thy = haftmann@25062: let haftmann@25062: val ty = Sign.the_const_constraint thy c; haftmann@25062: val typargs = Sign.const_typargs thy (c, ty); haftmann@25062: val typidx = find_index (fn TVar ((v, _), _) => Name.aT = v | _ => false) typargs; haftmann@25062: val ty' = Term.fastype_of rhs; haftmann@25062: in haftmann@25062: thy haftmann@25062: |> (ClassData.map o Graph.map_node class o map_class_data o apsnd) haftmann@25062: (fn (defs, operations) => haftmann@25062: (case some_def of NONE => defs | SOME def => def :: defs, haftmann@25062: (c, ((rhs, ty'), (ty, typidx))) :: operations)) haftmann@25062: end; haftmann@24218: haftmann@24589: haftmann@24589: (** rule calculation, tactics and methods **) haftmann@24589: wenzelm@25024: val class_prefix = Logic.const_of_class o Sign.base_name; wenzelm@25024: haftmann@25062: fun calculate_morphism class cs = haftmann@25062: let haftmann@25062: val subst_typ = Term.map_type_tfree (fn var as (v, sort) => haftmann@25062: if v = Name.aT then TVar ((v, 0), [class]) else TVar ((v, 0), sort)); haftmann@25062: fun subst_aterm (t as Free (v, ty)) = (case AList.lookup (op =) cs v haftmann@25062: of SOME (c, _) => Const (c, ty) haftmann@25062: | NONE => t) haftmann@25062: | subst_aterm t = t; haftmann@25062: val subst_term = map_aterms subst_aterm #> map_types subst_typ; haftmann@25062: in haftmann@25062: Morphism.identity haftmann@25062: $> Morphism.term_morphism subst_term haftmann@25062: $> Morphism.typ_morphism subst_typ haftmann@25062: end; haftmann@25062: haftmann@25038: fun class_intro thy class sups = haftmann@24589: let haftmann@24589: fun class_elim class = wenzelm@25020: case (#axioms o AxClass.get_info thy) class wenzelm@25020: of [thm] => SOME (Drule.unconstrainTs thm) haftmann@24589: | [] => NONE; haftmann@25038: val pred_intro = case Locale.intros thy class haftmann@24589: of ([ax_intro], [intro]) => intro |> OF_LAST ax_intro |> SOME haftmann@24589: | ([intro], []) => SOME intro haftmann@24589: | ([], [intro]) => SOME intro haftmann@24589: | _ => NONE; haftmann@24589: val pred_intro' = pred_intro haftmann@24589: |> Option.map (fn intro => intro OF map_filter class_elim sups); wenzelm@24930: val class_intro = (#intro o AxClass.get_info thy) class; haftmann@24589: val raw_intro = case pred_intro' haftmann@24589: of SOME pred_intro => class_intro |> OF_LAST pred_intro haftmann@24589: | NONE => class_intro; haftmann@24589: val sort = Sign.super_classes thy class; wenzelm@24847: val typ = TVar ((Name.aT, 0), sort); haftmann@24589: val defs = these_defs thy sups; haftmann@24589: in haftmann@24589: raw_intro haftmann@24589: |> Drule.instantiate' [SOME (Thm.ctyp_of thy typ)] [] haftmann@24589: |> strip_all_ofclass thy sort haftmann@24589: |> Thm.strip_shyps haftmann@24589: |> MetaSimplifier.rewrite_rule defs haftmann@24589: |> Drule.unconstrainTs haftmann@24589: end; haftmann@24589: haftmann@24589: fun class_interpretation class facts defs thy = haftmann@24589: let haftmann@25038: val params = these_params thy [class]; haftmann@25038: val { inst = ((_, inst), _), ... } = the_class_data thy class; wenzelm@25020: val tac = ALLGOALS (ProofContext.fact_tac facts); haftmann@25038: val prfx = class_prefix class; haftmann@24589: in haftmann@25038: thy haftmann@25038: |> fold_map (get_remove_global_constraint o fst o snd) params haftmann@25038: ||> prove_interpretation tac ((false, prfx), []) (Locale.Locale class) (inst, defs) haftmann@25038: |-> (fn cs => fold (Sign.add_const_constraint o apsnd SOME) cs) haftmann@24589: end; 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; wenzelm@24930: val axclass_intros = map_filter (try (#intro o AxClass.get_info thy)) classes; haftmann@24218: in wenzelm@24930: (ALLGOALS (Method.insert_tac facts THEN' wenzelm@24930: REPEAT_ALL_NEW (resolve_tac (class_trivs @ class_intros @ axclass_intros))) wenzelm@24930: THEN Tactic.distinct_subgoals_tac) st haftmann@24218: end; haftmann@24218: haftmann@24218: fun default_intro_classes_tac [] = intro_classes_tac [] wenzelm@24930: | default_intro_classes_tac _ = no_tac; 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@24589: (** classes and class target **) haftmann@24218: haftmann@25002: (* class context syntax *) haftmann@24748: haftmann@25062: fun internal_remove_constraint base_sort (c, (_, (ty, _))) ctxt = haftmann@24914: let haftmann@25062: val ty' = ty haftmann@25062: |> map_atyps (fn ty as TVar ((v, 0), _) => haftmann@25062: if v = Name.aT then TVar ((v, 0), base_sort) else ty) haftmann@25062: |> SOME; haftmann@25062: in ProofContext.add_const_constraint (c, ty') ctxt end; haftmann@25002: haftmann@25002: fun remove_constraint class c ctxt = haftmann@25002: let haftmann@25002: val thy = ProofContext.theory_of ctxt; haftmann@25062: val base_sort = (#base_sort o the_class_data thy) class; haftmann@25002: val SOME entry = local_operation thy [class] c; haftmann@24914: in haftmann@25062: internal_remove_constraint base_sort (c, entry) ctxt haftmann@24914: end; haftmann@24914: haftmann@25062: fun sort_term_check sups base_sort ts ctxt = haftmann@24748: let haftmann@24901: val thy = ProofContext.theory_of ctxt; haftmann@25062: val local_operation = local_operation thy sups o fst; haftmann@25062: val typargs = Consts.typargs (ProofContext.consts_of ctxt); haftmann@25002: val constraints = these_operations thy sups |> (map o apsnd) (fst o snd); haftmann@25062: fun check_typ (c, ty) (TFree (v, _)) t = if v = Name.aT haftmann@25062: then apfst (AList.update (op =) ((c, ty), t)) else I haftmann@25062: | check_typ (c, ty) (TVar (vi, _)) t = if TypeInfer.is_param vi haftmann@25062: then apfst (AList.update (op =) ((c, ty), t)) haftmann@24901: #> apsnd (insert (op =) vi) else I haftmann@25062: | check_typ _ _ _ = I; haftmann@25062: fun check_const (c, ty) ((t, _), (_, idx)) = haftmann@25062: check_typ (c, ty) (nth (typargs (c, ty)) idx) t; haftmann@25062: fun add_const (Const c_ty) = Option.map (check_const c_ty) (local_operation c_ty) haftmann@25062: |> the_default I haftmann@24901: | add_const _ = I; haftmann@24901: val (cs, typarams) = (fold o fold_aterms) add_const ts ([], []); haftmann@25062: val subst_typ = map_type_tvar (fn var as (vi, _) => haftmann@25062: if member (op =) typarams vi then TFree (Name.aT, base_sort) else TVar var); haftmann@25062: val subst_term = map_aterms haftmann@24901: (fn t as Const (c, ty) => the_default t (AList.lookup (op =) cs (c, ty)) | t => t) haftmann@25062: #> map_types subst_typ; haftmann@25062: val ts' = map subst_term ts; haftmann@24901: val ctxt' = fold (ProofContext.add_const_constraint o apsnd SOME) constraints ctxt; wenzelm@25060: in if eq_list (op aconv) (ts, ts') then NONE else SOME (ts', ctxt') end; haftmann@24748: haftmann@25062: val uncheck = ref true; haftmann@25002: haftmann@25002: fun sort_term_uncheck sups ts ctxt = haftmann@25002: let haftmann@25062: (*FIXME abbreviations*) haftmann@25002: val thy = ProofContext.theory_of ctxt; haftmann@25062: fun rew_app f (t1 $ t2) = rew_app f t1 $ f t2 haftmann@25062: | rew_app f t = t; haftmann@25062: val rews = map (fn (c, ((t, ty), _)) => (t, Const (c, ty))) (these_operations thy sups); haftmann@25062: val rews' = (map o apfst o rew_app) (Pattern.rewrite_term thy rews []) rews; haftmann@25062: val _ = map (Thm.cterm_of thy o Logic.mk_equals) rews'; haftmann@25062: val ts' = if ! uncheck haftmann@25062: then map (Pattern.rewrite_term thy rews' []) ts else ts; wenzelm@25060: in if eq_list (op aconv) (ts, ts') then NONE else SOME (ts', ctxt) end; haftmann@25002: haftmann@25062: fun init_class_ctxt sups base_sort ctxt = haftmann@24914: let haftmann@24914: val operations = these_operations (ProofContext.theory_of ctxt) sups; haftmann@25062: fun standard_infer_types ts ctxt = haftmann@25062: let haftmann@25062: val ts' = ProofContext.standard_infer_types ctxt ts; haftmann@25062: in if eq_list (op aconv) (ts, ts') then NONE else SOME (ts', ctxt) end; haftmann@24914: in haftmann@24914: ctxt haftmann@24914: |> Variable.declare_term haftmann@25062: (Logic.mk_type (TFree (Name.aT, base_sort))) haftmann@25062: |> fold (internal_remove_constraint base_sort) operations haftmann@25062: |> Context.proof_map (Syntax.add_term_check 1 "class" haftmann@25062: (sort_term_check sups base_sort) haftmann@25062: #> Syntax.add_term_check 1 "standard" standard_infer_types haftmann@25062: #> Syntax.add_term_uncheck (~10) "class" (sort_term_uncheck sups)) haftmann@24914: end; haftmann@24901: haftmann@24901: fun init class ctxt = haftmann@24914: init_class_ctxt [class] haftmann@25062: ((#base_sort o the_class_data (ProofContext.theory_of ctxt)) class) ctxt; haftmann@24914: haftmann@24748: haftmann@24589: (* class definition *) haftmann@24218: haftmann@24218: local haftmann@24218: haftmann@24748: fun gen_class_spec prep_class prep_expr process_expr thy raw_supclasses raw_includes_elems = haftmann@24218: let haftmann@24748: val supclasses = map (prep_class thy) raw_supclasses; haftmann@25062: val (sups, base_sort) = sups_base_sort thy supclasses; haftmann@24748: val supsort = Sign.minimize_sort thy supclasses; haftmann@25038: val suplocales = map Locale.Locale sups; haftmann@24748: val (raw_elems, includes) = fold_rev (fn Locale.Elem e => apfst (cons e) haftmann@24748: | Locale.Expr i => apsnd (cons (prep_expr thy i))) raw_includes_elems ([], []); haftmann@24748: val supexpr = Locale.Merge suplocales; haftmann@24748: val supparams = (map fst o Locale.parameters_of_expr thy) supexpr; haftmann@25002: val supconsts = AList.make (the o AList.lookup (op =) (these_params thy sups)) haftmann@24748: (map fst supparams); haftmann@24748: val mergeexpr = Locale.Merge (suplocales @ includes); haftmann@24748: val constrain = Element.Constrains ((map o apsnd o map_atyps) wenzelm@24847: (fn TFree (_, sort) => TFree (Name.aT, sort)) supparams); haftmann@24748: in haftmann@24748: ProofContext.init thy haftmann@24748: |> Locale.cert_expr supexpr [constrain] haftmann@24748: |> snd haftmann@25062: |> init_class_ctxt sups base_sort haftmann@24748: |> process_expr Locale.empty raw_elems haftmann@24748: |> fst haftmann@25062: |> (fn elems => ((((sups, supconsts), (supsort, base_sort, mergeexpr)), haftmann@24748: (*FIXME*) if null includes then constrain :: elems else elems))) haftmann@24748: end; haftmann@24748: haftmann@24748: val read_class_spec = gen_class_spec Sign.intern_class Locale.intern_expr Locale.read_expr; haftmann@24748: val check_class_spec = gen_class_spec (K I) (K I) Locale.cert_expr; haftmann@24748: wenzelm@24968: fun define_class_params (name, raw_superclasses) raw_consts raw_dep_axioms other_consts thy = wenzelm@24968: let wenzelm@24968: val superclasses = map (Sign.certify_class thy) raw_superclasses; wenzelm@24968: val consts = (map o apfst o apsnd) (Sign.certify_typ thy) raw_consts; wenzelm@24968: fun add_const ((c, ty), syn) = Sign.declare_const [] (c, ty, syn) #>> Term.dest_Const; wenzelm@24968: fun mk_axioms cs thy = wenzelm@24968: raw_dep_axioms thy cs wenzelm@24968: |> (map o apsnd o map) (Sign.cert_prop thy) wenzelm@24968: |> rpair thy; haftmann@25002: fun constrain_typs class = (map o apsnd o Term.map_type_tfree) haftmann@25002: (fn (v, _) => TFree (v, [class])) wenzelm@24968: in wenzelm@24968: thy wenzelm@24968: |> Sign.add_path (Logic.const_of_class name) wenzelm@24968: |> fold_map add_const consts wenzelm@24968: ||> Sign.restore_naming thy wenzelm@24968: |-> (fn cs => mk_axioms cs wenzelm@24968: #-> (fn axioms_prop => AxClass.define_class (name, superclasses) wenzelm@24968: (map fst cs @ other_consts) axioms_prop haftmann@25002: #-> (fn class => `(fn _ => constrain_typs class cs) haftmann@25002: #-> (fn cs' => `(fn thy => AxClass.get_info thy class) haftmann@25002: #-> (fn {axioms, ...} => fold (Sign.add_const_constraint o apsnd SOME) cs' haftmann@25002: #> pair (class, (cs', axioms))))))) wenzelm@24968: end; wenzelm@24968: haftmann@25002: fun gen_class prep_spec prep_param bname haftmann@24748: raw_supclasses raw_includes_elems raw_other_consts thy = haftmann@24748: let haftmann@25038: val class = Sign.full_name thy bname; haftmann@25062: val (((sups, supconsts), (supsort, base_sort, mergeexpr)), elems_syn) = haftmann@24748: prep_spec thy raw_supclasses raw_includes_elems; wenzelm@24968: val other_consts = map (tap (Sign.the_const_type thy) o prep_param thy) raw_other_consts; haftmann@24589: fun mk_inst class param_names cs = haftmann@24589: Symtab.empty haftmann@24589: |> fold2 (fn v => fn (c, ty) => Symtab.update (v, Const haftmann@24589: (c, Term.map_type_tfree (fn (v, _) => TFree (v, [class])) ty))) param_names cs; haftmann@25062: fun fork_syntax (Element.Fixes xs) = haftmann@25062: fold_map (fn (c, ty, syn) => cons (c, syn) #> pair (c, ty, NoSyn)) xs haftmann@25062: #>> Element.Fixes haftmann@25062: | fork_syntax x = pair x; haftmann@25062: val (elems, global_syn) = fold_map fork_syntax elems_syn []; haftmann@25062: fun globalize (c, ty) = haftmann@25062: ((c, Term.map_type_tfree (K (TFree (Name.aT, base_sort))) ty), haftmann@25062: (the_default NoSyn o AList.lookup (op =) global_syn) c); haftmann@25038: fun extract_params thy = haftmann@24218: let haftmann@25062: val params = map fst (Locale.parameters_of thy class); haftmann@24218: in haftmann@25062: (params, (map globalize o snd o chop (length supconsts)) params) haftmann@24218: end; haftmann@25038: fun extract_assumes 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: fun prep_asm ((name, atts), ts) = wenzelm@25024: ((Sign.base_name name, map (Attrib.attribute_i thy) atts), haftmann@24589: (map o map_aterms) subst ts); haftmann@24218: in haftmann@25038: Locale.global_asms_of thy class haftmann@24218: |> map prep_asm haftmann@24218: end; haftmann@24218: in haftmann@24218: thy haftmann@24748: |> Locale.add_locale_i (SOME "") bname mergeexpr elems haftmann@25038: |> snd haftmann@25038: |> ProofContext.theory (`extract_params haftmann@25062: #-> (fn (all_params, params) => wenzelm@24968: define_class_params (bname, supsort) params haftmann@25038: (extract_assumes params) other_consts haftmann@25038: #-> (fn (_, (consts, axioms)) => haftmann@25038: `(fn thy => class_intro thy class sups) haftmann@24218: #-> (fn class_intro => haftmann@25062: PureThy.note_thmss_qualified "" (NameSpace.append class classN) haftmann@25062: [((introN, []), [([class_intro], [])])] haftmann@25062: #-> (fn [(_, [class_intro])] => haftmann@25038: add_class_data ((class, sups), haftmann@25062: (map fst params ~~ consts, base_sort, haftmann@25062: mk_inst class (map fst all_params) (map snd supconsts @ consts), haftmann@25062: calculate_morphism class (supconsts @ (map (fst o fst) params ~~ consts)), class_intro)) haftmann@25038: #> class_interpretation class axioms [] haftmann@25062: ))))) haftmann@25038: |> pair class haftmann@24218: end; haftmann@24218: haftmann@24218: in haftmann@24218: wenzelm@24968: val class_cmd = gen_class read_class_spec ((#1 o Term.dest_Const) oo Sign.read_const); haftmann@24748: val class = gen_class check_class_spec (K I); haftmann@24218: haftmann@24218: end; (*local*) haftmann@24218: haftmann@24218: haftmann@24589: (* definition in class target *) haftmann@24218: haftmann@25062: fun add_const_in_class class ((c, mx), rhs) thy = haftmann@24218: let wenzelm@25024: val prfx = class_prefix class; wenzelm@25024: val thy' = thy |> Sign.add_path prfx; haftmann@25062: val phi = morphism thy' class; wenzelm@25024: haftmann@25062: val c' = Sign.full_name thy' c; haftmann@25062: val ((mx', _), _) = fork_mixfix true true mx; haftmann@25062: val rhs' = (map_types Logic.unvarifyT o Morphism.term phi) rhs; haftmann@24218: val ty' = Term.fastype_of rhs'; haftmann@24657: val def = (c, Logic.mk_equals (Const (c', ty'), rhs')); haftmann@25062: val c'' = NameSpace.full (Sign.naming_of thy' |> NameSpace.add_path prfx) c; haftmann@24218: in wenzelm@25024: thy' haftmann@25038: |> Sign.hide_consts_i false [c''] wenzelm@25024: |> Sign.declare_const [] (c, ty', mx') |> snd haftmann@24218: |> Sign.parent_path haftmann@24218: |> Sign.sticky_prefix prfx haftmann@25062: |> yield_singleton (PureThy.add_defs_i false) (def, []) haftmann@25062: |>> Thm.symmetric haftmann@25062: |-> (fn def => class_interpretation class [def] haftmann@25062: [(map_types Logic.unvarifyT o Thm.prop_of) def] haftmann@25062: #> register_operation class ((c', rhs), SOME def)) haftmann@24218: |> Sign.restore_naming thy haftmann@24914: |> pair c' haftmann@24218: end; haftmann@24218: haftmann@24901: haftmann@24901: (* abbreviation in class target *) haftmann@24901: haftmann@25062: fun add_abbrev_in_class class prmode ((c, mx), rhs) thy = haftmann@24836: let wenzelm@25024: val prfx = class_prefix class; haftmann@25062: val phi = morphism thy class; haftmann@25062: haftmann@25038: val naming = Sign.naming_of thy |> NameSpace.add_path prfx |> NameSpace.add_path prfx; haftmann@25038: (*FIXME*) wenzelm@25024: val c' = NameSpace.full naming c; haftmann@25062: val rhs' = (map_types Logic.unvarifyT o Morphism.term phi) rhs; wenzelm@25020: val ty' = Term.fastype_of rhs'; haftmann@24836: in haftmann@24836: thy wenzelm@25024: |> Sign.notation true prmode [(Const (c', ty'), mx)] haftmann@25062: |> register_operation class ((c', rhs), NONE) haftmann@24914: |> pair c' haftmann@24836: end; haftmann@24836: haftmann@24218: end;