haftmann@29358: (* Title: Pure/Isar/class_target.ML haftmann@24218: Author: Florian Haftmann, TU Muenchen haftmann@24218: haftmann@29358: Type classes derived from primitive axclasses and locales - mechanisms. haftmann@24218: *) haftmann@24218: haftmann@29358: signature CLASS_TARGET = haftmann@24218: sig haftmann@25462: (*classes*) haftmann@29358: val register: class -> class list -> ((string * typ) * (string * typ)) list haftmann@29526: -> sort -> morphism -> thm option -> thm option -> thm haftmann@29526: -> theory -> theory haftmann@29558: val register_subclass: class * class -> morphism option -> Element.witness option haftmann@29558: -> morphism -> theory -> theory haftmann@25485: haftmann@29526: val is_class: theory -> class -> bool haftmann@29526: val base_sort: theory -> class -> sort haftmann@29526: val rules: theory -> class -> thm option * thm haftmann@29526: val these_params: theory -> sort -> (string * (class * (string * typ))) list haftmann@29526: val these_defs: theory -> sort -> thm list haftmann@29577: val these_operations: theory -> sort -> (string * (class * (typ * term))) list haftmann@29526: val print_classes: theory -> unit haftmann@29526: haftmann@29358: val begin: class list -> sort -> Proof.context -> Proof.context haftmann@25311: val init: class -> theory -> Proof.context wenzelm@28017: val declare: class -> Properties.T haftmann@27690: -> (string * mixfix) * term -> theory -> theory wenzelm@28017: val abbrev: class -> Syntax.mode -> Properties.T wenzelm@25104: -> (string * mixfix) * term -> theory -> theory haftmann@29526: val class_prefix: string -> string haftmann@25083: val refresh_syntax: class -> Proof.context -> Proof.context haftmann@29632: val redeclare_operations: theory -> sort -> Proof.context -> Proof.context haftmann@25485: haftmann@25462: (*instances*) haftmann@26247: val init_instantiation: string list * (string * sort) list * sort haftmann@26247: -> theory -> local_theory haftmann@26247: val instantiation_instance: (local_theory -> local_theory) haftmann@26247: -> local_theory -> Proof.state haftmann@26247: val prove_instantiation_instance: (Proof.context -> tactic) haftmann@26247: -> local_theory -> local_theory haftmann@28666: val prove_instantiation_exit: (Proof.context -> tactic) haftmann@28666: -> local_theory -> theory haftmann@28666: val prove_instantiation_exit_result: (morphism -> 'a -> 'b) haftmann@28666: -> (Proof.context -> 'b -> tactic) -> 'a -> local_theory -> 'b * theory haftmann@25485: val conclude_instantiation: local_theory -> local_theory haftmann@25603: val instantiation_param: local_theory -> string -> string option haftmann@25485: val confirm_declaration: string -> local_theory -> local_theory haftmann@25603: val pretty_instantiation: local_theory -> Pretty.T haftmann@26259: val type_name: string -> string haftmann@25485: haftmann@29526: val intro_classes_tac: thm list -> tactic haftmann@29526: val default_intro_tac: Proof.context -> thm list -> tactic haftmann@29526: haftmann@25462: (*old axclass layer*) haftmann@25462: val axclass_cmd: bstring * xstring list wenzelm@28084: -> (Attrib.binding * string list) list haftmann@25462: -> theory -> class * theory haftmann@25462: val classrel_cmd: xstring * xstring -> theory -> Proof.state 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@29358: structure Class_Target : CLASS_TARGET = haftmann@24218: struct haftmann@24218: 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@28715: haftmann@28715: (* static part *) haftmann@24218: consts: (string * string) list haftmann@24836: (*locale parameter ~> constant name*), haftmann@25062: base_sort: sort, haftmann@29545: base_morph: morphism haftmann@29439: (*static part of canonical morphism*), haftmann@25618: assm_intro: thm option, haftmann@25618: of_class: thm, haftmann@25618: axiom: thm option, haftmann@28715: haftmann@28715: (* dynamic part *) haftmann@24657: defs: thm list, haftmann@25368: operations: (string * (class * (typ * term))) list haftmann@28715: haftmann@24657: }; haftmann@24218: haftmann@29526: fun rep_class_data (ClassData data) = data; haftmann@29526: fun mk_class_data ((consts, base_sort, base_morph, assm_intro, of_class, axiom), haftmann@25368: (defs, operations)) = haftmann@29526: ClassData { consts = consts, base_sort = base_sort, haftmann@29439: base_morph = base_morph, assm_intro = assm_intro, of_class = of_class, axiom = axiom, haftmann@25618: defs = defs, operations = operations }; haftmann@29526: fun map_class_data f (ClassData { consts, base_sort, base_morph, assm_intro, haftmann@28715: of_class, axiom, defs, operations }) = haftmann@29526: mk_class_data (f ((consts, base_sort, base_morph, assm_intro, of_class, axiom), haftmann@25368: (defs, operations))); haftmann@25038: fun merge_class_data _ (ClassData { consts = consts, haftmann@29526: base_sort = base_sort, base_morph = base_morph, assm_intro = assm_intro, haftmann@25618: of_class = of_class, axiom = axiom, defs = defs1, operations = operations1 }, haftmann@29526: ClassData { consts = _, base_sort = _, base_morph = _, assm_intro = _, haftmann@25618: of_class = _, axiom = _, defs = defs2, operations = operations2 }) = haftmann@29526: mk_class_data ((consts, base_sort, base_morph, assm_intro, of_class, axiom), 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@29509: val heritage = Graph.all_preds o ClassData.get; haftmann@29509: 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@26518: (map o apsnd) haftmann@26518: (fn c => (class, (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@29526: val base_sort = #base_sort oo the_class_data; haftmann@29526: haftmann@29526: fun rules thy class = haftmann@29526: let val { axiom, of_class, ... } = the_class_data thy class haftmann@29526: in (axiom, of_class) end; haftmann@29526: haftmann@29526: fun all_assm_intros thy = haftmann@25618: Graph.fold (fn (_, (data, _)) => fold (insert Thm.eq_thm) haftmann@25618: ((the_list o #assm_intro o rep_class_data) data)) (ClassData.get thy) []; haftmann@24218: haftmann@29526: fun these_defs thy = maps (#defs o the_class_data thy) o ancestry thy; haftmann@29526: fun these_operations thy = maps (#operations o the_class_data thy) o ancestry thy; haftmann@29358: haftmann@29526: val base_morphism = #base_morph oo the_class_data; haftmann@29526: fun morphism thy class = base_morphism thy class haftmann@29526: $> Element.eq_morphism thy (these_defs thy [class]); haftmann@28715: 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: ((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@29526: fun register class sups params base_sort morph haftmann@29358: axiom assm_intro of_class thy = haftmann@25002: let haftmann@25368: val operations = map (fn (v_ty as (_, ty), (c, _)) => haftmann@25683: (c, (class, (ty, Free v_ty)))) params; haftmann@25038: val add_class = Graph.new_node (class, haftmann@25683: mk_class_data (((map o pairself) fst params, base_sort, haftmann@29526: morph, assm_intro, of_class, axiom), ([], operations))) haftmann@29526: #> fold (curry Graph.add_edge class) sups; haftmann@25618: in ClassData.map add_class thy end; haftmann@24218: haftmann@29526: fun activate_defs class thms thy = haftmann@29526: let haftmann@29526: val eq_morph = Element.eq_morphism thy thms; haftmann@29526: fun amend cls thy = Locale.amend_registration eq_morph haftmann@29526: (cls, morphism thy cls) thy; haftmann@29526: in fold amend (heritage thy [class]) thy end; haftmann@29526: haftmann@25368: fun register_operation class (c, (t, some_def)) thy = haftmann@25062: let haftmann@29358: val base_sort = base_sort thy class; haftmann@29439: val prep_typ = map_type_tfree haftmann@29439: (fn (v, sort) => if Name.aT = v haftmann@29439: then TFree (v, base_sort) else TVar ((v, 0), 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@29526: |> is_some some_def ? activate_defs class (the_list some_def) haftmann@25062: end; haftmann@24218: haftmann@29558: fun register_subclass (sub, sup) some_dep_morph some_wit export thy = haftmann@25618: let haftmann@29558: val intros = (snd o rules thy) sup :: map_filter I haftmann@29558: [Option.map (Drule.standard' o Element.conclude_witness) some_wit, haftmann@29558: (fst o rules thy) sub]; haftmann@29610: val tac = EVERY (map (TRYALL o Tactic.rtac) intros); haftmann@29558: val classrel = Goal.prove_global thy [] [] (Logic.mk_classrel (sub, sup)) haftmann@29558: (K tac); haftmann@29509: val diff_sort = Sign.complete_sort thy [sup] haftmann@29509: |> subtract (op =) (Sign.complete_sort thy [sub]); haftmann@25618: in haftmann@25618: thy haftmann@25618: |> AxClass.add_classrel classrel haftmann@25618: |> ClassData.map (Graph.add_edge (sub, sup)) haftmann@29526: |> activate_defs sub (these_defs thy diff_sort) haftmann@29610: |> fold (fn dep_morph => Locale.add_dependency sub (sup, haftmann@29610: dep_morph $> Element.satisfy_morphism (the_list some_wit) $> export)) haftmann@29610: (the_list some_dep_morph) haftmann@29558: |> (fn thy => fold_rev Locale.activate_global_facts haftmann@29558: (Locale.get_registrations thy) thy) haftmann@24218: end; haftmann@24218: haftmann@24218: haftmann@24589: (** classes and class target **) haftmann@24218: haftmann@25002: (* class context syntax *) haftmann@24748: haftmann@29577: fun these_unchecks thy = haftmann@29577: map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) o these_operations thy; haftmann@29577: haftmann@29632: fun redeclare_const thy c = haftmann@29632: let val b = Sign.base_name c haftmann@29632: in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end; haftmann@29632: haftmann@29577: fun synchronize_class_syntax sort base_sort ctxt = haftmann@24914: let wenzelm@25344: val thy = ProofContext.theory_of ctxt; haftmann@26596: val algebra = Sign.classes_of thy; haftmann@29577: val operations = these_operations thy sort; haftmann@26518: fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort))); haftmann@26518: val primary_constraints = haftmann@25368: (map o apsnd) (subst_class_typ base_sort o fst o snd) operations; haftmann@26518: val secondary_constraints = haftmann@25368: (map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations; haftmann@26518: fun improve (c, ty) = (case AList.lookup (op =) primary_constraints c haftmann@26238: of SOME ty' => (case try (Type.raw_match (ty', ty)) Vartab.empty haftmann@26238: of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0) haftmann@26596: of SOME (_, ty' as TVar (tvar as (vi, sort))) => haftmann@26238: if TypeInfer.is_param vi haftmann@26596: andalso Sorts.sort_le algebra (base_sort, sort) haftmann@26596: then SOME (ty', TFree (Name.aT, base_sort)) haftmann@26596: else NONE haftmann@26238: | _ => NONE) haftmann@26238: | NONE => NONE) haftmann@26238: | NONE => NONE) haftmann@26238: fun subst (c, ty) = Option.map snd (AList.lookup (op =) operations c); haftmann@29577: val unchecks = these_unchecks thy sort; haftmann@25083: in haftmann@25083: ctxt haftmann@29632: |> fold (redeclare_const thy o fst) primary_constraints haftmann@26518: |> Overloading.map_improvable_syntax (K (((primary_constraints, secondary_constraints), haftmann@26730: (((improve, subst), true), unchecks)), false)) haftmann@26518: |> Overloading.set_primary_constraints haftmann@25083: end; haftmann@25083: haftmann@25083: fun refresh_syntax class ctxt = haftmann@25002: let haftmann@25002: val thy = ProofContext.theory_of ctxt; haftmann@29358: val base_sort = base_sort thy class; haftmann@26238: in synchronize_class_syntax [class] base_sort ctxt end; haftmann@25002: haftmann@29632: fun redeclare_operations thy sort = haftmann@29632: fold (redeclare_const thy o fst) (these_operations thy sort); haftmann@29632: haftmann@29577: fun begin sort base_sort ctxt = haftmann@25083: ctxt haftmann@25083: |> Variable.declare_term haftmann@25083: (Logic.mk_type (TFree (Name.aT, base_sort))) haftmann@29577: |> synchronize_class_syntax sort base_sort haftmann@26238: |> Overloading.add_improvable_syntax; haftmann@24901: haftmann@25311: fun init class thy = haftmann@25311: thy haftmann@29509: |> Locale.init class haftmann@29358: |> begin [class] (base_sort thy class); haftmann@24914: haftmann@24748: haftmann@27690: (* class target *) haftmann@27690: haftmann@29526: val class_prefix = Logic.const_of_class o Sign.base_name; haftmann@29526: haftmann@27690: fun declare class pos ((c, mx), dict) thy = haftmann@27690: let haftmann@29577: val morph = morphism thy class; haftmann@29577: val b = Morphism.binding morph (Binding.name c); haftmann@29577: val b_def = Morphism.binding morph (Binding.name (c ^ "_dict")); haftmann@29577: val c' = Sign.full_name thy b; haftmann@29439: val dict' = Morphism.term morph dict; haftmann@29439: val ty' = Term.fastype_of dict'; haftmann@29577: val def_eq = Logic.mk_equals (Const (c', ty'), dict') haftmann@29577: |> map_types Type.strip_sorts; haftmann@27690: in haftmann@29577: thy haftmann@29577: |> Sign.declare_const pos ((b, Type.strip_sorts ty'), mx) haftmann@29577: |> snd haftmann@29577: |> Thm.add_def false false (b_def, def_eq) haftmann@29577: |>> Thm.varifyT haftmann@29577: |-> (fn def_thm => PureThy.store_thm (b_def, def_thm) haftmann@29577: #> snd haftmann@29577: #> register_operation class (c', (dict', SOME (Thm.symmetric def_thm)))) haftmann@27690: |> Sign.add_const_constraint (c', SOME ty') haftmann@27690: end; haftmann@27690: haftmann@27690: fun abbrev class prmode pos ((c, mx), rhs) thy = haftmann@27690: let haftmann@29577: val morph = morphism thy class; haftmann@29577: val unchecks = these_unchecks thy [class]; haftmann@29577: val b = Morphism.binding morph (Binding.name c); haftmann@29577: val c' = Sign.full_name thy b; haftmann@27690: val rhs' = Pattern.rewrite_term thy unchecks [] rhs; haftmann@27690: val ty' = Term.fastype_of rhs'; haftmann@29577: val rhs'' = map_types ((*Type.strip_sorts o *)Logic.varifyT) rhs'; haftmann@27690: in haftmann@29577: thy haftmann@29577: |> Sign.add_abbrev (#1 prmode) pos (b, rhs'') haftmann@29577: |> snd haftmann@27690: |> Sign.add_const_constraint (c', SOME ty') haftmann@27690: |> Sign.notation true prmode [(Const (c', ty'), mx)] haftmann@27690: |> not (#1 prmode = PrintMode.input) ? register_operation class (c', (rhs', NONE)) haftmann@27690: end; haftmann@27690: haftmann@27690: haftmann@25462: (** instantiation target **) haftmann@25462: haftmann@25462: (* bookkeeping *) haftmann@25462: haftmann@25462: datatype instantiation = Instantiation of { haftmann@25864: arities: string list * (string * sort) list * sort, haftmann@25462: params: ((string * string) * (string * typ)) list haftmann@25603: (*(instantiation parameter, 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@25462: haftmann@25462: (* syntax *) haftmann@25462: haftmann@26238: fun synchronize_inst_syntax ctxt = haftmann@25462: let haftmann@26259: val Instantiation { arities = (_, _, sort), params = params } = Instantiation.get ctxt; haftmann@26238: val thy = ProofContext.theory_of ctxt; haftmann@26238: fun subst (c, ty) = case AxClass.inst_tyco_of thy (c, ty) haftmann@26238: of SOME tyco => (case AList.lookup (op =) params (c, tyco) haftmann@26238: of SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty) haftmann@26238: | NONE => NONE) haftmann@26238: | NONE => NONE; haftmann@26238: val unchecks = haftmann@26238: map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params; haftmann@26238: in haftmann@26238: ctxt haftmann@26259: |> Overloading.map_improvable_syntax haftmann@26730: (fn (((primary_constraints, _), (((improve, _), _), _)), _) => haftmann@26730: (((primary_constraints, []), (((improve, subst), false), unchecks)), false)) haftmann@26238: 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)) wenzelm@25999: ::: Scan.repeat ((Scan.many1 is_valid >> implode) --| junk)); haftmann@25485: in haftmann@25485: explode #> scan_valids #> implode haftmann@25485: end; haftmann@25485: haftmann@26259: fun type_name "*" = "prod" haftmann@26259: | type_name "+" = "sum" haftmann@29526: | type_name s = sanatize_name (NameSpace.base s); haftmann@26259: haftmann@26518: fun resort_terms pp algebra consts constraints ts = haftmann@26518: let haftmann@26518: fun matchings (Const (c_ty as (c, _))) = (case constraints c haftmann@26518: of NONE => I haftmann@26518: | SOME sorts => fold2 (curry (Sorts.meet_sort algebra)) haftmann@26518: (Consts.typargs consts c_ty) sorts) haftmann@26518: | matchings _ = I haftmann@26518: val tvartab = (fold o fold_aterms) matchings ts Vartab.empty wenzelm@26642: handle Sorts.CLASS_ERROR e => error (Sorts.class_error pp e); haftmann@27089: val inst = map_type_tvar haftmann@27089: (fn (vi, sort) => TVar (vi, the_default sort (Vartab.lookup tvartab vi))); haftmann@26518: in if Vartab.is_empty tvartab then NONE else SOME ((map o map_types) inst ts) end; haftmann@26518: haftmann@25864: fun init_instantiation (tycos, vs, sort) thy = haftmann@25462: let haftmann@25536: val _ = if null tycos then error "At least one arity must be given" else (); haftmann@29969: val params = these_params thy (filter (can (AxClass.get_info thy)) sort); haftmann@29632: fun get_param tyco (param, (_, (c, ty))) = haftmann@29632: if can (AxClass.param_of_inst thy) (c, tyco) haftmann@25603: then NONE else SOME ((c, tyco), haftmann@25864: (param ^ "_" ^ type_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty)); haftmann@26518: val inst_params = map_product get_param tycos params |> map_filter I; haftmann@26518: val primary_constraints = map (apsnd haftmann@26518: (map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) params; wenzelm@26939: val pp = Syntax.pp_global thy; haftmann@26518: val algebra = Sign.classes_of thy haftmann@26518: |> fold (fn tyco => Sorts.add_arities pp haftmann@26518: (tyco, map (fn class => (class, map snd vs)) sort)) tycos; haftmann@26518: val consts = Sign.consts_of thy; haftmann@26518: val improve_constraints = AList.lookup (op =) haftmann@26518: (map (fn (_, (class, (c, _))) => (c, [[class]])) params); haftmann@26518: fun resort_check ts ctxt = case resort_terms pp algebra consts improve_constraints ts haftmann@26518: of NONE => NONE haftmann@26518: | SOME ts' => SOME (ts', ctxt); haftmann@26518: fun improve (c, ty) = case AxClass.inst_tyco_of thy (c, ty) haftmann@26329: of SOME tyco => (case AList.lookup (op =) inst_params (c, tyco) haftmann@29969: of SOME (_, ty') => if Type.typ_instance (Sign.tsig_of thy) (ty', ty) haftmann@29969: then SOME (ty, ty') else NONE haftmann@26259: | NONE => NONE) haftmann@26259: | NONE => NONE; haftmann@25485: in haftmann@25485: thy haftmann@25485: |> ProofContext.init haftmann@26329: |> Instantiation.put (mk_instantiation ((tycos, vs, sort), inst_params)) wenzelm@27281: |> fold (Variable.declare_typ o TFree) vs haftmann@26329: |> fold (Variable.declare_names o Free o snd) inst_params haftmann@26259: |> (Overloading.map_improvable_syntax o apfst) haftmann@29969: (K ((primary_constraints, []), (((improve, K NONE), false), []))) haftmann@26259: |> Overloading.add_improvable_syntax haftmann@26518: |> Context.proof_map (Syntax.add_term_check 0 "resorting" resort_check) haftmann@26238: |> synchronize_inst_syntax haftmann@25485: end; haftmann@25485: haftmann@26238: fun confirm_declaration c = (map_instantiation o apsnd) haftmann@26238: (filter_out (fn (_, (c', _)) => c' = c)) haftmann@26238: #> LocalTheory.target synchronize_inst_syntax haftmann@26238: haftmann@25485: fun gen_instantiation_instance do_proof after_qed lthy = haftmann@25485: let haftmann@25864: val (tycos, vs, sort) = (#arities o the_instantiation) lthy; haftmann@25864: val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, 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@28666: fun prove_instantiation_exit tac = prove_instantiation_instance tac haftmann@28666: #> LocalTheory.exit_global; haftmann@28666: haftmann@28666: fun prove_instantiation_exit_result f tac x lthy = haftmann@28666: let haftmann@29439: val morph = ProofContext.export_morphism lthy haftmann@28666: (ProofContext.init (ProofContext.theory_of lthy)); haftmann@29439: val y = f morph x; haftmann@28666: in haftmann@28666: lthy haftmann@28666: |> prove_instantiation_exit (fn ctxt => tac ctxt y) haftmann@28666: |> pair y haftmann@28666: end; haftmann@28666: haftmann@25462: fun conclude_instantiation lthy = haftmann@25462: let haftmann@25485: val { arities, params } = the_instantiation lthy; haftmann@25864: val (tycos, vs, sort) = arities; haftmann@25462: val thy = ProofContext.theory_of lthy; haftmann@25597: val _ = map (fn tyco => if Sign.of_sort thy haftmann@25864: (Type (tyco, map TFree vs), 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@25603: fun pretty_instantiation lthy = haftmann@25603: let haftmann@25603: val { arities, params } = the_instantiation lthy; haftmann@25864: val (tycos, vs, sort) = arities; haftmann@25603: val thy = ProofContext.theory_of lthy; haftmann@25864: fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort); haftmann@25603: fun pr_param ((c, _), (v, ty)) = haftmann@25864: (Pretty.block o Pretty.breaks) [Pretty.str v, Pretty.str "==", wenzelm@26939: (Pretty.str o Sign.extern_const thy) c, Pretty.str "::", Syntax.pretty_typ_global thy ty]; haftmann@25603: in haftmann@25603: (Pretty.block o Pretty.fbreaks) haftmann@25603: (Pretty.str "instantiation" :: map pr_arity tycos @ map pr_param params) haftmann@25603: end; haftmann@25603: haftmann@29526: haftmann@29526: (** tactics and methods **) haftmann@29526: haftmann@29526: fun intro_classes_tac facts st = haftmann@29526: let haftmann@29526: val thy = Thm.theory_of_thm st; haftmann@29526: val classes = Sign.all_classes thy; haftmann@29526: val class_trivs = map (Thm.class_triv thy) classes; haftmann@29526: val class_intros = map_filter (try (#intro o AxClass.get_info thy)) classes; haftmann@29526: val assm_intros = all_assm_intros thy; haftmann@29526: in haftmann@29526: Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st haftmann@29526: end; haftmann@29526: haftmann@29526: fun default_intro_tac ctxt [] = haftmann@29577: intro_classes_tac [] ORELSE Locale.intro_locales_tac true ctxt [] haftmann@29526: | default_intro_tac _ _ = no_tac; haftmann@29526: haftmann@29526: fun default_tac rules ctxt facts = haftmann@29526: HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE haftmann@29526: default_intro_tac ctxt facts; haftmann@29526: haftmann@29526: val _ = Context.>> (Context.map_theory haftmann@29526: (Method.add_methods haftmann@29526: [("intro_classes", Method.no_args (Method.METHOD intro_classes_tac), haftmann@29526: "back-chain introduction rules of classes"), haftmann@29526: ("default", Method.thms_ctxt_args (Method.METHOD oo default_tac), haftmann@29526: "apply some intro/elim rule")])); haftmann@29526: haftmann@24218: end; haftmann@25683: