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@25462: (*classes*) 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 haftmann@25485: haftmann@25311: val init: class -> theory -> Proof.context haftmann@25485: val logical_const: string -> Markup.property list wenzelm@25104: -> (string * mixfix) * term -> theory -> theory haftmann@25485: val syntactic_const: string -> Syntax.mode -> Markup.property list wenzelm@25104: -> (string * mixfix) * term -> theory -> theory haftmann@25083: val refresh_syntax: class -> Proof.context -> Proof.context haftmann@25485: haftmann@24589: val intro_classes_tac: thm list -> tactic haftmann@24589: val default_intro_classes_tac: thm list -> tactic haftmann@25195: val prove_subclass: class * class -> thm list -> Proof.context haftmann@25195: -> theory -> theory haftmann@25485: haftmann@25485: val class_prefix: string -> string haftmann@25485: val is_class: theory -> class -> bool haftmann@25485: val these_params: theory -> sort -> (string * (string * typ)) list haftmann@24589: val print_classes: theory -> unit haftmann@24423: haftmann@25462: (*instances*) haftmann@25536: val init_instantiation: string list * sort list * sort -> theory -> local_theory haftmann@25536: val prep_spec: local_theory -> term -> term haftmann@25485: val instantiation_instance: (local_theory -> local_theory) -> local_theory -> Proof.state haftmann@25485: val prove_instantiation_instance: (Proof.context -> tactic) -> local_theory -> local_theory haftmann@25485: val conclude_instantiation: local_theory -> local_theory haftmann@25485: val instantiation_param: Proof.context -> string -> string option haftmann@25485: val confirm_declaration: string -> local_theory -> local_theory haftmann@25485: haftmann@25462: (*old axclass layer*) haftmann@25462: val axclass_cmd: bstring * xstring list haftmann@25462: -> ((bstring * Attrib.src list) * string list) list haftmann@25462: -> theory -> class * theory haftmann@25462: val classrel_cmd: xstring * xstring -> theory -> Proof.state haftmann@25462: haftmann@25462: (*old instance layer*) haftmann@25536: val instance_arity: (theory -> theory) -> arity -> theory -> Proof.state haftmann@25536: val instance_arity_cmd: bstring * xstring list * xstring -> theory -> Proof.state 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@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: haftmann@25195: fun prove_interpretation_in tac after_qed (name, expr) = haftmann@25195: Locale.interpretation_in_locale haftmann@25195: (ProofContext.theory after_qed) (name, expr) haftmann@25195: #> Proof.global_terminal_proof haftmann@25195: (Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE) haftmann@25195: #> ProofContext.theory_of; haftmann@25195: 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@25485: (** primitive axclass and instance commands **) 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@25536: o 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@25502: val instance_arity_cmd = haftmann@25502: gen_instance (Logic.mk_arities oo Sign.read_arity) AxClass.add_arity I; haftmann@24589: val classrel = haftmann@25536: gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel)) AxClass.add_classrel I; haftmann@24589: val classrel_cmd = haftmann@25536: gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) AxClass.add_classrel I; haftmann@24218: haftmann@24218: end; (*local*) 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@25083: inst: term option list haftmann@25083: (*canonical interpretation*), haftmann@25062: morphism: morphism, haftmann@25062: (*partial morphism of canonical interpretation*) haftmann@24657: intro: thm, haftmann@24657: defs: thm list, haftmann@25368: operations: (string * (class * (typ * term))) list 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@25368: (defs, operations)) = haftmann@25062: ClassData { consts = consts, base_sort = base_sort, inst = inst, haftmann@25062: morphism = morphism, intro = intro, defs = defs, haftmann@25368: operations = operations }; haftmann@25062: fun map_class_data f (ClassData { consts, base_sort, inst, morphism, intro, haftmann@25368: defs, operations }) = haftmann@25062: mk_class_data (f ((consts, base_sort, inst, morphism, intro), haftmann@25368: (defs, operations))); haftmann@25038: fun merge_class_data _ (ClassData { consts = consts, haftmann@25062: base_sort = base_sort, inst = inst, morphism = morphism, intro = intro, haftmann@25368: defs = defs1, operations = operations1 }, haftmann@25062: ClassData { consts = _, base_sort = _, inst = _, morphism = _, intro = _, haftmann@25368: defs = defs2, operations = operations2 }) = haftmann@25062: mk_class_data ((consts, base_sort, inst, morphism, intro), haftmann@24914: (Thm.merge_thms (defs1, defs2), haftmann@25368: 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@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@25163: fun add_class_data ((class, superclasses), (cs, base_sort, inst, phi, intro)) thy = haftmann@25002: let haftmann@25368: val operations = map (fn (v_ty as (_, ty), (c, _)) => haftmann@25368: (c, (class, (ty, Free v_ty)))) cs; haftmann@25002: val cs = (map o pairself) fst cs; haftmann@25038: val add_class = Graph.new_node (class, haftmann@25368: mk_class_data ((cs, base_sort, map (SOME o Const) 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@25368: fun register_operation class (c, (t, some_def)) thy = haftmann@25062: let haftmann@25368: val base_sort = (#base_sort o the_class_data thy) class; haftmann@25239: val prep_typ = map_atyps haftmann@25368: (fn TVar (vi as (v, _), sort) => if Name.aT = v haftmann@25368: then TFree (v, base_sort) else TVar (vi, sort)); haftmann@25368: val t' = map_types prep_typ t; haftmann@25368: val ty' = Term.fastype_of t'; haftmann@25062: in haftmann@25062: thy haftmann@25062: |> (ClassData.map o Graph.map_node class o map_class_data o apsnd) haftmann@25368: (fn (defs, operations) => haftmann@25096: (fold cons (the_list some_def) defs, haftmann@25368: (c, (class, (ty', t'))) :: 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@25209: 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@25083: val inst = (#inst o 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 ballarin@25094: ||> prove_interpretation tac ((false, prfx), []) (Locale.Locale class) ballarin@25094: (inst, map (fn def => (("", []), def)) 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 haftmann@25268: Method.intros_tac (class_trivs @ class_intros @ axclass_intros) facts 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@25195: fun subclass_rule thy (sub, sup) = haftmann@25195: let haftmann@25195: val ctxt = Locale.init sub thy; haftmann@25195: val ctxt_thy = ProofContext.init thy; haftmann@25195: val props = haftmann@25195: Locale.global_asms_of thy sup haftmann@25195: |> maps snd haftmann@25195: |> map (ObjectLogic.ensure_propT thy); haftmann@25195: fun tac { prems, context } = haftmann@25195: Locale.intro_locales_tac true context prems haftmann@25195: ORELSE ALLGOALS assume_tac; haftmann@25195: in haftmann@25195: Goal.prove_multi ctxt [] [] props tac haftmann@25195: |> map (Assumption.export false ctxt ctxt_thy) haftmann@25195: |> Variable.export ctxt ctxt_thy haftmann@25195: end; haftmann@25195: haftmann@25195: fun prove_single_subclass (sub, sup) thms ctxt thy = haftmann@25195: let haftmann@25195: val ctxt_thy = ProofContext.init thy; haftmann@25195: val subclass_rule = Conjunction.intr_balanced thms haftmann@25195: |> Assumption.export false ctxt ctxt_thy haftmann@25195: |> singleton (Variable.export ctxt ctxt_thy); haftmann@25195: val sub_inst = Thm.ctyp_of thy (TVar ((Name.aT, 0), [sub])); haftmann@25195: val sub_ax = #axioms (AxClass.get_info thy sub); haftmann@25195: val classrel = haftmann@25195: #intro (AxClass.get_info thy sup) haftmann@25195: |> Drule.instantiate' [SOME sub_inst] [] haftmann@25195: |> OF_LAST (subclass_rule OF sub_ax) haftmann@25195: |> strip_all_ofclass thy (Sign.super_classes thy sup) haftmann@25195: |> Thm.strip_shyps haftmann@25195: in haftmann@25195: thy haftmann@25195: |> AxClass.add_classrel classrel haftmann@25195: |> prove_interpretation_in (ALLGOALS (ProofContext.fact_tac thms)) haftmann@25195: I (sub, Locale.Locale sup) haftmann@25195: |> ClassData.map (Graph.add_edge (sub, sup)) haftmann@25195: end; haftmann@25195: haftmann@25195: fun prove_subclass (sub, sup) thms ctxt thy = haftmann@25195: let haftmann@25268: val classes = ClassData.get thy; haftmann@25268: val is_sup = not o null o curry (Graph.irreducible_paths classes) sub; haftmann@25268: val supclasses = Graph.all_succs classes [sup] |> filter_out is_sup; haftmann@25195: fun transform sup' = subclass_rule thy (sup, sup') |> map (fn thm => thm OF thms); haftmann@25195: in haftmann@25195: thy haftmann@25195: |> fold_rev (fn sup' => prove_single_subclass (sub, sup') haftmann@25195: (transform sup') ctxt) supclasses haftmann@25195: end; haftmann@25195: haftmann@24218: haftmann@24589: (** classes and class target **) haftmann@24218: haftmann@25002: (* class context syntax *) haftmann@24748: haftmann@25083: structure ClassSyntax = ProofDataFun( haftmann@25083: type T = { haftmann@25368: local_constraints: (string * typ) list, haftmann@25368: global_constraints: (string * typ) list, haftmann@25083: base_sort: sort, haftmann@25368: operations: (string * (typ * term)) list, haftmann@25195: unchecks: (term * term) list, haftmann@25083: passed: bool haftmann@25368: }; haftmann@25368: fun init _ = { haftmann@25368: local_constraints = [], haftmann@25368: global_constraints = [], haftmann@25368: base_sort = [], haftmann@25368: operations = [], haftmann@25368: unchecks = [], haftmann@25368: passed = true haftmann@25368: };; haftmann@25083: ); haftmann@25083: wenzelm@25344: fun synchronize_syntax sups base_sort ctxt = haftmann@24914: let wenzelm@25344: val thy = ProofContext.theory_of ctxt; haftmann@25368: fun subst_class_typ sort = map_atyps haftmann@25368: (fn TFree _ => TVar ((Name.aT, 0), sort) | ty' => ty'); haftmann@25083: val operations = these_operations thy sups; haftmann@25368: val local_constraints = haftmann@25368: (map o apsnd) (subst_class_typ base_sort o fst o snd) operations; haftmann@25368: val global_constraints = haftmann@25368: (map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations; wenzelm@25318: fun declare_const (c, _) = wenzelm@25318: let val b = Sign.base_name c wenzelm@25344: in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end; haftmann@25368: val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) operations; haftmann@25083: in haftmann@25083: ctxt haftmann@25368: |> fold declare_const local_constraints haftmann@25368: |> fold (ProofContext.add_const_constraint o apsnd SOME) local_constraints haftmann@25368: |> ClassSyntax.put { haftmann@25368: local_constraints = local_constraints, haftmann@25368: global_constraints = global_constraints, haftmann@25083: base_sort = base_sort, haftmann@25368: operations = (map o apsnd) snd operations, haftmann@25195: unchecks = unchecks, haftmann@25083: passed = false haftmann@25368: } haftmann@25083: end; haftmann@25083: haftmann@25083: fun refresh_syntax class 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; wenzelm@25344: in synchronize_syntax [class] base_sort ctxt end; haftmann@24914: haftmann@25368: val mark_passed = ClassSyntax.map haftmann@25368: (fn { local_constraints, global_constraints, base_sort, operations, unchecks, passed } => haftmann@25368: { local_constraints = local_constraints, global_constraints = global_constraints, haftmann@25368: base_sort = base_sort, operations = operations, unchecks = unchecks, passed = true }); haftmann@25083: haftmann@25083: fun sort_term_check ts ctxt = haftmann@24748: let haftmann@25368: val { local_constraints, global_constraints, base_sort, operations, passed, ... } = haftmann@25368: ClassSyntax.get ctxt; haftmann@25368: fun check_improve (Const (c, ty)) = (case AList.lookup (op =) local_constraints c haftmann@25368: of SOME ty0 => (case try (Type.raw_match (ty0, ty)) Vartab.empty haftmann@25368: of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0) haftmann@25368: of SOME (_, TVar (tvar as (vi, _))) => haftmann@25368: if TypeInfer.is_param vi then cons tvar else I haftmann@25368: | _ => I) haftmann@25368: | NONE => I) haftmann@25368: | NONE => I) haftmann@25368: | check_improve _ = I; haftmann@25368: val improvements = (fold o fold_aterms) check_improve ts []; haftmann@25368: val ts' = (map o map_types o map_atyps) (fn ty as TVar tvar => haftmann@25368: if member (op =) improvements tvar haftmann@25368: then TFree (Name.aT, base_sort) else ty | ty => ty) ts; haftmann@25368: fun check t0 = Envir.expand_term (fn Const (c, ty) => (case AList.lookup (op =) operations c haftmann@25368: of SOME (ty0, t) => haftmann@25368: if Type.typ_instance (ProofContext.tsig_of ctxt) (ty, ty0) haftmann@25368: then SOME (ty0, check t) else NONE haftmann@25368: | NONE => NONE) haftmann@25368: | _ => NONE) t0; haftmann@25368: val ts'' = map check ts'; haftmann@25368: in if eq_list (op aconv) (ts, ts'') andalso passed then NONE haftmann@25083: else haftmann@25083: ctxt haftmann@25368: |> fold (ProofContext.add_const_constraint o apsnd SOME) global_constraints haftmann@25083: |> mark_passed haftmann@25368: |> pair ts'' haftmann@25083: |> SOME haftmann@25083: end; haftmann@24748: haftmann@25083: fun sort_term_uncheck ts ctxt = haftmann@25002: let haftmann@25002: val thy = ProofContext.theory_of ctxt; haftmann@25368: val unchecks = (#unchecks o ClassSyntax.get) ctxt; haftmann@25462: val ts' = map (Pattern.rewrite_term thy unchecks []) ts; wenzelm@25060: in if eq_list (op aconv) (ts, ts') then NONE else SOME (ts', ctxt) end; haftmann@25002: wenzelm@25344: fun init_ctxt sups base_sort ctxt = haftmann@25083: ctxt haftmann@25083: |> Variable.declare_term haftmann@25083: (Logic.mk_type (TFree (Name.aT, base_sort))) wenzelm@25344: |> synchronize_syntax sups base_sort haftmann@25083: |> Context.proof_map ( haftmann@25083: Syntax.add_term_check 0 "class" sort_term_check haftmann@25103: #> Syntax.add_term_uncheck 0 "class" sort_term_uncheck) haftmann@24901: haftmann@25311: fun init class thy = haftmann@25311: thy haftmann@25311: |> Locale.init class wenzelm@25344: |> init_ctxt [class] ((#base_sort o the_class_data thy) class); 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@25209: val sups = filter (is_class thy) supclasses; haftmann@25209: fun the_base_sort class = lookup_class_data thy class haftmann@25209: |> Option.map #base_sort haftmann@25209: |> the_default [class]; haftmann@25209: val base_sort = Sign.minimize_sort thy (maps the_base_sort 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 wenzelm@25344: |> init_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; haftmann@25083: fun add_const ((c, ty), syn) = haftmann@25083: Sign.declare_const [] (c, Type.strip_sorts 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@25163: fun mk_inst class cs = haftmann@25163: (map o apsnd o Term.map_type_tfree) (fn (v, _) => TFree (v, [class])) 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 []; wenzelm@25326: 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@25311: |> ProofContext.theory_of haftmann@25311: |> `extract_params haftmann@25311: |-> (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@25163: mk_inst class (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@25311: )))) haftmann@25268: |> init class haftmann@25038: |> pair class haftmann@24218: end; haftmann@24218: wenzelm@25326: fun read_const thy = #1 o Term.dest_Const o ProofContext.read_const (ProofContext.init thy); wenzelm@25326: haftmann@24218: in haftmann@24218: wenzelm@25326: val class_cmd = gen_class read_class_spec read_const; haftmann@24748: val class = gen_class check_class_spec (K I); haftmann@24218: haftmann@24218: end; (*local*) haftmann@24218: haftmann@24218: haftmann@25485: (* class target *) haftmann@24218: haftmann@25485: fun logical_const class pos ((c, mx), dict) 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@25239: val dict' = Morphism.term phi dict; haftmann@25239: val dict_def = map_types Logic.unvarifyT dict'; haftmann@25239: val ty' = Term.fastype_of dict_def; haftmann@25083: val ty'' = Type.strip_sorts ty'; haftmann@25239: val def_eq = Logic.mk_equals (Const (c', ty'), dict_def); haftmann@24218: in wenzelm@25024: thy' haftmann@25096: |> Sign.declare_const pos (c, ty'', mx) |> snd haftmann@25518: |> Thm.add_def false false (c, def_eq) haftmann@25062: |>> Thm.symmetric haftmann@25083: |-> (fn def => class_interpretation class [def] [Thm.prop_of def] haftmann@25368: #> register_operation class (c', (dict', SOME (Thm.varifyT def)))) haftmann@24218: |> Sign.restore_naming thy haftmann@25083: |> Sign.add_const_constraint (c', SOME ty') haftmann@24218: end; haftmann@24218: haftmann@25485: fun syntactic_const class prmode pos ((c, mx), rhs) thy = haftmann@24836: let wenzelm@25024: val prfx = class_prefix class; haftmann@25096: val thy' = thy |> Sign.add_path prfx; haftmann@25062: val phi = morphism thy class; haftmann@25062: haftmann@25096: val c' = Sign.full_name thy' c; haftmann@25146: val rews = map (Logic.dest_equals o Thm.prop_of) (these_defs thy' [class]) haftmann@25146: val rhs' = (Pattern.rewrite_term thy rews [] o Morphism.term phi) rhs; haftmann@25239: val ty' = Logic.unvarifyT (Term.fastype_of rhs'); haftmann@24836: in haftmann@25096: thy' haftmann@25146: |> Sign.add_abbrev (#1 prmode) pos (c, map_types Type.strip_sorts rhs') |> snd haftmann@25096: |> Sign.add_const_constraint (c', SOME ty') wenzelm@25024: |> Sign.notation true prmode [(Const (c', ty'), mx)] haftmann@25368: |> register_operation class (c', (rhs', NONE)) haftmann@25096: |> Sign.restore_naming thy haftmann@24836: end; haftmann@24836: haftmann@25462: haftmann@25462: (** instantiation target **) haftmann@25462: haftmann@25462: (* bookkeeping *) haftmann@25462: haftmann@25462: datatype instantiation = Instantiation of { haftmann@25536: arities: string list * sort list * sort, haftmann@25462: params: ((string * string) * (string * typ)) list haftmann@25574: (*(instantiation const, type constructor), (local instantiation parameter, typ)*) haftmann@25462: } haftmann@25462: haftmann@25462: structure Instantiation = ProofDataFun haftmann@25462: ( haftmann@25462: type T = instantiation haftmann@25536: fun init _ = Instantiation { arities = ([], [], []), params = [] }; haftmann@25462: ); haftmann@25462: haftmann@25485: fun mk_instantiation (arities, params) = haftmann@25485: Instantiation { arities = arities, params = params }; haftmann@25514: fun get_instantiation lthy = case Instantiation.get (LocalTheory.target_of lthy) haftmann@25485: of Instantiation data => data; haftmann@25514: fun map_instantiation f = (LocalTheory.target o Instantiation.map) haftmann@25514: (fn Instantiation { arities, params } => mk_instantiation (f (arities, params))); haftmann@25462: haftmann@25514: fun the_instantiation lthy = case get_instantiation lthy haftmann@25536: of { arities = ([], [], []), ... } => error "No instantiation target" haftmann@25485: | data => data; haftmann@25462: haftmann@25485: val instantiation_params = #params o get_instantiation; haftmann@25462: haftmann@25514: fun instantiation_param lthy v = instantiation_params lthy haftmann@25462: |> find_first (fn (_, (v', _)) => v = v') haftmann@25462: |> Option.map (fst o fst); haftmann@25462: haftmann@25514: fun confirm_declaration c = (map_instantiation o apsnd) haftmann@25485: (filter_out (fn (_, (c', _)) => c' = c)); haftmann@25485: haftmann@25462: haftmann@25462: (* syntax *) haftmann@25462: haftmann@25597: fun subst_param thy params = map_aterms (fn t as Const (c, ty) => haftmann@25597: (case AxClass.inst_tyco_of thy (c, ty) haftmann@25536: of SOME tyco => (case AList.lookup (op =) params (c, tyco) haftmann@25536: of SOME v_ty => Free v_ty haftmann@25536: | NONE => t) haftmann@25536: | NONE => t) haftmann@25536: | t => t); haftmann@25536: haftmann@25536: fun prep_spec lthy = haftmann@25536: let haftmann@25536: val thy = ProofContext.theory_of lthy; haftmann@25536: val params = instantiation_params lthy; haftmann@25536: in subst_param thy params end; haftmann@25536: haftmann@25514: fun inst_term_check ts lthy = haftmann@25462: let haftmann@25514: val params = instantiation_params lthy; haftmann@25514: val tsig = ProofContext.tsig_of lthy; haftmann@25514: val thy = ProofContext.theory_of lthy; haftmann@25462: haftmann@25597: fun check_improve (Const (c, ty)) = (case AxClass.inst_tyco_of thy (c, ty) haftmann@25462: of SOME tyco => (case AList.lookup (op =) params (c, tyco) haftmann@25502: of SOME (_, ty') => perhaps (try (Type.typ_match tsig (ty, ty'))) haftmann@25462: | NONE => I) haftmann@25462: | NONE => I) haftmann@25462: | check_improve _ = I; haftmann@25462: val improvement = (fold o fold_aterms) check_improve ts Vartab.empty; haftmann@25462: val ts' = (map o map_types) (Envir.typ_subst_TVars improvement) ts; haftmann@25536: val ts'' = map (subst_param thy params) ts'; haftmann@25514: in if eq_list (op aconv) (ts, ts'') then NONE else SOME (ts'', lthy) end; haftmann@25462: haftmann@25514: fun inst_term_uncheck ts lthy = haftmann@25462: let haftmann@25514: val params = instantiation_params lthy; haftmann@25462: val ts' = (map o map_aterms) (fn t as Free (v, ty) => haftmann@25462: (case get_first (fn ((c, _), (v', _)) => if v = v' then SOME c else NONE) params haftmann@25462: of SOME c => Const (c, ty) haftmann@25462: | NONE => t) haftmann@25462: | t => t) ts; haftmann@25514: in if eq_list (op aconv) (ts, ts') then NONE else SOME (ts', lthy) end; haftmann@25462: haftmann@25462: haftmann@25462: (* target *) haftmann@25462: haftmann@25485: val sanatize_name = (*FIXME*) haftmann@25485: let haftmann@25574: fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s haftmann@25574: orelse s = "'" orelse s = "_"; haftmann@25485: val is_junk = not o is_valid andf Symbol.is_regular; haftmann@25485: val junk = Scan.many is_junk; haftmann@25485: val scan_valids = Symbol.scanner "Malformed input" haftmann@25485: ((junk |-- haftmann@25485: (Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode) haftmann@25485: --| junk)) haftmann@25485: -- Scan.repeat ((Scan.many1 is_valid >> implode) --| junk) >> op ::); haftmann@25485: in haftmann@25485: explode #> scan_valids #> implode haftmann@25485: end; haftmann@25485: haftmann@25536: fun init_instantiation (tycos, sorts, sort) thy = haftmann@25462: let haftmann@25536: val _ = if null tycos then error "At least one arity must be given" else (); haftmann@25536: val _ = map (the_class_data thy) sort; haftmann@25536: val vs = map TFree (Name.names Name.context Name.aT sorts); haftmann@25485: fun type_name "*" = "prod" haftmann@25485: | type_name "+" = "sum" haftmann@25485: | type_name s = sanatize_name (NameSpace.base s); (*FIXME*) haftmann@25597: fun get_param tyco (param, (c, ty)) = if can (AxClass.param_of_inst thy) (c, tyco) haftmann@25597: then NONE else SOME ((AxClass.unoverload_const thy (c, ty), tyco), haftmann@25536: (param ^ "_" ^ type_name tyco, map_atyps (K (Type (tyco, vs))) ty)); haftmann@25536: val params = map_product get_param tycos (these_params thy sort) |> map_filter I; haftmann@25485: in haftmann@25485: thy haftmann@25485: |> ProofContext.init haftmann@25536: |> Instantiation.put (mk_instantiation ((tycos, sorts, sort), params)) haftmann@25536: |> fold (Variable.declare_term o Logic.mk_type) vs haftmann@25574: |> fold (Variable.declare_names o Free o snd) params haftmann@25536: |> fold (fn tyco => ProofContext.add_arity (tyco, sorts, sort)) tycos haftmann@25485: |> Context.proof_map ( haftmann@25485: Syntax.add_term_check 0 "instance" inst_term_check haftmann@25485: #> Syntax.add_term_uncheck 0 "instance" inst_term_uncheck) haftmann@25485: end; haftmann@25485: haftmann@25485: fun gen_instantiation_instance do_proof after_qed lthy = haftmann@25485: let haftmann@25536: val (tycos, sorts, sort) = (#arities o the_instantiation) lthy; haftmann@25536: val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, sorts, sort)) tycos; haftmann@25462: fun after_qed' results = haftmann@25462: LocalTheory.theory (fold (AxClass.add_arity o Thm.varifyT) results) haftmann@25462: #> after_qed; haftmann@25462: in haftmann@25462: lthy haftmann@25462: |> do_proof after_qed' arities_proof haftmann@25462: end; haftmann@25462: haftmann@25485: val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts => haftmann@25462: Proof.theorem_i NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts)); haftmann@25462: haftmann@25485: fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed => haftmann@25502: fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t haftmann@25502: (fn {context, ...} => tac context)) ts) lthy) I; haftmann@25462: haftmann@25462: fun conclude_instantiation lthy = haftmann@25462: let haftmann@25485: val { arities, params } = the_instantiation lthy; haftmann@25536: val (tycos, sorts, sort) = arities; haftmann@25462: val thy = ProofContext.theory_of lthy; haftmann@25597: val _ = map (fn tyco => if Sign.of_sort thy haftmann@25462: (Type (tyco, map TFree (Name.names Name.context Name.aT sorts)), sort) haftmann@25462: then () else error ("Missing instance proof for type " ^ quote (Sign.extern_type thy tyco))) haftmann@25597: tycos; haftmann@25597: in lthy end; haftmann@25462: haftmann@24218: end;