wenzelm@12014: (* Title: Pure/Isar/locale.ML wenzelm@11896: ID: $Id$ ballarin@15206: Author: Clemens Ballarin, TU Muenchen; Markus Wenzel, LMU/TU Muenchen wenzelm@11896: wenzelm@12058: Locales -- Isar proof contexts as meta-level predicates, with local wenzelm@12529: syntax and implicit structures. wenzelm@12529: ballarin@19942: Draws basic ideas from Florian Kammueller's original version of wenzelm@12529: locales, but uses the richer infrastructure of Isar instead of the raw ballarin@20035: meta-logic. Furthermore, structured import of contexts (with merge ballarin@20035: and rename operations) are provided, as well as type-inference of the wenzelm@13375: signature parts, and predicate definitions of the specification text. ballarin@14446: ballarin@17437: Interpretation enables the reuse of theorems of locales in other ballarin@17437: contexts, namely those defined by theories, structured proofs and ballarin@17437: locales themselves. ballarin@17437: ballarin@14446: See also: ballarin@14446: ballarin@14446: [1] Clemens Ballarin. Locales and Locale Expressions in Isabelle/Isar. ballarin@14446: In Stefano Berardi et al., Types for Proofs and Programs: International ballarin@15099: Workshop, TYPES 2003, Torino, Italy, LNCS 3085, pages 34-50, 2004. ballarin@19942: [2] Clemens Ballarin. Interpretation of Locales in Isabelle: Managing ballarin@19942: Dependencies between Locales. Technical Report TUM-I0607, Technische ballarin@19942: Universitaet Muenchen, 2006. ballarin@19942: [3] Clemens Ballarin. Interpretation of Locales in Isabelle: Theories and ballarin@19942: Proof Contexts. In J.M. Borwein and W.M. Farmer, MKM 2006, LNAI 4108, ballarin@19942: pages 31-43, 2006. wenzelm@11896: *) wenzelm@11896: ballarin@16169: (* TODO: ballarin@16169: - beta-eta normalisation of interpretation parameters ballarin@16169: - dangling type frees in locales ballarin@16620: - test subsumption of interpretations when merging theories ballarin@16169: *) ballarin@16169: wenzelm@11896: signature LOCALE = wenzelm@11896: sig wenzelm@12273: datatype expr = wenzelm@12273: Locale of string | ballarin@16102: Rename of expr * (string * mixfix option) option list | wenzelm@12273: Merge of expr list wenzelm@12273: val empty: expr wenzelm@18137: datatype 'a element = Elem of 'a | Expr of expr wenzelm@21035: val map_elem: ('a -> 'b) -> 'a element -> 'b element ballarin@15206: wenzelm@16458: val intern: theory -> xstring -> string wenzelm@16458: val extern: theory -> string -> xstring ballarin@19991: val init: string -> theory -> Proof.context ballarin@15206: ballarin@18795: (* The specification of a locale *) ballarin@18795: val parameters_of: theory -> string -> haftmann@18917: ((string * typ) * mixfix) list ballarin@19276: val parameters_of_expr: theory -> expr -> ballarin@19276: ((string * typ) * mixfix) list ballarin@18795: val local_asms_of: theory -> string -> ballarin@18795: ((string * Attrib.src list) * term list) list ballarin@18795: val global_asms_of: theory -> string -> ballarin@18795: ((string * Attrib.src list) * term list) list haftmann@22523: val intros: theory -> string -> haftmann@22523: thm list * thm list ballarin@18795: wenzelm@18899: (* Processing of locale statements *) wenzelm@18137: val read_context_statement: xstring option -> Element.context element list -> wenzelm@19585: (string * string list) list list -> Proof.context -> ballarin@19991: string option * Proof.context * Proof.context * (term * term list) list list wenzelm@21035: val read_context_statement_i: string option -> Element.context element list -> wenzelm@21035: (string * string list) list list -> Proof.context -> wenzelm@21035: string option * Proof.context * Proof.context * (term * term list) list list wenzelm@18137: val cert_context_statement: string option -> Element.context_i element list -> wenzelm@19585: (term * term list) list list -> Proof.context -> ballarin@19991: string option * Proof.context * Proof.context * (term * term list) list list wenzelm@19780: val read_expr: expr -> Element.context list -> Proof.context -> wenzelm@19780: Element.context_i list * Proof.context wenzelm@19780: val cert_expr: expr -> Element.context_i list -> Proof.context -> wenzelm@19780: Element.context_i list * Proof.context ballarin@15596: ballarin@15596: (* Diagnostic functions *) wenzelm@12758: val print_locales: theory -> unit wenzelm@18137: val print_locale: theory -> bool -> expr -> Element.context list -> unit ballarin@17138: val print_global_registrations: bool -> string -> theory -> unit wenzelm@18617: val print_local_registrations': bool -> string -> Proof.context -> unit wenzelm@18617: val print_local_registrations: bool -> string -> Proof.context -> unit wenzelm@18137: haftmann@22351: val add_locale: string option -> bstring -> expr -> Element.context list -> theory wenzelm@20965: -> string * Proof.context haftmann@22351: val add_locale_i: string option -> bstring -> expr -> Element.context_i list -> theory wenzelm@20965: -> string * Proof.context ballarin@15596: schirmer@21225: (* Tactics *) schirmer@21225: val intro_locales_tac: bool -> Proof.context -> thm list -> tactic schirmer@21225: ballarin@15696: (* Storing results *) wenzelm@18806: val add_thmss: string -> string -> wenzelm@18806: ((string * Attrib.src list) * (thm list * Attrib.src list) list) list -> wenzelm@21582: Proof.context -> Proof.context wenzelm@21665: val add_type_syntax: string -> (morphism -> Context.generic -> Context.generic) -> wenzelm@21499: Proof.context -> Proof.context wenzelm@21665: val add_term_syntax: string -> (morphism -> Context.generic -> Context.generic) -> wenzelm@21499: Proof.context -> Proof.context wenzelm@21665: val add_declaration: string -> (morphism -> Context.generic -> Context.generic) -> wenzelm@21499: Proof.context -> Proof.context wenzelm@18137: haftmann@22300: val interpretation_i: (Proof.context -> Proof.context) -> ballarin@22658: (bool * string) * Attrib.src list -> expr -> ballarin@22658: (typ Symtab.table * term Symtab.table) * (term * term) list -> haftmann@22300: theory -> Proof.state wenzelm@21005: val interpretation: (Proof.context -> Proof.context) -> ballarin@22658: (bool * string) * Attrib.src list -> expr -> ballarin@22658: string option list * (string * string) list -> wenzelm@21005: theory -> Proof.state wenzelm@21005: val interpretation_in_locale: (Proof.context -> Proof.context) -> wenzelm@21005: xstring * expr -> theory -> Proof.state haftmann@22300: val interpret_i: (Proof.state -> Proof.state Seq.seq) -> ballarin@22658: (bool * string) * Attrib.src list -> expr -> ballarin@22658: (typ Symtab.table * term Symtab.table) * (term * term) list -> haftmann@22300: bool -> Proof.state -> Proof.state wenzelm@21005: val interpret: (Proof.state -> Proof.state Seq.seq) -> ballarin@22658: (bool * string) * Attrib.src list -> expr -> ballarin@22658: string option list * (string * string) list -> wenzelm@21005: bool -> Proof.state -> Proof.state wenzelm@11896: end; wenzelm@12839: wenzelm@12289: structure Locale: LOCALE = wenzelm@11896: struct wenzelm@11896: wenzelm@23228: (* legacy operations *) wenzelm@23228: wenzelm@23228: fun gen_merge_lists _ xs [] = xs wenzelm@23228: | gen_merge_lists _ [] ys = ys wenzelm@23228: | gen_merge_lists eq xs ys = xs @ filter_out (member eq xs) ys; wenzelm@23228: wenzelm@23228: fun merge_lists xs ys = gen_merge_lists (op =) xs ys; wenzelm@23228: fun merge_alists xs = gen_merge_lists (eq_fst (op =)) xs; wenzelm@23228: ballarin@19931: fun legacy_unvarifyT thm = ballarin@19931: let ballarin@19931: val cT = Thm.ctyp_of (Thm.theory_of_thm thm); wenzelm@22700: val tvars = rev (Thm.fold_terms Term.add_tvars thm []); wenzelm@20307: val tfrees = map (fn ((x, _), S) => SOME (cT (TFree (x, S)))) tvars; ballarin@19931: in Drule.instantiate' tfrees [] thm end; ballarin@19931: ballarin@19931: fun legacy_unvarify raw_thm = ballarin@19931: let ballarin@19931: val thm = legacy_unvarifyT raw_thm; ballarin@19931: val ct = Thm.cterm_of (Thm.theory_of_thm thm); wenzelm@22700: val vars = rev (Thm.fold_terms Term.add_vars thm []); wenzelm@20307: val frees = map (fn ((x, _), T) => SOME (ct (Free (x, T)))) vars; ballarin@19931: in Drule.instantiate' [] frees thm end; wenzelm@19780: wenzelm@20307: wenzelm@12273: (** locale elements and expressions **) wenzelm@11896: wenzelm@18137: datatype ctxt = datatype Element.ctxt; wenzelm@17355: wenzelm@12273: datatype expr = wenzelm@12273: Locale of string | ballarin@16102: Rename of expr * (string * mixfix option) option list | wenzelm@12273: Merge of expr list; wenzelm@11896: wenzelm@12273: val empty = Merge []; wenzelm@12273: wenzelm@18137: datatype 'a element = ballarin@15206: Elem of 'a | Expr of expr; wenzelm@12273: wenzelm@18137: fun map_elem f (Elem e) = Elem (f e) wenzelm@18137: | map_elem _ (Expr e) = Expr e; wenzelm@18137: wenzelm@21665: type decl = (morphism -> Context.generic -> Context.generic) * stamp; wenzelm@21499: wenzelm@12070: type locale = ballarin@19931: {axiom: Element.witness list, ballarin@19942: (* For locales that define predicates this is [A [A]], where A is the locale ballarin@20317: specification. Otherwise []. ballarin@20317: Only required to generate the right witnesses for locales with predicates. *) wenzelm@22573: imports: expr, (*dynamic imports*) ballarin@19783: elems: (Element.context_i * stamp) list, ballarin@19783: (* Static content, neither Fixes nor Constrains elements *) wenzelm@21499: params: ((string * typ) * mixfix) list, (*all params*) ballarin@19931: lparams: string list, (*local params*) wenzelm@21665: decls: decl list * decl list, (*type/term_syntax declarations*) ballarin@19931: regs: ((string * string list) * Element.witness list) list, wenzelm@19780: (* Registrations: indentifiers and witnesses of locales interpreted in the locale. *) ballarin@19931: intros: thm list * thm list} ballarin@19931: (* Introduction rules: of delta predicate and locale predicate. *) wenzelm@11896: wenzelm@15703: (* CB: an internal (Int) locale element was either imported or included, wenzelm@15703: an external (Ext) element appears directly in the locale text. *) wenzelm@15703: wenzelm@15703: datatype ('a, 'b) int_ext = Int of 'a | Ext of 'b; wenzelm@15703: wenzelm@15703: wenzelm@15703: ballarin@16736: (** management of registrations in theories and proof contexts **) wenzelm@11896: ballarin@15837: structure Registrations : ballarin@15837: sig ballarin@15837: type T ballarin@15837: val empty: T ballarin@15837: val join: T * T -> T ballarin@20069: val dest: theory -> T -> ballarin@22658: (term list * ballarin@22658: (((bool * string) * Attrib.src list) * Element.witness list * ballarin@22658: Element.witness Termtab.table)) list wenzelm@16458: val lookup: theory -> T * term list -> ballarin@22658: (((bool * string) * Attrib.src list) * Element.witness list * ballarin@22658: Element.witness Termtab.table) option haftmann@22351: val insert: theory -> term list * ((bool * string) * Attrib.src list) -> T -> haftmann@22351: T * (term list * (((bool * string) * Attrib.src list) * Element.witness list)) list wenzelm@19780: val add_witness: term list -> Element.witness -> T -> T ballarin@22658: val add_equation: term list -> Element.witness -> T -> T ballarin@15837: end = ballarin@15837: struct ballarin@22658: (* A registration is indexed by parameter instantiation. Its components are: ballarin@22658: - parameters and prefix ballarin@22658: (if the Boolean flag is set, only accesses containing the prefix are generated, wenzelm@23418: otherwise the prefix may be omitted when accessing theorems etc.) ballarin@22658: - theorems (actually witnesses) instantiating locale assumptions ballarin@22658: - theorems (equations, actually witnesses) interpreting derived concepts ballarin@22658: and indexed by lhs ballarin@22658: *) ballarin@22658: type T = (((bool * string) * Attrib.src list) * Element.witness list * ballarin@22658: Element.witness Termtab.table) Termtab.table; ballarin@15837: ballarin@15837: val empty = Termtab.empty; ballarin@15837: ballarin@15837: (* term list represented as single term, for simultaneous matching *) ballarin@15837: fun termify ts = wenzelm@18343: Term.list_comb (Const ("", map fastype_of ts ---> propT), ts); ballarin@15837: fun untermify t = ballarin@15837: let fun ut (Const _) ts = ts ballarin@15837: | ut (s $ t) ts = ut s (t::ts) ballarin@15837: in ut t [] end; ballarin@15837: ballarin@22658: (* joining of registrations: ballarin@22658: - prefix and attributes of right theory; ballarin@22658: - witnesses are equal, no attempt to subsumption testing; ballarin@22658: - union of equalities, if conflicting (i.e. two eqns with equal lhs) ballarin@22658: eqn of right theory takes precedence *) ballarin@22658: fun join (r1, r2) = Termtab.join (fn _ => fn ((_, _, e1), (n, w, e2)) => ballarin@22658: (n, w, Termtab.join (fn _ => fn (_, e) => e) (e1, e2))) (r1, r2); ballarin@15837: ballarin@20069: fun dest_transfer thy regs = ballarin@22658: Termtab.dest regs |> map (apsnd (fn (n, ws, es) => ballarin@22658: (n, map (Element.transfer_witness thy) ws, Termtab.map (Element.transfer_witness thy) es))); ballarin@20069: ballarin@20069: fun dest thy regs = dest_transfer thy regs |> map (apfst untermify); ballarin@15837: ballarin@15837: (* registrations that subsume t *) wenzelm@17203: fun subsumers thy t regs = ballarin@20069: filter (fn (t', _) => Pattern.matches thy (t', t)) (dest_transfer thy regs); ballarin@15837: ballarin@15837: (* look up registration, pick one that subsumes the query *) wenzelm@19780: fun lookup thy (regs, ts) = ballarin@15837: let ballarin@15837: val t = termify ts; wenzelm@19780: val subs = subsumers thy t regs; wenzelm@21483: in wenzelm@21483: (case subs of ballarin@15837: [] => NONE ballarin@22658: | ((t', (attn, thms, eqns)) :: _) => wenzelm@21483: let wenzelm@19780: val (tinst, inst) = Pattern.match thy (t', t) (Vartab.empty, Vartab.empty); ballarin@15837: (* thms contain Frees, not Vars *) wenzelm@21665: val tinst' = tinst |> Vartab.dest (* FIXME Vartab.map (!?) *) wenzelm@19810: |> map (fn ((x, 0), (_, T)) => (x, Logic.legacy_unvarifyT T)) ballarin@15837: |> Symtab.make; ballarin@15837: val inst' = inst |> Vartab.dest wenzelm@19810: |> map (fn ((x, 0), (_, t)) => (x, Logic.legacy_unvarify t)) ballarin@15837: |> Symtab.make; wenzelm@21483: val inst_witness = Element.morph_witness (Element.inst_morphism thy (tinst', inst')); ballarin@22658: in SOME (attn, map inst_witness thms, Termtab.map inst_witness eqns) end) ballarin@15837: end; ballarin@15837: ballarin@15837: (* add registration if not subsumed by ones already present, ballarin@15837: additionally returns registrations that are strictly subsumed *) ballarin@20069: fun insert thy (ts, attn) regs = ballarin@15837: let ballarin@15837: val t = termify ts; ballarin@20069: val subs = subsumers thy t regs ; ballarin@15837: in (case subs of ballarin@15837: [] => let ballarin@15837: val sups = ballarin@20069: filter (fn (t', _) => Pattern.matches thy (t, t')) (dest_transfer thy regs); ballarin@22658: val sups' = map (apfst untermify) sups |> map (fn (ts, (n, w, d)) => (ts, (n, w))) ballarin@22658: in (Termtab.update (t, (attn, [], Termtab.empty)) regs, sups') end ballarin@15837: | _ => (regs, [])) ballarin@15837: end; ballarin@15837: ballarin@15837: (* add witness theorem to registration, ballarin@16169: only if instantiation is exact, otherwise exception Option raised *) ballarin@15837: fun add_witness ts thm regs = ballarin@15837: let ballarin@15837: val t = termify ts; ballarin@22658: val (x, thms, eqns) = the (Termtab.lookup regs t); ballarin@15837: in ballarin@22658: Termtab.update (t, (x, thm::thms, eqns)) regs ballarin@15837: end; ballarin@22658: ballarin@22658: (* add equation to registration, replaces previous equation with same lhs; ballarin@22658: only if instantiation is exact, otherwise exception Option raised; ballarin@22658: exception TERM raised if not a meta equality *) ballarin@22658: fun add_equation ts thm regs = ballarin@22658: let ballarin@22658: val t = termify ts; ballarin@22658: val (x, thms, eqns) = the (Termtab.lookup regs t); ballarin@22658: in ballarin@22658: Termtab.update (t, (x, thms, Termtab.update (thm |> Element.witness_prop |> Logic.dest_equals |> fst, thm) eqns)) regs ballarin@22658: end; ballarin@22658: (* TODO: avoid code clone *) ballarin@15837: end; ballarin@15837: ballarin@16736: ballarin@15837: (** theory data **) ballarin@15596: wenzelm@16458: structure GlobalLocalesData = TheoryDataFun wenzelm@22846: ( ballarin@15837: type T = NameSpace.T * locale Symtab.table * Registrations.T Symtab.table; ballarin@15596: (* 1st entry: locale namespace, ballarin@15596: 2nd entry: locales of the theory, ballarin@15837: 3rd entry: registrations, indexed by locale name *) wenzelm@11896: ballarin@15596: val empty = (NameSpace.empty, Symtab.empty, Symtab.empty); wenzelm@12063: val copy = I; wenzelm@16458: val extend = I; wenzelm@12289: wenzelm@21665: fun join_locales _ wenzelm@22573: ({axiom, imports, elems, params, lparams, decls = (decls1, decls2), regs, intros}: locale, wenzelm@21665: {elems = elems', decls = (decls1', decls2'), regs = regs', ...}: locale) = ballarin@19931: {axiom = axiom, wenzelm@22573: imports = imports, haftmann@17496: elems = gen_merge_lists (eq_snd (op =)) elems elems', ballarin@16736: params = params, ballarin@19278: lparams = lparams, wenzelm@21499: decls = wenzelm@21665: (Library.merge (eq_snd (op =)) (decls1, decls1'), wenzelm@21665: Library.merge (eq_snd (op =)) (decls2, decls2')), ballarin@19931: regs = merge_alists regs regs', ballarin@19931: intros = intros}; wenzelm@16458: fun merge _ ((space1, locs1, regs1), (space2, locs2, regs2)) = wenzelm@21499: (NameSpace.merge (space1, space2), Symtab.join join_locales (locs1, locs2), wenzelm@19025: Symtab.join (K Registrations.join) (regs1, regs2)); wenzelm@22846: ); wenzelm@15801: wenzelm@15801: ballarin@15624: ballarin@15624: (** context data **) wenzelm@11896: wenzelm@16458: structure LocalLocalesData = ProofDataFun wenzelm@22846: ( wenzelm@22846: type T = Registrations.T Symtab.table; (*registrations, indexed by locale name*) ballarin@15624: fun init _ = Symtab.empty; wenzelm@22846: ); wenzelm@12289: wenzelm@12277: wenzelm@12277: (* access locales *) wenzelm@12277: wenzelm@22846: fun print_locales thy = wenzelm@22846: let val (space, locs, _) = GlobalLocalesData.get thy in wenzelm@22846: Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (space, locs))) wenzelm@22846: |> Pretty.writeln wenzelm@22846: end; ballarin@15624: wenzelm@16458: val intern = NameSpace.intern o #1 o GlobalLocalesData.get; wenzelm@16458: val extern = NameSpace.extern o #1 o GlobalLocalesData.get; ballarin@15624: wenzelm@16144: fun declare_locale name thy = wenzelm@16144: thy |> GlobalLocalesData.map (fn (space, locs, regs) => wenzelm@16458: (Sign.declare_name thy name space, locs, regs)); wenzelm@11896: ballarin@15596: fun put_locale name loc = ballarin@15624: GlobalLocalesData.map (fn (space, locs, regs) => wenzelm@17412: (space, Symtab.update (name, loc) locs, regs)); wenzelm@17412: wenzelm@17412: fun get_locale thy name = Symtab.lookup (#2 (GlobalLocalesData.get thy)) name; wenzelm@11896: wenzelm@12014: fun the_locale thy name = wenzelm@12014: (case get_locale thy name of skalberg@15531: SOME loc => loc skalberg@15531: | NONE => error ("Unknown locale " ^ quote name)); wenzelm@11896: wenzelm@18806: fun change_locale name f thy = wenzelm@18806: let wenzelm@22573: val {axiom, imports, elems, params, lparams, decls, regs, intros} = ballarin@19931: the_locale thy name; wenzelm@22573: val (axiom', imports', elems', params', lparams', decls', regs', intros') = wenzelm@22573: f (axiom, imports, elems, params, lparams, decls, regs, intros); wenzelm@18806: in wenzelm@22573: put_locale name {axiom = axiom', imports = imports', elems = elems', wenzelm@21499: params = params', lparams = lparams', decls = decls', regs = regs', ballarin@19931: intros = intros'} thy wenzelm@18806: end; wenzelm@18806: wenzelm@12046: ballarin@15596: (* access registrations *) ballarin@15596: ballarin@15696: (* Ids of global registrations are varified, ballarin@15696: Ids of local registrations are not. ballarin@22658: Witness thms of registrations are never varified. ballarin@22658: Varification of eqns as varification of ids. *) ballarin@15696: ballarin@15624: (* retrieve registration from theory or context *) ballarin@15624: ballarin@20069: fun gen_get_registrations get thy_of thy_ctxt name = wenzelm@17412: case Symtab.lookup (get thy_ctxt) name of ballarin@15696: NONE => [] ballarin@20069: | SOME reg => Registrations.dest (thy_of thy_ctxt) reg; ballarin@15696: ballarin@15696: val get_global_registrations = ballarin@20069: gen_get_registrations (#3 o GlobalLocalesData.get) I; ballarin@15696: val get_local_registrations = ballarin@20069: gen_get_registrations LocalLocalesData.get ProofContext.theory_of; ballarin@15696: wenzelm@16458: fun gen_get_registration get thy_of thy_ctxt (name, ps) = wenzelm@17412: case Symtab.lookup (get thy_ctxt) name of ballarin@15624: NONE => NONE wenzelm@16458: | SOME reg => Registrations.lookup (thy_of thy_ctxt) (reg, ps); ballarin@15624: ballarin@15624: val get_global_registration = wenzelm@16458: gen_get_registration (#3 o GlobalLocalesData.get) I; ballarin@15624: val get_local_registration = wenzelm@16458: gen_get_registration LocalLocalesData.get ProofContext.theory_of; ballarin@15596: wenzelm@18343: val test_global_registration = is_some oo get_global_registration; wenzelm@18343: val test_local_registration = is_some oo get_local_registration; ballarin@15624: fun smart_test_registration ctxt id = ballarin@15624: let ballarin@15624: val thy = ProofContext.theory_of ctxt; ballarin@15624: in ballarin@15624: test_global_registration thy id orelse test_local_registration ctxt id ballarin@15624: end; ballarin@15624: ballarin@15624: ballarin@15837: (* add registration to theory or context, ignored if subsumed *) ballarin@15624: wenzelm@16458: fun gen_put_registration map_data thy_of (name, ps) attn thy_ctxt = ballarin@15837: map_data (fn regs => ballarin@15837: let wenzelm@16458: val thy = thy_of thy_ctxt; wenzelm@18343: val reg = the_default Registrations.empty (Symtab.lookup regs name); wenzelm@16458: val (reg', sups) = Registrations.insert thy (ps, attn) reg; ballarin@15837: val _ = if not (null sups) then warning ballarin@15837: ("Subsumed interpretation(s) of locale " ^ wenzelm@16458: quote (extern thy name) ^ ballarin@22658: "\nwith the following prefix(es):" ^ haftmann@22351: commas_quote (map (fn (_, (((_, s), _), _)) => s) sups)) ballarin@15837: else (); wenzelm@17412: in Symtab.update (name, reg') regs end) thy_ctxt; ballarin@15624: ballarin@15624: val put_global_registration = ballarin@15624: gen_put_registration (fn f => wenzelm@16458: GlobalLocalesData.map (fn (space, locs, regs) => (space, locs, f regs))) I; ballarin@15837: val put_local_registration = wenzelm@16458: gen_put_registration LocalLocalesData.map ProofContext.theory_of; ballarin@15596: wenzelm@18806: fun put_registration_in_locale name id = wenzelm@22573: change_locale name (fn (axiom, imports, elems, params, lparams, decls, regs, intros) => wenzelm@22573: (axiom, imports, elems, params, lparams, decls, regs @ [(id, [])], intros)); ballarin@17000: ballarin@15624: ballarin@22658: (* add witness theorem to registration, ignored if registration not present *) ballarin@15596: wenzelm@18123: fun add_witness (name, ps) thm = wenzelm@18123: Symtab.map_entry name (Registrations.add_witness ps thm); wenzelm@18123: wenzelm@18123: fun add_global_witness id thm = wenzelm@18123: GlobalLocalesData.map (fn (space, locs, regs) => (space, locs, add_witness id thm regs)); wenzelm@18123: wenzelm@18123: val add_local_witness = LocalLocalesData.map oo add_witness; ballarin@15596: wenzelm@18806: fun add_witness_in_locale name id thm = wenzelm@22573: change_locale name (fn (axiom, imports, elems, params, lparams, decls, regs, intros) => ballarin@17000: let ballarin@17000: fun add (id', thms) = wenzelm@18806: if id = id' then (id', thm :: thms) else (id', thms); wenzelm@22573: in (axiom, imports, elems, params, lparams, decls, map add regs, intros) end); ballarin@15596: ballarin@14215: ballarin@22658: (* add equation to registration, ignored if registration not present *) ballarin@22658: ballarin@22658: fun add_equation (name, ps) thm = ballarin@22658: Symtab.map_entry name (Registrations.add_equation ps thm); ballarin@22658: ballarin@22658: fun add_global_equation id thm = ballarin@22658: GlobalLocalesData.map (fn (space, locs, regs) => (space, locs, add_equation id thm regs)); ballarin@22658: ballarin@22658: val add_local_equation = LocalLocalesData.map oo add_equation; ballarin@22658: ballarin@22658: ballarin@15624: (* printing of registrations *) ballarin@15596: ballarin@17138: fun gen_print_registrations get_regs mk_ctxt msg show_wits loc thy_ctxt = ballarin@15596: let wenzelm@15703: val ctxt = mk_ctxt thy_ctxt; wenzelm@15703: val thy = ProofContext.theory_of ctxt; ballarin@22658: (* TODO: nice effect of show_wits on equations. *) wenzelm@15703: val prt_term = Pretty.quote o ProofContext.pretty_term ctxt; ballarin@17096: fun prt_inst ts = ballarin@17096: Pretty.enclose "(" ")" (Pretty.breaks (map prt_term ts)); ballarin@22658: fun prt_attn ((false, prfx), atts) = ballarin@22658: Pretty.breaks (Pretty.str prfx :: Pretty.str "(optional)" :: ballarin@22658: Attrib.pretty_attribs ctxt atts) ballarin@22658: | prt_attn ((true, prfx), atts) = ballarin@22658: Pretty.breaks (Pretty.str prfx :: Attrib.pretty_attribs ctxt atts); ballarin@22658: fun prt_eqns [] = Pretty.str "no equations." ballarin@22658: | prt_eqns eqns = Pretty.block (Pretty.str "equations:" :: Pretty.brk 1 :: ballarin@22658: Pretty.breaks (map (Element.pretty_witness ctxt) eqns)); ballarin@22658: fun prt_core ts eqns = ballarin@22658: [prt_inst ts, Pretty.fbrk, prt_eqns (Termtab.dest eqns |> map snd)]; ballarin@22658: fun prt_witns [] = Pretty.str "no witnesses." ballarin@22658: | prt_witns witns = Pretty.block (Pretty.str "witnesses:" :: Pretty.brk 1 :: ballarin@22658: Pretty.breaks (map (Element.pretty_witness ctxt) witns)) ballarin@22658: fun prt_reg (ts, (((_, ""), []), witns, eqns)) = ballarin@17138: if show_wits ballarin@22658: then Pretty.block (prt_core ts eqns @ [Pretty.fbrk, prt_witns witns]) ballarin@22658: else Pretty.block (prt_core ts eqns) ballarin@22658: | prt_reg (ts, (attn, witns, eqns)) = ballarin@17138: if show_wits ballarin@22658: then Pretty.block ((prt_attn attn @ [Pretty.str ":", Pretty.brk 1] @ ballarin@22658: prt_core ts eqns @ [Pretty.fbrk, prt_witns witns])) ballarin@17096: else Pretty.block ((prt_attn attn @ ballarin@22658: [Pretty.str ":", Pretty.brk 1] @ prt_core ts eqns)); wenzelm@15703: wenzelm@16458: val loc_int = intern thy loc; ballarin@15624: val regs = get_regs thy_ctxt; wenzelm@17412: val loc_regs = Symtab.lookup regs loc_int; ballarin@15596: in ballarin@15596: (case loc_regs of wenzelm@17355: NONE => Pretty.str ("no interpretations in " ^ msg) ballarin@15763: | SOME r => let ballarin@20069: val r' = Registrations.dest thy r; ballarin@22658: val r'' = Library.sort_wrt (fn (_, (((_, prfx), _), _, _)) => prfx) r'; wenzelm@17355: in Pretty.big_list ("interpretations in " ^ msg ^ ":") ballarin@17096: (map prt_reg r'') ballarin@15763: end) ballarin@15596: |> Pretty.writeln ballarin@15596: end; ballarin@15596: ballarin@15624: val print_global_registrations = ballarin@15624: gen_print_registrations (#3 o GlobalLocalesData.get) wenzelm@15703: ProofContext.init "theory"; ballarin@15624: val print_local_registrations' = ballarin@15624: gen_print_registrations LocalLocalesData.get wenzelm@15703: I "context"; ballarin@17138: fun print_local_registrations show_wits loc ctxt = ballarin@17138: (print_global_registrations show_wits loc (ProofContext.theory_of ctxt); ballarin@17138: print_local_registrations' show_wits loc ctxt); ballarin@15624: ballarin@15596: wenzelm@12277: (* diagnostics *) wenzelm@12273: wenzelm@12277: fun err_in_locale ctxt msg ids = wenzelm@12277: let wenzelm@16458: val thy = ProofContext.theory_of ctxt; wenzelm@12529: fun prt_id (name, parms) = wenzelm@16458: [Pretty.block (Pretty.breaks (map Pretty.str (extern thy name :: parms)))]; wenzelm@19482: val prt_ids = flat (separate [Pretty.str " +", Pretty.brk 1] (map prt_id ids)); wenzelm@12502: val err_msg = wenzelm@12529: if forall (equal "" o #1) ids then msg wenzelm@12502: else msg ^ "\n" ^ Pretty.string_of (Pretty.block wenzelm@12502: (Pretty.str "The error(s) above occurred in locale:" :: Pretty.brk 1 :: prt_ids)); wenzelm@18678: in error err_msg end; wenzelm@12063: ballarin@15206: fun err_in_locale' ctxt msg ids' = err_in_locale ctxt msg (map fst ids'); wenzelm@12277: wenzelm@12277: ballarin@19783: fun pretty_ren NONE = Pretty.str "_" ballarin@19783: | pretty_ren (SOME (x, NONE)) = Pretty.str x ballarin@19783: | pretty_ren (SOME (x, SOME syn)) = ballarin@19783: Pretty.block [Pretty.str x, Pretty.brk 1, Syntax.pretty_mixfix syn]; ballarin@19783: ballarin@19783: fun pretty_expr thy (Locale name) = Pretty.str (extern thy name) ballarin@19783: | pretty_expr thy (Rename (expr, xs)) = ballarin@19783: Pretty.block [pretty_expr thy expr, Pretty.brk 1, Pretty.block (map pretty_ren xs |> Pretty.breaks)] ballarin@19783: | pretty_expr thy (Merge es) = ballarin@19783: Pretty.separate "+" (map (pretty_expr thy) es) |> Pretty.block; ballarin@19783: ballarin@19783: fun err_in_expr _ msg (Merge []) = error msg ballarin@19783: | err_in_expr ctxt msg expr = ballarin@19783: error (msg ^ "\n" ^ Pretty.string_of (Pretty.block ballarin@19783: [Pretty.str "The error(s) above occured in locale expression:", Pretty.brk 1, ballarin@19783: pretty_expr (ProofContext.theory_of ctxt) expr])); ballarin@19783: wenzelm@12046: wenzelm@12529: (** structured contexts: rename + merge + implicit type instantiation **) wenzelm@12529: wenzelm@12529: (* parameter types *) wenzelm@12529: wenzelm@12529: fun frozen_tvars ctxt Ts = wenzelm@19914: #1 (Variable.importT_inst (map Logic.mk_type Ts) ctxt) wenzelm@19900: |> map (fn ((xi, S), T) => (xi, (S, T))); wenzelm@12529: wenzelm@12529: fun unify_frozen ctxt maxidx Ts Us = wenzelm@12529: let wenzelm@18343: fun paramify NONE i = (NONE, i) wenzelm@18343: | paramify (SOME T) i = apfst SOME (TypeInfer.paramify_dummies T i); wenzelm@12529: wenzelm@18343: val (Ts', maxidx') = fold_map paramify Ts maxidx; wenzelm@18343: val (Us', maxidx'') = fold_map paramify Us maxidx'; wenzelm@16947: val thy = ProofContext.theory_of ctxt; ballarin@14215: wenzelm@18190: fun unify (SOME T, SOME U) env = (Sign.typ_unify thy (U, T) env wenzelm@18190: handle Type.TUNIFY => raise TYPE ("unify_frozen: failed to unify types", [U, T], [])) wenzelm@18190: | unify _ env = env; wenzelm@18190: val (unifier, _) = fold unify (Ts' ~~ Us') (Vartab.empty, maxidx''); skalberg@15570: val Vs = map (Option.map (Envir.norm_type unifier)) Us'; wenzelm@19482: val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (map_filter I Vs)); skalberg@15570: in map (Option.map (Envir.norm_type unifier')) Vs end; wenzelm@12529: wenzelm@21665: fun params_of elemss = wenzelm@21665: distinct (eq_fst (op = : string * string -> bool)) (maps (snd o fst) elemss); wenzelm@21665: wenzelm@21665: fun params_of' elemss = wenzelm@21665: distinct (eq_fst (op = : string * string -> bool)) (maps (snd o fst o fst) elemss); wenzelm@21665: wenzelm@21665: wenzelm@21665: fun params_prefix params = space_implode "_" params; ballarin@16102: ballarin@14508: ballarin@14508: (* CB: param_types has the following type: skalberg@15531: ('a * 'b option) list -> ('a * 'b) list *) wenzelm@19482: fun param_types ps = map_filter (fn (_, NONE) => NONE | (x, SOME T) => SOME (x, T)) ps; wenzelm@12529: wenzelm@12529: haftmann@19932: fun merge_syntax ctxt ids ss = Symtab.merge (op = : mixfix * mixfix -> bool) ss wenzelm@23655: handle Symtab.DUP x => err_in_locale ctxt wenzelm@23655: ("Conflicting syntax for parameter: " ^ quote x) (map fst ids); ballarin@16102: ballarin@16102: ballarin@17000: (* Distinction of assumed vs. derived identifiers. ballarin@17000: The former may have axioms relating assumptions of the context to ballarin@17000: assumptions of the specification fragment (for locales with wenzelm@19780: predicates). The latter have witnesses relating assumptions of the ballarin@17000: specification fragment to assumptions of other (assumed) specification ballarin@17000: fragments. *) ballarin@17000: ballarin@17000: datatype 'a mode = Assumed of 'a | Derived of 'a; ballarin@17000: ballarin@17000: fun map_mode f (Assumed x) = Assumed (f x) ballarin@17000: | map_mode f (Derived x) = Derived (f x); ballarin@17000: wenzelm@18123: wenzelm@12529: (* flatten expressions *) wenzelm@11896: wenzelm@12510: local wenzelm@12502: wenzelm@18137: fun unify_parms ctxt fixed_parms raw_parmss = wenzelm@12502: let wenzelm@16458: val thy = ProofContext.theory_of ctxt; wenzelm@12502: val maxidx = length raw_parmss; wenzelm@12502: val idx_parmss = (0 upto maxidx - 1) ~~ raw_parmss; wenzelm@12502: wenzelm@12502: fun varify i = Term.map_type_tfree (fn (a, S) => TVar ((a, i), S)); wenzelm@12529: fun varify_parms (i, ps) = map (apsnd (varify i)) (param_types ps); wenzelm@19482: val parms = fixed_parms @ maps varify_parms idx_parmss; wenzelm@12502: wenzelm@18137: fun unify T U envir = Sign.typ_unify thy (U, T) envir ballarin@15206: handle Type.TUNIFY => wenzelm@23418: let schirmer@22339: val T' = Envir.norm_type (fst envir) T; schirmer@22339: val U' = Envir.norm_type (fst envir) U; wenzelm@23418: val prt = Sign.string_of_typ thy; schirmer@22339: in wenzelm@18137: raise TYPE ("unify_parms: failed to unify types " ^ schirmer@22339: prt U' ^ " and " ^ prt T', [U', T'], []) wenzelm@18137: end; wenzelm@18137: fun unify_list (T :: Us) = fold (unify T) Us wenzelm@18137: | unify_list [] = I; wenzelm@18952: val (unifier, _) = fold unify_list (map #2 (Symtab.dest (Symtab.make_list parms))) wenzelm@18137: (Vartab.empty, maxidx); wenzelm@12502: wenzelm@19061: val parms' = map (apsnd (Envir.norm_type unifier)) (distinct (eq_fst (op =)) parms); wenzelm@12502: val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (map #2 parms')); wenzelm@12502: wenzelm@12502: fun inst_parms (i, ps) = wenzelm@23178: List.foldr Term.add_typ_tfrees [] (map_filter snd ps) wenzelm@19482: |> map_filter (fn (a, S) => wenzelm@12502: let val T = Envir.norm_type unifier' (TVar ((a, i), S)) wenzelm@18137: in if T = TFree (a, S) then NONE else SOME (a, T) end) wenzelm@18137: |> Symtab.make; wenzelm@18137: in map inst_parms idx_parmss end; wenzelm@12502: wenzelm@12529: in wenzelm@12502: wenzelm@12529: fun unify_elemss _ _ [] = [] wenzelm@12529: | unify_elemss _ [] [elems] = [elems] wenzelm@12529: | unify_elemss ctxt fixed_parms elemss = wenzelm@12502: let wenzelm@18137: val thy = ProofContext.theory_of ctxt; wenzelm@21483: val phis = unify_parms ctxt fixed_parms (map (snd o fst o fst) elemss) wenzelm@21483: |> map (Element.instT_morphism thy); wenzelm@21483: fun inst ((((name, ps), mode), elems), phi) = wenzelm@21483: (((name, map (apsnd (Option.map (Morphism.typ phi))) ps), wenzelm@21483: map_mode (map (Element.morph_witness phi)) mode), wenzelm@21483: map (Element.morph_ctxt phi) elems); wenzelm@21483: in map inst (elemss ~~ phis) end; wenzelm@12502: ballarin@17000: wenzelm@19810: fun renaming xs parms = zip_options parms xs wenzelm@19810: handle Library.UnequalLengths => wenzelm@19810: error ("Too many arguments in renaming: " ^ wenzelm@19810: commas (map (fn NONE => "_" | SOME x => quote (fst x)) xs)); wenzelm@19810: wenzelm@19810: ballarin@19783: (* params_of_expr: ballarin@19783: Compute parameters (with types and syntax) of locale expression. ballarin@19783: *) ballarin@19783: ballarin@19783: fun params_of_expr ctxt fixed_params expr (prev_parms, prev_types, prev_syn) = ballarin@19783: let ballarin@19783: val thy = ProofContext.theory_of ctxt; ballarin@19783: ballarin@19783: fun merge_tenvs fixed tenv1 tenv2 = ballarin@19783: let ballarin@19783: val [env1, env2] = unify_parms ctxt fixed ballarin@19783: [tenv1 |> Symtab.dest |> map (apsnd SOME), ballarin@19783: tenv2 |> Symtab.dest |> map (apsnd SOME)] ballarin@19783: in ballarin@19783: Symtab.merge (op =) (Symtab.map (Element.instT_type env1) tenv1, ballarin@19783: Symtab.map (Element.instT_type env2) tenv2) ballarin@19783: end; ballarin@19783: ballarin@19783: fun merge_syn expr syn1 syn2 = haftmann@19932: Symtab.merge (op = : mixfix * mixfix -> bool) (syn1, syn2) wenzelm@23655: handle Symtab.DUP x => err_in_expr ctxt wenzelm@23655: ("Conflicting syntax for parameter: " ^ quote x) expr; wenzelm@20366: ballarin@19783: fun params_of (expr as Locale name) = ballarin@19783: let wenzelm@22573: val {imports, params, ...} = the_locale thy name; ballarin@19783: val parms = map (fst o fst) params; wenzelm@22573: val (parms', types', syn') = params_of imports; ballarin@19783: val all_parms = merge_lists parms' parms; ballarin@19783: val all_types = merge_tenvs [] types' (params |> map fst |> Symtab.make); ballarin@19783: val all_syn = merge_syn expr syn' (params |> map (apfst fst) |> Symtab.make); ballarin@19783: in (all_parms, all_types, all_syn) end ballarin@19783: | params_of (expr as Rename (e, xs)) = ballarin@19783: let ballarin@19783: val (parms', types', syn') = params_of e; ballarin@19783: val ren = renaming xs parms'; ballarin@19783: (* renaming may reduce number of parameters *) ballarin@19783: val new_parms = map (Element.rename ren) parms' |> distinct (op =); ballarin@19783: val ren_syn = syn' |> Symtab.dest |> map (Element.rename_var ren); ballarin@19783: val new_syn = fold (Symtab.insert (op =)) ren_syn Symtab.empty ballarin@19783: handle Symtab.DUP x => ballarin@19783: err_in_expr ctxt ("Conflicting syntax for parameter: " ^ quote x) expr; wenzelm@22700: val syn_types = map (apsnd (fn mx => wenzelm@22700: SOME (Type.freeze_type (#1 (TypeInfer.paramify_dummies (Syntax.mixfixT mx) 0))))) wenzelm@22700: (Symtab.dest new_syn); ballarin@19783: val ren_types = types' |> Symtab.dest |> map (apfst (Element.rename ren)); wenzelm@20366: val (env :: _) = unify_parms ctxt [] ballarin@19783: ((ren_types |> map (apsnd SOME)) :: map single syn_types); ballarin@19783: val new_types = fold (Symtab.insert (op =)) ballarin@19783: (map (apsnd (Element.instT_type env)) ren_types) Symtab.empty; ballarin@19783: in (new_parms, new_types, new_syn) end ballarin@19783: | params_of (Merge es) = ballarin@19783: fold (fn e => fn (parms, types, syn) => ballarin@19783: let ballarin@19783: val (parms', types', syn') = params_of e ballarin@19783: in ballarin@19783: (merge_lists parms parms', merge_tenvs [] types types', ballarin@19783: merge_syn e syn syn') ballarin@19783: end) ballarin@19783: es ([], Symtab.empty, Symtab.empty) ballarin@19783: ballarin@19783: val (parms, types, syn) = params_of expr; ballarin@19783: in ballarin@19783: (merge_lists prev_parms parms, merge_tenvs fixed_params prev_types types, ballarin@19783: merge_syn expr prev_syn syn) ballarin@19783: end; ballarin@19783: ballarin@19783: fun make_params_ids params = [(("", params), ([], Assumed []))]; ballarin@19783: fun make_raw_params_elemss (params, tenv, syn) = ballarin@19783: [((("", map (fn p => (p, Symtab.lookup tenv p)) params), Assumed []), ballarin@19783: Int [Fixes (map (fn p => ballarin@19783: (p, Symtab.lookup tenv p, Symtab.lookup syn p |> the)) params)])]; ballarin@19783: ballarin@19783: ballarin@15596: (* flatten_expr: ballarin@15596: Extend list of identifiers by those new in locale expression expr. ballarin@15596: Compute corresponding list of lists of locale elements (one entry per ballarin@15596: identifier). ballarin@15596: ballarin@15596: Identifiers represent locale fragments and are in an extended form: ballarin@15596: ((name, ps), (ax_ps, axs)) ballarin@15596: (name, ps) is the locale name with all its parameters. ballarin@15596: (ax_ps, axs) is the locale axioms with its parameters; ballarin@15596: axs are always taken from the top level of the locale hierarchy, ballarin@15596: hence axioms may contain additional parameters from later fragments: ballarin@15596: ps subset of ax_ps. axs is either singleton or empty. ballarin@15596: ballarin@15596: Elements are enriched by identifier-like information: ballarin@15596: (((name, ax_ps), axs), elems) ballarin@15596: The parameters in ax_ps are the axiom parameters, but enriched by type ballarin@15596: info: now each entry is a pair of string and typ option. Axioms are ballarin@15596: type-instantiated. ballarin@15596: ballarin@15596: *) ballarin@15596: ballarin@16102: fun flatten_expr ctxt ((prev_idents, prev_syntax), expr) = wenzelm@12014: let wenzelm@12014: val thy = ProofContext.theory_of ctxt; wenzelm@12263: ballarin@17000: fun rename_parms top ren ((name, ps), (parms, mode)) = ballarin@19783: ((name, map (Element.rename ren) ps), ballarin@19783: if top ballarin@19783: then (map (Element.rename ren) parms, wenzelm@21483: map_mode (map (Element.morph_witness (Element.rename_morphism ren))) mode) ballarin@19783: else (parms, mode)); wenzelm@12263: ballarin@20167: (* add (name, ps) and its registrations, recursively; adjust hyps of witnesses *) ballarin@17000: ballarin@20167: fun add_with_regs ((name, pTs), mode) (wits, ids, visited) = ballarin@20167: if member (fn (a, (b, _)) => a = b) visited (name, map #1 pTs) ballarin@20167: then (wits, ids, visited) ballarin@20167: else wenzelm@20366: let wenzelm@20366: val {params, regs, ...} = the_locale thy name; wenzelm@20366: val pTs' = map #1 params; wenzelm@20366: val ren = map #1 pTs' ~~ map (fn (x, _) => (x, NONE)) pTs; wenzelm@20366: (* dummy syntax, since required by rename *) wenzelm@20366: val pTs'' = map (fn ((p, _), (_, T)) => (p, T)) (pTs ~~ pTs'); wenzelm@20366: val [env] = unify_parms ctxt pTs [map (apsnd SOME) pTs'']; wenzelm@20366: (* propagate parameter types, to keep them consistent *) wenzelm@20366: val regs' = map (fn ((name, ps), wits) => wenzelm@20366: ((name, map (Element.rename ren) ps), wenzelm@20366: map (Element.transfer_witness thy) wits)) regs; wenzelm@20366: val new_regs = regs'; wenzelm@20366: val new_ids = map fst new_regs; wenzelm@21483: val new_idTs = wenzelm@21483: map (apsnd (map (fn p => (p, (the o AList.lookup (op =) pTs) p)))) new_ids; ballarin@17096: wenzelm@20366: val new_wits = new_regs |> map (#2 #> map wenzelm@21483: (Element.morph_witness wenzelm@21483: (Element.instT_morphism thy env $> wenzelm@21483: Element.rename_morphism ren $> wenzelm@21483: Element.satisfy_morphism wits))); wenzelm@20366: val new_ids' = map (fn (id, wits) => wenzelm@20366: (id, ([], Derived wits))) (new_ids ~~ new_wits); wenzelm@20366: val new_idTs' = map (fn ((n, pTs), (_, ([], mode))) => wenzelm@20366: ((n, pTs), mode)) (new_idTs ~~ new_ids'); wenzelm@20366: val new_id = ((name, map #1 pTs), ([], mode)); wenzelm@20366: val (wits', ids', visited') = fold add_with_regs new_idTs' ballarin@20167: (wits @ flat new_wits, ids, visited @ [new_id]); wenzelm@20366: in wenzelm@20366: (wits', ids' @ [new_id], visited') wenzelm@20366: end; ballarin@17000: ballarin@17000: (* distribute top-level axioms over assumed ids *) ballarin@17000: ballarin@17000: fun axiomify all_ps ((name, parms), (_, Assumed _)) axioms = ballarin@17000: let ballarin@17000: val {elems, ...} = the_locale thy name; wenzelm@19482: val ts = maps wenzelm@19482: (fn (Assumes asms, _) => maps (map #1 o #2) asms ballarin@17000: | _ => []) wenzelm@19482: elems; wenzelm@19018: val (axs1, axs2) = chop (length ts) axioms; ballarin@17000: in (((name, parms), (all_ps, Assumed axs1)), axs2) end ballarin@17000: | axiomify all_ps (id, (_, Derived ths)) axioms = ballarin@17000: ((id, (all_ps, Derived ths)), axioms); ballarin@17000: ballarin@17096: (* identifiers of an expression *) ballarin@17096: ballarin@15206: fun identify top (Locale name) = ballarin@15596: (* CB: ids_ax is a list of tuples of the form ((name, ps), axs), ballarin@15206: where name is a locale name, ps a list of parameter names and axs ballarin@15206: a list of axioms relating to the identifier, axs is empty unless ballarin@15206: identify at top level (top = true); ballarin@14215: parms is accumulated list of parameters *) wenzelm@12289: let wenzelm@22573: val {axiom, imports, params, ...} = the_locale thy name; ballarin@19278: val ps = map (#1 o #1) params; wenzelm@22573: val (ids', parms') = identify false imports; ballarin@15206: (* acyclic import dependencies *) ballarin@19931: ballarin@20167: val (_, ids'', _) = add_with_regs ((name, map #1 params), Assumed []) ([], ids', ids'); ballarin@20167: val ids_ax = if top then fst (fold_map (axiomify ps) ids'' axiom) else ids''; ballarin@20035: in (ids_ax, merge_lists parms' ps) end ballarin@15206: | identify top (Rename (e, xs)) = wenzelm@12273: let ballarin@20035: val (ids', parms') = identify top e; wenzelm@12839: val ren = renaming xs parms' wenzelm@18678: handle ERROR msg => err_in_locale' ctxt msg ids'; ballarin@17096: wenzelm@19061: val ids'' = distinct (eq_fst (op =)) (map (rename_parms top ren) ids'); wenzelm@19482: val parms'' = distinct (op =) (maps (#2 o #1) ids''); ballarin@20035: in (ids'', parms'') end ballarin@15206: | identify top (Merge es) = ballarin@20035: fold (fn e => fn (ids, parms) => ballarin@17000: let ballarin@20035: val (ids', parms') = identify top e ballarin@17000: in ballarin@20035: (merge_alists ids ids', merge_lists parms parms') ballarin@17000: end) ballarin@20035: es ([], []); ballarin@15206: ballarin@20035: fun inst_wit all_params (t, th) = let ballarin@15206: val {hyps, prop, ...} = Thm.rep_thm th; wenzelm@16861: val ps = map (apsnd SOME) (fold Term.add_frees (prop :: hyps) []); ballarin@15206: val [env] = unify_parms ctxt all_params [ps]; wenzelm@18137: val t' = Element.instT_term env t; wenzelm@18137: val th' = Element.instT_thm thy env th; wenzelm@18123: in (t', th') end; ballarin@17000: ballarin@20035: fun eval all_params tenv syn ((name, params), (locale_params, mode)) = ballarin@20035: let ballarin@20035: val {params = ps_mx, elems = elems_stamped, ...} = the_locale thy name; ballarin@20035: val elems = map fst elems_stamped; ballarin@20035: val ps = map fst ps_mx; ballarin@20035: fun lookup_syn x = (case Symtab.lookup syn x of SOME Structure => NONE | opt => opt); ballarin@20035: val locale_params' = map (fn p => (p, Symtab.lookup tenv p |> the)) locale_params; ballarin@20035: val mode' = map_mode (map (Element.map_witness (inst_wit all_params))) mode; ballarin@20035: val ren = map fst ps ~~ map (fn p => (p, lookup_syn p)) params; ballarin@20035: val [env] = unify_parms ctxt all_params [map (apfst (Element.rename ren) o apsnd SOME) ps]; wenzelm@21499: val elem_morphism = wenzelm@21499: Element.rename_morphism ren $> wenzelm@21665: Morphism.name_morphism (NameSpace.qualified (params_prefix params)) $> wenzelm@21499: Element.instT_morphism thy env; wenzelm@21499: val elems' = map (Element.morph_ctxt elem_morphism) elems; wenzelm@21499: in (((name, map (apsnd SOME) locale_params'), mode'), elems') end; ballarin@20035: ballarin@20035: (* parameters, their types and syntax *) ballarin@20035: val (all_params', tenv, syn) = params_of_expr ctxt [] expr ([], Symtab.empty, Symtab.empty); ballarin@20035: val all_params = map (fn p => (p, Symtab.lookup tenv p |> the)) all_params'; ballarin@20035: (* compute identifiers and syntax, merge with previous ones *) ballarin@20035: val (ids, _) = identify true expr; haftmann@20951: val idents = subtract (eq_fst (op =)) prev_idents ids; ballarin@20035: val syntax = merge_syntax ctxt ids (syn, prev_syntax); ballarin@20035: (* type-instantiate elements *) ballarin@20035: val final_elemss = map (eval all_params tenv syntax) idents; ballarin@16102: in ((prev_idents @ idents, syntax), final_elemss) end; wenzelm@12046: wenzelm@12510: end; wenzelm@12510: wenzelm@12070: wenzelm@12529: (* activate elements *) wenzelm@12273: wenzelm@12510: local wenzelm@12510: wenzelm@21686: fun axioms_export axs _ As = wenzelm@21686: (Element.satisfy_thm axs #> Drule.implies_intr_list (Library.drop (length axs, As)), fn t => t); wenzelm@12263: ballarin@17000: ballarin@17000: (* NB: derived ids contain only facts at this stage *) ballarin@17000: ballarin@19931: fun activate_elem _ _ ((ctxt, mode), Fixes fixes) = wenzelm@18671: ((ctxt |> ProofContext.add_fixes_i fixes |> snd, mode), []) ballarin@19931: | activate_elem _ _ ((ctxt, mode), Constrains _) = ballarin@17000: ((ctxt, mode), []) ballarin@19931: | activate_elem ax_in_ctxt _ ((ctxt, Assumed axs), Assumes asms) = wenzelm@13399: let wenzelm@18728: val asms' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) asms; wenzelm@19482: val ts = maps (map #1 o #2) asms'; wenzelm@19018: val (ps, qs) = chop (length ts) axs; wenzelm@17856: val (_, ctxt') = wenzelm@21370: ctxt |> fold Variable.auto_fixes ts ballarin@19931: |> ProofContext.add_assms_i (axioms_export (if ax_in_ctxt then ps else [])) asms'; ballarin@17000: in ((ctxt', Assumed qs), []) end ballarin@19931: | activate_elem _ _ ((ctxt, Derived ths), Assumes asms) = ballarin@17000: ((ctxt, Derived ths), []) ballarin@19931: | activate_elem _ _ ((ctxt, Assumed axs), Defines defs) = ballarin@15596: let wenzelm@18728: val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs; wenzelm@19732: val asms = defs' |> map (fn ((name, atts), (t, ps)) => wenzelm@19732: let val ((c, _), t') = LocalDefs.cert_def ctxt t wenzelm@20872: in (t', ((Thm.def_name_optional c name, atts), [(t', ps)])) end); wenzelm@17856: val (_, ctxt') = wenzelm@21370: ctxt |> fold (Variable.auto_fixes o #1) asms wenzelm@19732: |> ProofContext.add_assms_i LocalDefs.def_export (map #2 asms); ballarin@17000: in ((ctxt', Assumed axs), []) end ballarin@19931: | activate_elem _ _ ((ctxt, Derived ths), Defines defs) = ballarin@17000: ((ctxt, Derived ths), []) wenzelm@21441: | activate_elem _ is_ext ((ctxt, mode), Notes (kind, facts)) = ballarin@15596: let wenzelm@18728: val facts' = Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of ctxt)) facts; wenzelm@21441: val (res, ctxt') = ctxt |> ProofContext.note_thmss_i kind facts'; ballarin@17000: in ((ctxt', mode), if is_ext then res else []) end; wenzelm@12502: ballarin@19931: fun activate_elems ax_in_ctxt (((name, ps), mode), elems) ctxt = ballarin@17033: let wenzelm@18123: val thy = ProofContext.theory_of ctxt; ballarin@17033: val ((ctxt', _), res) = wenzelm@21441: foldl_map (activate_elem ax_in_ctxt (name = "")) wenzelm@21441: ((ProofContext.qualified_names ctxt, mode), elems) wenzelm@21441: handle ERROR msg => err_in_locale ctxt msg [(name, map fst ps)]; ballarin@15696: val ctxt'' = if name = "" then ctxt' ballarin@15696: else let ballarin@15696: val ps' = map (fn (n, SOME T) => Free (n, T)) ps; haftmann@22351: val ctxt'' = put_local_registration (name, ps') ((true, ""), []) ctxt' ballarin@17000: in case mode of wenzelm@18123: Assumed axs => wenzelm@19780: fold (add_local_witness (name, ps') o wenzelm@19780: Element.assume_witness thy o Element.witness_prop) axs ctxt'' wenzelm@18123: | Derived ths => fold (add_local_witness (name, ps')) ths ctxt'' ballarin@15696: end wenzelm@16144: in (ProofContext.restore_naming ctxt ctxt'', res) end; wenzelm@13399: ballarin@19931: fun activate_elemss ax_in_ctxt prep_facts = ballarin@17000: fold_map (fn (((name, ps), mode), raw_elems) => fn ctxt => ballarin@17000: let ballarin@17000: val elems = map (prep_facts ctxt) raw_elems; wenzelm@19482: val (ctxt', res) = apsnd flat ballarin@19931: (activate_elems ax_in_ctxt (((name, ps), mode), elems) ctxt); wenzelm@21530: val elems' = elems |> map (Element.map_ctxt_attrib Args.closure); ballarin@19931: in (((((name, ps), mode), elems'), res), ctxt') end); wenzelm@12834: wenzelm@12546: in wenzelm@12546: ballarin@15206: (* CB: activate_facts prep_facts (ctxt, elemss), ballarin@15206: where elemss is a list of pairs consisting of identifiers and ballarin@15206: context elements, extends ctxt by the context elements yielding ballarin@15206: ctxt' and returns (ctxt', (elemss', facts)). ballarin@15206: Identifiers in the argument are of the form ((name, ps), axs) and ballarin@15206: assumptions use the axioms in the identifiers to set up exporters ballarin@15206: in ctxt'. elemss' does not contain identifiers and is obtained ballarin@15206: from elemss and the intermediate context with prep_facts. wenzelm@15703: If read_facts or cert_facts is used for prep_facts, these also remove ballarin@14508: the internal/external markers from elemss. *) ballarin@14508: ballarin@19931: fun activate_facts ax_in_ctxt prep_facts (ctxt, args) = ballarin@19931: let val ((elemss, factss), ctxt') = activate_elemss ax_in_ctxt prep_facts args ctxt |>> split_list wenzelm@19482: in (ctxt', (elemss, flat factss)) end; wenzelm@15703: wenzelm@12510: end; wenzelm@12510: wenzelm@12307: ballarin@15696: wenzelm@18137: (** prepare locale elements **) wenzelm@12529: wenzelm@12529: (* expressions *) wenzelm@12529: wenzelm@16458: fun intern_expr thy (Locale xname) = Locale (intern thy xname) wenzelm@16458: | intern_expr thy (Merge exprs) = Merge (map (intern_expr thy) exprs) wenzelm@16458: | intern_expr thy (Rename (expr, xs)) = Rename (intern_expr thy expr, xs); wenzelm@12529: wenzelm@12529: wenzelm@12529: (* propositions and bindings *) wenzelm@12529: ballarin@17000: (* flatten (ctxt, prep_expr) ((ids, syn), expr) ballarin@17000: normalises expr (which is either a locale ballarin@14508: expression or a single context element) wrt. ballarin@14508: to the list ids of already accumulated identifiers. ballarin@19783: It returns ((ids', syn'), elemss) where ids' is an extension of ids ballarin@14508: with identifiers generated for expr, and elemss is the list of ballarin@16102: context elements generated from expr. ballarin@16102: syn and syn' are symtabs mapping parameter names to their syntax. syn' ballarin@16102: is an extension of syn. ballarin@16102: For details, see flatten_expr. ballarin@16102: ballarin@15596: Additionally, for a locale expression, the elems are grouped into a single ballarin@15596: Int; individual context elements are marked Ext. In this case, the ballarin@15596: identifier-like information of the element is as follows: ballarin@15596: - for Fixes: (("", ps), []) where the ps have type info NONE ballarin@15596: - for other elements: (("", []), []). ballarin@15206: The implementation of activate_facts relies on identifier names being ballarin@15206: empty strings for external elements. ballarin@15596: *) ballarin@14508: ballarin@16102: fun flatten (ctxt, _) ((ids, syn), Elem (Fixes fixes)) = let wenzelm@18137: val ids' = ids @ [(("", map #1 fixes), ([], Assumed []))] ballarin@16102: in wenzelm@18137: ((ids', wenzelm@18137: merge_syntax ctxt ids' wenzelm@18137: (syn, Symtab.make (map (fn fx => (#1 fx, #3 fx)) fixes)) wenzelm@23655: handle Symtab.DUP x => err_in_locale ctxt wenzelm@23655: ("Conflicting syntax for parameter: " ^ quote x) ballarin@16102: (map #1 ids')), wenzelm@18137: [((("", map (rpair NONE o #1) fixes), Assumed []), Ext (Fixes fixes))]) ballarin@16102: end ballarin@16102: | flatten _ ((ids, syn), Elem elem) = ballarin@17000: ((ids @ [(("", []), ([], Assumed []))], syn), [((("", []), Assumed []), Ext elem)]) ballarin@16102: | flatten (ctxt, prep_expr) ((ids, syn), Expr expr) = ballarin@16102: apsnd (map (apsnd Int)) (flatten_expr ctxt ((ids, syn), prep_expr expr)); ballarin@14508: wenzelm@12529: local wenzelm@12529: wenzelm@12839: local wenzelm@12839: wenzelm@12727: fun declare_int_elem (ctxt, Fixes fixes) = wenzelm@18671: (ctxt |> ProofContext.add_fixes_i (map (fn (x, T, mx) => wenzelm@18671: (x, Option.map (Term.map_type_tfree (TypeInfer.param 0)) T, mx)) fixes) |> snd, []) wenzelm@12727: | declare_int_elem (ctxt, _) = (ctxt, []); wenzelm@12529: wenzelm@18671: fun declare_ext_elem prep_vars (ctxt, Fixes fixes) = wenzelm@18671: let val (vars, _) = prep_vars fixes ctxt wenzelm@18671: in (ctxt |> ProofContext.add_fixes_i vars |> snd, []) end wenzelm@18671: | declare_ext_elem prep_vars (ctxt, Constrains csts) = wenzelm@18671: let val (_, ctxt') = prep_vars (map (fn (x, T) => (x, SOME T, NoSyn)) csts) ctxt wenzelm@18671: in (ctxt', []) end wenzelm@12529: | declare_ext_elem _ (ctxt, Assumes asms) = (ctxt, map #2 asms) wenzelm@19585: | declare_ext_elem _ (ctxt, Defines defs) = (ctxt, map (fn (_, (t, ps)) => [(t, ps)]) defs) wenzelm@21441: | declare_ext_elem _ (ctxt, Notes _) = (ctxt, []); wenzelm@12529: wenzelm@18671: fun declare_elems prep_vars (ctxt, (((name, ps), Assumed _), elems)) = ballarin@17000: let val (ctxt', propps) = ballarin@17000: (case elems of ballarin@17000: Int es => foldl_map declare_int_elem (ctxt, es) wenzelm@18671: | Ext e => foldl_map (declare_ext_elem prep_vars) (ctxt, [e])) wenzelm@18678: handle ERROR msg => err_in_locale ctxt msg [(name, map fst ps)] ballarin@17000: in (ctxt', propps) end ballarin@17000: | declare_elems _ (ctxt, ((_, Derived _), elems)) = (ctxt, []); wenzelm@12727: wenzelm@12839: in wenzelm@12839: wenzelm@18671: fun declare_elemss prep_vars fixed_params raw_elemss ctxt = wenzelm@12727: let ballarin@14215: (* CB: fix of type bug of goal in target with context elements. ballarin@14215: Parameters new in context elements must receive types that are ballarin@14215: distinct from types of parameters in target (fixed_params). *) ballarin@14215: val ctxt_with_fixed = wenzelm@19900: fold Variable.declare_term (map Free fixed_params) ctxt; wenzelm@12727: val int_elemss = wenzelm@12727: raw_elemss wenzelm@19482: |> map_filter (fn (id, Int es) => SOME (id, es) | _ => NONE) ballarin@14215: |> unify_elemss ctxt_with_fixed fixed_params; wenzelm@12727: val (_, raw_elemss') = wenzelm@12727: foldl_map (fn ((_, es) :: elemss, (id, Int _)) => (elemss, (id, Int es)) | x => x) wenzelm@12727: (int_elemss, raw_elemss); wenzelm@18671: in foldl_map (declare_elems prep_vars) (ctxt, raw_elemss') end; wenzelm@12529: wenzelm@12839: end; wenzelm@12529: wenzelm@12839: local wenzelm@12839: wenzelm@12839: val norm_term = Envir.beta_norm oo Term.subst_atomic; wenzelm@12839: wenzelm@16458: fun abstract_thm thy eq = wenzelm@16458: Thm.assume (Thm.cterm_of thy eq) |> Drule.gen_all |> Drule.abs_def; wenzelm@12502: wenzelm@18190: fun bind_def ctxt (name, ps) eq (xs, env, ths) = wenzelm@12839: let wenzelm@18831: val ((y, T), b) = LocalDefs.abs_def eq; wenzelm@13308: val b' = norm_term env b; wenzelm@16458: val th = abstract_thm (ProofContext.theory_of ctxt) eq; wenzelm@13308: fun err msg = err_in_locale ctxt (msg ^ ": " ^ quote y) [(name, map fst ps)]; wenzelm@12839: in wenzelm@21962: exists (equal y o #1) xs andalso wenzelm@21962: err "Attempt to define previously specified variable"; wenzelm@21962: exists (fn (Free (y', _), _) => y = y' | _ => false) env andalso wenzelm@21962: err "Attempt to redefine variable"; wenzelm@16861: (Term.add_frees b' xs, (Free (y, T), b') :: env, th :: ths) wenzelm@12839: end; wenzelm@12575: ballarin@17000: ballarin@17000: (* CB: for finish_elems (Int and Ext), ballarin@17000: extracts specification, only of assumed elements *) ballarin@15206: wenzelm@18190: fun eval_text _ _ _ (Fixes _) text = text wenzelm@18190: | eval_text _ _ _ (Constrains _) text = text wenzelm@18190: | eval_text _ (_, Assumed _) is_ext (Assumes asms) wenzelm@18190: (((exts, exts'), (ints, ints')), (xs, env, defs)) = wenzelm@13394: let wenzelm@19482: val ts = maps (map #1 o #2) asms; wenzelm@13394: val ts' = map (norm_term env) ts; wenzelm@13394: val spec' = wenzelm@13394: if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints')) wenzelm@13394: else ((exts, exts'), (ints @ ts, ints' @ ts')); wenzelm@16861: in (spec', (fold Term.add_frees ts' xs, env, defs)) end wenzelm@18190: | eval_text _ (_, Derived _) _ (Assumes _) text = text wenzelm@18190: | eval_text ctxt (id, Assumed _) _ (Defines defs) (spec, binds) = wenzelm@18190: (spec, fold (bind_def ctxt id o #1 o #2) defs binds) wenzelm@18190: | eval_text _ (_, Derived _) _ (Defines _) text = text wenzelm@18190: | eval_text _ _ _ (Notes _) text = text; wenzelm@13308: ballarin@17000: ballarin@17000: (* for finish_elems (Int), ballarin@17000: remove redundant elements of derived identifiers, ballarin@17000: turn assumptions and definitions into facts, wenzelm@21483: satisfy hypotheses of facts *) ballarin@17000: ballarin@17096: fun finish_derived _ _ (Assumed _) (Fixes fixes) = SOME (Fixes fixes) ballarin@17096: | finish_derived _ _ (Assumed _) (Constrains csts) = SOME (Constrains csts) ballarin@17096: | finish_derived _ _ (Assumed _) (Assumes asms) = SOME (Assumes asms) ballarin@17096: | finish_derived _ _ (Assumed _) (Defines defs) = SOME (Defines defs) ballarin@17096: ballarin@17000: | finish_derived _ _ (Derived _) (Fixes _) = NONE ballarin@17000: | finish_derived _ _ (Derived _) (Constrains _) = NONE wenzelm@21483: | finish_derived sign satisfy (Derived _) (Assumes asms) = asms ballarin@17096: |> map (apsnd (map (fn (a, _) => ([Thm.assume (cterm_of sign a)], [])))) wenzelm@21441: |> pair Thm.assumptionK |> Notes wenzelm@21483: |> Element.morph_ctxt satisfy |> SOME wenzelm@21483: | finish_derived sign satisfy (Derived _) (Defines defs) = defs ballarin@17096: |> map (apsnd (fn (d, _) => [([Thm.assume (cterm_of sign d)], [])])) wenzelm@21441: |> pair Thm.definitionK |> Notes wenzelm@21483: |> Element.morph_ctxt satisfy |> SOME ballarin@17000: wenzelm@21483: | finish_derived _ satisfy _ (Notes facts) = Notes facts wenzelm@21483: |> Element.morph_ctxt satisfy |> SOME; ballarin@17000: ballarin@15206: (* CB: for finish_elems (Ext) *) ballarin@15206: wenzelm@13308: fun closeup _ false elem = elem wenzelm@13308: | closeup ctxt true elem = wenzelm@12839: let wenzelm@13308: fun close_frees t = wenzelm@19900: let val frees = rev (filter_out (Variable.is_fixed ctxt o #1) (Term.add_frees t [])) wenzelm@13308: in Term.list_all_free (frees, t) end; wenzelm@13308: wenzelm@13308: fun no_binds [] = [] wenzelm@18678: | no_binds _ = error "Illegal term bindings in locale element"; wenzelm@13308: in wenzelm@13308: (case elem of wenzelm@13308: Assumes asms => Assumes (asms |> map (fn (a, propps) => wenzelm@19585: (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps))) wenzelm@13308: | Defines defs => Defines (defs |> map (fn (a, (t, ps)) => wenzelm@18831: (a, (close_frees (#2 (LocalDefs.cert_def ctxt t)), no_binds ps)))) wenzelm@13308: | e => e) wenzelm@13308: end; wenzelm@12839: wenzelm@12502: wenzelm@12839: fun finish_ext_elem parms _ (Fixes fixes, _) = Fixes (map (fn (x, _, mx) => haftmann@17271: (x, AList.lookup (op =) parms x, mx)) fixes) wenzelm@18899: | finish_ext_elem parms _ (Constrains _, _) = Constrains [] wenzelm@12839: | finish_ext_elem _ close (Assumes asms, propp) = wenzelm@12839: close (Assumes (map #1 asms ~~ propp)) wenzelm@12839: | finish_ext_elem _ close (Defines defs, propp) = wenzelm@19585: close (Defines (map #1 defs ~~ map (fn [(t, ps)] => (t, ps)) propp)) wenzelm@12839: | finish_ext_elem _ _ (Notes facts, _) = Notes facts; wenzelm@12839: ballarin@17000: ballarin@15206: (* CB: finish_parms introduces type info from parms to identifiers *) skalberg@15531: (* CB: only needed for types that have been NONE so far??? ballarin@15206: If so, which are these??? *) ballarin@15206: ballarin@17000: fun finish_parms parms (((name, ps), mode), elems) = haftmann@19932: (((name, map (fn (x, _) => (x, AList.lookup (op = : string * string -> bool) parms x)) ps), mode), elems); wenzelm@12839: ballarin@17000: fun finish_elems ctxt parms _ ((text, wits), ((id, Int e), _)) = wenzelm@12839: let ballarin@17000: val [(id' as (_, mode), es)] = unify_elemss ctxt parms [(id, e)]; ballarin@17000: val wits' = case mode of Assumed _ => wits | Derived ths => wits @ ths; wenzelm@18190: val text' = fold (eval_text ctxt id' false) es text; wenzelm@19482: val es' = map_filter wenzelm@21483: (finish_derived (ProofContext.theory_of ctxt) (Element.satisfy_morphism wits') mode) es; ballarin@17000: in ((text', wits'), (id', map Int es')) end ballarin@17000: | finish_elems ctxt parms do_close ((text, wits), ((id, Ext e), [propp])) = wenzelm@13308: let wenzelm@13308: val e' = finish_ext_elem parms (closeup ctxt do_close) (e, propp); wenzelm@18190: val text' = eval_text ctxt id true e' text; ballarin@17000: in ((text', wits), (id, [Ext e'])) end wenzelm@12839: wenzelm@12839: in wenzelm@12510: ballarin@15206: (* CB: only called by prep_elemss *) ballarin@15206: wenzelm@13375: fun finish_elemss ctxt parms do_close = wenzelm@13375: foldl_map (apsnd (finish_parms parms) o finish_elems ctxt parms do_close); wenzelm@12839: wenzelm@12839: end; wenzelm@12839: ballarin@16736: ballarin@19942: (* Remove duplicate Defines elements: temporary workaround to fix Afp/Category. *) ballarin@19942: ballarin@19942: fun defs_ord (defs1, defs2) = ballarin@19942: list_ord (fn ((_, (d1, _)), (_, (d2, _))) => ballarin@19942: Term.fast_term_ord (d1, d2)) (defs1, defs2); ballarin@19942: structure Defstab = ballarin@19942: TableFun(type key = ((string * Attrib.src list) * (term * term list)) list ballarin@19942: val ord = defs_ord); ballarin@19942: ballarin@19942: fun rem_dup_defs es ds = ballarin@19942: fold_map (fn e as (Defines defs) => (fn ds => ballarin@19942: if Defstab.defined ds defs ballarin@19942: then (Defines [], ds) ballarin@19942: else (e, Defstab.update (defs, ()) ds)) ballarin@19942: | e => (fn ds => (e, ds))) es ds; ballarin@19942: fun rem_dup_elemss (Int es) ds = apfst Int (rem_dup_defs es ds) ballarin@19942: | rem_dup_elemss (Ext e) ds = (Ext e, ds); ballarin@19942: fun rem_dup_defines raw_elemss = ballarin@19942: fold_map (fn (id as (_, (Assumed _)), es) => (fn ds => ballarin@19942: apfst (pair id) (rem_dup_elemss es ds)) ballarin@19942: | (id as (_, (Derived _)), es) => (fn ds => ballarin@19942: ((id, es), ds))) raw_elemss Defstab.empty |> #1; ballarin@19942: ballarin@16736: (* CB: type inference and consistency checks for locales. ballarin@16736: ballarin@16736: Works by building a context (through declare_elemss), extracting the ballarin@16736: required information and adjusting the context elements (finish_elemss). ballarin@16736: Can also universally close free vars in assms and defs. This is only ballarin@17000: needed for Ext elements and controlled by parameter do_close. ballarin@17000: ballarin@17000: Only elements of assumed identifiers are considered. ballarin@16736: *) ballarin@15127: wenzelm@18671: fun prep_elemss prep_vars prepp do_close context fixed_params raw_elemss raw_concl = wenzelm@12529: let ballarin@15127: (* CB: contexts computed in the course of this function are discarded. ballarin@15127: They are used for type inference and consistency checks only. *) ballarin@15206: (* CB: fixed_params are the parameters (with types) of the target locale, ballarin@15206: empty list if there is no target. *) ballarin@14508: (* CB: raw_elemss are list of pairs consisting of identifiers and ballarin@14508: context elements, the latter marked as internal or external. *) ballarin@19942: val raw_elemss = rem_dup_defines raw_elemss; wenzelm@18671: val (raw_ctxt, raw_proppss) = declare_elemss prep_vars fixed_params raw_elemss context; ballarin@14508: (* CB: raw_ctxt is context with additional fixed variables derived from ballarin@14508: the fixes elements in raw_elemss, ballarin@14508: raw_proppss contains assumptions and definitions from the ballarin@15206: external elements in raw_elemss. *) haftmann@18550: fun prep_prop raw_propp (raw_ctxt, raw_concl) = haftmann@18450: let haftmann@18450: (* CB: add type information from fixed_params to context (declare_term) *) haftmann@18450: (* CB: process patterns (conclusion and external elements only) *) haftmann@18450: val (ctxt, all_propp) = wenzelm@19900: prepp (fold Variable.declare_term (map Free fixed_params) raw_ctxt, raw_concl @ raw_propp); haftmann@18450: (* CB: add type information from conclusion and external elements to context *) wenzelm@19900: val ctxt = fold Variable.declare_term (maps (map fst) all_propp) ctxt; haftmann@18450: (* CB: resolve schematic variables (patterns) in conclusion and external elements. *) haftmann@18450: val all_propp' = map2 (curry (op ~~)) haftmann@18450: (#1 (#2 (ProofContext.bind_propp_schematic_i (ctxt, all_propp)))) (map (map snd) all_propp); wenzelm@19018: val (concl, propp) = chop (length raw_concl) all_propp'; haftmann@18550: in (propp, (ctxt, concl)) end ballarin@15206: haftmann@18550: val (proppss, (ctxt, concl)) = haftmann@18550: (fold_burrow o fold_burrow) prep_prop raw_proppss (raw_ctxt, raw_concl); wenzelm@12502: ballarin@15206: (* CB: obtain all parameters from identifier part of raw_elemss *) ballarin@15206: val xs = map #1 (params_of' raw_elemss); wenzelm@12727: val typing = unify_frozen ctxt 0 wenzelm@19900: (map (Variable.default_type raw_ctxt) xs) wenzelm@19900: (map (Variable.default_type ctxt) xs); wenzelm@12529: val parms = param_types (xs ~~ typing); ballarin@14508: (* CB: parms are the parameters from raw_elemss, with correct typing. *) wenzelm@12273: ballarin@14508: (* CB: extract information from assumes and defines elements ballarin@16169: (fixes, constrains and notes in raw_elemss don't have an effect on ballarin@16169: text and elemss), compute final form of context elements. *) ballarin@17000: val ((text, _), elemss) = finish_elemss ctxt parms do_close ballarin@17000: ((((([], []), ([], [])), ([], [], [])), []), raw_elemss ~~ proppss); ballarin@14508: (* CB: text has the following structure: ballarin@14508: (((exts, exts'), (ints, ints')), (xs, env, defs)) ballarin@14508: where ballarin@14508: exts: external assumptions (terms in external assumes elements) ballarin@14508: exts': dito, normalised wrt. env ballarin@14508: ints: internal assumptions (terms in internal assumes elements) ballarin@14508: ints': dito, normalised wrt. env ballarin@14508: xs: the free variables in exts' and ints' and rhss of definitions, ballarin@14508: this includes parameters except defined parameters ballarin@14508: env: list of term pairs encoding substitutions, where the first term ballarin@14508: is a free variable; substitutions represent defines elements and ballarin@14508: the rhs is normalised wrt. the previous env ballarin@14508: defs: theorems representing the substitutions from defines elements ballarin@14508: (thms are normalised wrt. env). ballarin@14508: elemss is an updated version of raw_elemss: ballarin@16169: - type info added to Fixes and modified in Constrains ballarin@14508: - axiom and definition statement replaced by corresponding one ballarin@14508: from proppss in Assumes and Defines ballarin@14508: - Facts unchanged ballarin@14508: *) wenzelm@13308: in ((parms, elemss, concl), text) end; wenzelm@12502: wenzelm@12502: in wenzelm@12502: wenzelm@18671: fun read_elemss x = prep_elemss ProofContext.read_vars ProofContext.read_propp_schematic x; wenzelm@18671: fun cert_elemss x = prep_elemss ProofContext.cert_vars ProofContext.cert_propp_schematic x; wenzelm@12529: wenzelm@12529: end; wenzelm@12529: wenzelm@12529: wenzelm@15703: (* facts and attributes *) wenzelm@12529: wenzelm@12529: local wenzelm@12529: wenzelm@20965: fun check_name name = wenzelm@18678: if NameSpace.is_qualified name then error ("Illegal qualified name: " ^ quote name) wenzelm@15703: else name; wenzelm@12529: wenzelm@21499: fun prep_facts _ _ _ ctxt (Int elem) = elem wenzelm@21499: |> Element.morph_ctxt (Morphism.thm_morphism (Thm.transfer (ProofContext.theory_of ctxt))) wenzelm@20965: | prep_facts prep_name get intern ctxt (Ext elem) = elem |> Element.map_ctxt wenzelm@15703: {var = I, typ = I, term = I, wenzelm@18678: name = prep_name, wenzelm@15703: fact = get ctxt, wenzelm@16458: attrib = Args.assignable o intern (ProofContext.theory_of ctxt)}; wenzelm@12529: wenzelm@12529: in wenzelm@12529: wenzelm@20965: fun read_facts x = prep_facts check_name ProofContext.get_thms Attrib.intern_src x; wenzelm@20965: fun cert_facts x = prep_facts I (K I) (K I) x; wenzelm@12529: wenzelm@12529: end; wenzelm@12529: wenzelm@12529: ballarin@19931: (* Get the specification of a locale *) ballarin@18795: wenzelm@19780: (*The global specification is made from the parameters and global wenzelm@19780: assumptions, the local specification from the parameters and the wenzelm@19780: local assumptions.*) ballarin@18795: ballarin@18795: local ballarin@18795: ballarin@18795: fun gen_asms_of get thy name = ballarin@18795: let ballarin@18795: val ctxt = ProofContext.init thy; ballarin@18795: val (_, raw_elemss) = flatten (ctxt, I) (([], Symtab.empty), Expr (Locale name)); ballarin@18795: val ((_, elemss, _), _) = read_elemss false ctxt [] raw_elemss []; ballarin@18795: in ballarin@18890: elemss |> get wenzelm@19780: |> maps (fn (_, es) => map (fn Int e => e) es) wenzelm@19780: |> maps (fn Assumes asms => asms | _ => []) ballarin@18795: |> map (apsnd (map fst)) ballarin@18795: end; ballarin@18795: ballarin@18795: in ballarin@18795: ballarin@18795: fun parameters_of thy name = ballarin@19278: the_locale thy name |> #params; ballarin@18795: ballarin@19276: fun parameters_of_expr thy expr = ballarin@19276: let ballarin@19276: val ctxt = ProofContext.init thy; ballarin@19783: val pts = params_of_expr ctxt [] (intern_expr thy expr) ballarin@19783: ([], Symtab.empty, Symtab.empty); ballarin@19783: val raw_params_elemss = make_raw_params_elemss pts; ballarin@19276: val ((_, syn), raw_elemss) = flatten (ctxt, intern_expr thy) ballarin@19276: (([], Symtab.empty), Expr expr); ballarin@19783: val ((parms, _, _), _) = ballarin@19783: read_elemss false ctxt [] (raw_params_elemss @ raw_elemss) []; ballarin@19276: in map (fn p as (n, _) => (p, Symtab.lookup syn n |> the)) parms end; ballarin@19276: ballarin@18795: fun local_asms_of thy name = ballarin@18890: gen_asms_of (single o Library.last_elem) thy name; ballarin@18795: ballarin@18795: fun global_asms_of thy name = ballarin@18890: gen_asms_of I thy name; ballarin@18795: wenzelm@19780: end; ballarin@18795: haftmann@22523: fun intros thy = haftmann@22523: #intros o the o Symtab.lookup (#2 (GlobalLocalesData.get thy)); haftmann@22523: (*returns introduction rule for delta predicate and locale predicate haftmann@22523: as a pair of singleton lists*) haftmann@22523: ballarin@18795: wenzelm@22573: (* full context statements: imports + elements + conclusion *) wenzelm@12529: wenzelm@12529: local wenzelm@12529: wenzelm@12529: fun prep_context_statement prep_expr prep_elemss prep_facts wenzelm@22573: do_close fixed_params imports elements raw_concl context = wenzelm@12529: let wenzelm@16458: val thy = ProofContext.theory_of context; wenzelm@13375: ballarin@19783: val (import_params, import_tenv, import_syn) = wenzelm@22573: params_of_expr context fixed_params (prep_expr thy imports) ballarin@19783: ([], Symtab.empty, Symtab.empty); ballarin@19783: val includes = map_filter (fn Expr e => SOME e | Elem _ => NONE) elements; ballarin@19783: val (incl_params, incl_tenv, incl_syn) = fold (params_of_expr context fixed_params) ballarin@19783: (map (prep_expr thy) includes) (import_params, import_tenv, import_syn); ballarin@19783: ballarin@19783: val ((import_ids, _), raw_import_elemss) = wenzelm@22573: flatten (context, prep_expr thy) (([], Symtab.empty), Expr imports); ballarin@14215: (* CB: normalise "includes" among elements *) wenzelm@16458: val ((ids, syn), raw_elemsss) = foldl_map (flatten (context, prep_expr thy)) ballarin@19783: ((import_ids, incl_syn), elements); ballarin@15696: wenzelm@19482: val raw_elemss = flat raw_elemsss; ballarin@14508: (* CB: raw_import_elemss @ raw_elemss is the normalised list of ballarin@14508: context elements obtained from import and elements. *) ballarin@19783: (* Now additional elements for parameters are inserted. *) ballarin@19783: val import_params_ids = make_params_ids import_params; ballarin@19783: val incl_params_ids = ballarin@19783: make_params_ids (incl_params \\ import_params); ballarin@19783: val raw_import_params_elemss = ballarin@19783: make_raw_params_elemss (import_params, incl_tenv, incl_syn); ballarin@19783: val raw_incl_params_elemss = ballarin@19783: make_raw_params_elemss (incl_params \\ import_params, incl_tenv, incl_syn); wenzelm@13375: val ((parms, all_elemss, concl), (spec, (_, _, defs))) = prep_elemss do_close ballarin@19783: context fixed_params ballarin@19783: (raw_import_params_elemss @ raw_import_elemss @ raw_incl_params_elemss @ raw_elemss) raw_concl; ballarin@19783: ballarin@15696: (* replace extended ids (for axioms) by ids *) ballarin@19783: val (import_ids', incl_ids) = chop (length import_ids) ids; ballarin@20035: val all_ids = import_params_ids @ import_ids' @ incl_params_ids @ incl_ids; ballarin@17000: val all_elemss' = map (fn (((_, ps), _), (((n, ps'), mode), elems)) => haftmann@17485: (((n, map (fn p => (p, (the o AList.lookup (op =) ps') p)) ps), mode), elems)) ballarin@20035: (all_ids ~~ all_elemss); ballarin@19783: (* CB: all_elemss and parms contain the correct parameter types *) ballarin@15696: ballarin@19783: val (ps, qs) = chop (length raw_import_params_elemss + length raw_import_elemss) all_elemss'; ballarin@15206: val (import_ctxt, (import_elemss, _)) = ballarin@19931: activate_facts false prep_facts (context, ps); ballarin@14215: ballarin@15206: val (ctxt, (elemss, _)) = wenzelm@21441: activate_facts false prep_facts (ProofContext.set_stmt true import_ctxt, qs); wenzelm@12834: in ballarin@19783: ((((import_ctxt, import_elemss), (ctxt, elemss, syn)), ballarin@19991: (parms, spec, defs)), concl) wenzelm@12834: end; wenzelm@12529: wenzelm@18806: fun prep_statement prep_locale prep_ctxt raw_locale elems concl ctxt = wenzelm@12529: let wenzelm@12529: val thy = ProofContext.theory_of ctxt; wenzelm@16458: val locale = Option.map (prep_locale thy) raw_locale; wenzelm@22573: val (fixed_params, imports) = wenzelm@18806: (case locale of ballarin@19931: NONE => ([], empty) skalberg@15531: | SOME name => ballarin@19931: let val {params = ps, ...} = the_locale thy name ballarin@19931: in (map fst ps, Locale name) end); ballarin@19991: val ((((locale_ctxt, _), (elems_ctxt, _, _)), _), concl') = wenzelm@22573: prep_ctxt false fixed_params imports elems concl ctxt; ballarin@19991: in (locale, locale_ctxt, elems_ctxt, concl') end; wenzelm@13399: wenzelm@22573: fun prep_expr prep imports body ctxt = wenzelm@19780: let wenzelm@22573: val (((_, import_elemss), (ctxt', elemss, _)), _) = prep imports body ctxt; wenzelm@19780: val all_elems = maps snd (import_elemss @ elemss); wenzelm@19780: in (all_elems, ctxt') end; wenzelm@19780: wenzelm@12529: in wenzelm@12529: wenzelm@18806: val read_ctxt = prep_context_statement intern_expr read_elemss read_facts; wenzelm@18806: val cert_ctxt = prep_context_statement (K I) cert_elemss cert_facts; ballarin@14215: wenzelm@22573: fun read_context imports body ctxt = #1 (read_ctxt true [] imports (map Elem body) [] ctxt); wenzelm@22573: fun cert_context imports body ctxt = #1 (cert_ctxt true [] imports (map Elem body) [] ctxt); wenzelm@12502: wenzelm@19780: val read_expr = prep_expr read_context; wenzelm@19780: val cert_expr = prep_expr cert_context; wenzelm@19780: wenzelm@21035: fun read_context_statement loc = prep_statement intern read_ctxt loc; wenzelm@21035: fun read_context_statement_i loc = prep_statement (K I) read_ctxt loc; wenzelm@21035: fun cert_context_statement loc = prep_statement (K I) cert_ctxt loc; wenzelm@18806: wenzelm@12502: end; wenzelm@11896: wenzelm@11896: wenzelm@21665: (* init *) wenzelm@21665: wenzelm@21665: fun init loc = wenzelm@21665: ProofContext.init wenzelm@21665: #> (#2 o cert_context_statement (SOME loc) [] []); wenzelm@21665: wenzelm@21665: wenzelm@13336: (* print locale *) wenzelm@12070: wenzelm@22573: fun print_locale thy show_facts imports body = wenzelm@22573: let val (all_elems, ctxt) = read_expr imports body (ProofContext.init thy) in wenzelm@18137: Pretty.big_list "locale elements:" (all_elems ballarin@17316: |> (if show_facts then I else filter (fn Notes _ => false | _ => true)) wenzelm@21701: |> map (Element.pretty_ctxt ctxt) |> filter_out null wenzelm@21701: |> map Pretty.chunks) wenzelm@13336: |> Pretty.writeln wenzelm@12277: end; wenzelm@12070: wenzelm@12070: wenzelm@12706: wenzelm@16144: (** store results **) wenzelm@12702: wenzelm@19018: (* naming of interpreted theorems *) ballarin@15696: haftmann@22351: fun global_note_prefix_i kind (fully_qualified, prfx) args thy = wenzelm@16144: thy wenzelm@22796: |> Sign.qualified_names wenzelm@22796: |> (if fully_qualified then Sign.sticky_prefix prfx else Sign.add_path prfx) wenzelm@16144: |> PureThy.note_thmss_i kind args wenzelm@22796: ||> Sign.restore_naming thy; ballarin@15696: haftmann@22351: fun local_note_prefix_i kind (fully_qualified, prfx) args ctxt = wenzelm@16144: ctxt wenzelm@19061: |> ProofContext.qualified_names haftmann@22351: |> (if fully_qualified then ProofContext.sticky_prefix prfx else ProofContext.add_path prfx) wenzelm@21441: |> ProofContext.note_thmss_i kind args wenzelm@19780: ||> ProofContext.restore_naming ctxt; wenzelm@16144: ballarin@15696: ballarin@22658: (* collect witnesses and equations up to a particular target for global ballarin@22658: registration; requires parameters and flattened list of identifiers ballarin@17138: instead of recomputing it from the target *) ballarin@17138: ballarin@17138: fun collect_global_witnesses thy parms ids vts = let ballarin@17138: val ts = map Logic.unvarify vts; ballarin@17138: val (parms, parmTs) = split_list parms; wenzelm@19810: val parmvTs = map Logic.varifyT parmTs; ballarin@17138: val vtinst = fold (Sign.typ_match thy) (parmvTs ~~ map Term.fastype_of ts) Vartab.empty; ballarin@17138: val tinst = Vartab.dest vtinst |> map (fn ((x, 0), (_, T)) => (x, T)) wenzelm@18137: |> Symtab.make; ballarin@17138: (* replace parameter names in ids by instantiations *) ballarin@17138: val vinst = Symtab.make (parms ~~ vts); wenzelm@17412: fun vinst_names ps = map (the o Symtab.lookup vinst) ps; ballarin@17138: val inst = Symtab.make (parms ~~ ts); ballarin@22658: val inst_ids = map (apfst (apsnd vinst_names)) ids; ballarin@22658: val assumed_ids' = map_filter (fn (id, (_, Assumed _)) => SOME id | _ => NONE) inst_ids; ballarin@22658: val wits = maps (#2 o the o get_global_registration thy) assumed_ids'; ballarin@22658: ballarin@22658: val ids' = map fst inst_ids; ballarin@22658: (* TODO: code duplication with activate_facts_elemss *) ballarin@22658: fun add_eqns id eqns = ballarin@22658: let ballarin@22658: val eqns' = case get_global_registration thy id ballarin@22658: of NONE => eqns ballarin@22658: | SOME (_, _, eqns') => Termtab.join (fn t => fn (_, e) => e) (eqns, eqns') ballarin@22658: in ((id, eqns'), eqns') end; ballarin@22658: val eqns = fold_map add_eqns ids' Termtab.empty |> snd |> Termtab.dest |> map snd; ballarin@22658: in ((tinst, inst), wits, eqns) end; ballarin@17138: ballarin@17138: ballarin@15696: (* store instantiations of args for all registered interpretations ballarin@15696: of the theory *) ballarin@15696: wenzelm@21441: fun note_thmss_registrations target (kind, args) thy = ballarin@15596: let ballarin@19278: val parms = the_locale thy target |> #params |> map fst; wenzelm@16458: val ids = flatten (ProofContext.init thy, intern_expr thy) ballarin@22658: (([], Symtab.empty), Expr (Locale target)) |> fst |> fst; ballarin@15696: ballarin@15696: val regs = get_global_registrations thy target; ballarin@15696: ballarin@15696: (* add args to thy for all registrations *) ballarin@15596: ballarin@22658: fun activate (vts, (((fully_qualified, prfx), atts2), _, _)) thy = ballarin@15696: let ballarin@22658: val (insts, prems, eqns) = collect_global_witnesses thy parms ids vts; wenzelm@20911: val attrib = Attrib.attribute_i thy; wenzelm@22241: val inst_atts = Args.morph_values wenzelm@22241: (Morphism.name_morphism (NameSpace.qualified prfx) $> wenzelm@22241: Element.inst_morphism thy insts $> wenzelm@22241: Element.satisfy_morphism prems $> ballarin@22756: Morphism.term_morphism (MetaSimplifier.rewrite_term thy (map Element.conclude_witness eqns) []) $> ballarin@22756: Morphism.thm_morphism (MetaSimplifier.rewrite_rule (map Element.conclude_witness eqns) #> Drule.standard)); wenzelm@20911: val inst_thm = ballarin@22658: Element.inst_thm thy insts #> Element.satisfy_thm prems #> ballarin@22756: MetaSimplifier.rewrite_rule (map Element.conclude_witness eqns) #> ballarin@22658: Drule.standard; wenzelm@20911: val args' = args |> map (fn ((name, atts), bs) => wenzelm@20911: ((name, map (attrib o inst_atts) atts), wenzelm@20911: bs |> map (fn (ths, more_atts) => wenzelm@20911: (map inst_thm ths, map attrib (map inst_atts more_atts @ atts2))))); haftmann@22351: in global_note_prefix_i kind (fully_qualified, prfx) args' thy |> snd end; wenzelm@18190: in fold activate regs thy end; ballarin@15596: ballarin@15596: wenzelm@20911: (* locale results *) wenzelm@12958: wenzelm@21441: fun add_thmss loc kind args ctxt = wenzelm@12706: let wenzelm@21582: val (ctxt', ([(_, [Notes args'])], _)) = wenzelm@21441: activate_facts true cert_facts wenzelm@21441: (ctxt, [((("", []), Assumed []), [Ext (Notes (kind, args))])]); wenzelm@20911: val ctxt'' = ctxt' |> ProofContext.theory wenzelm@20911: (change_locale loc wenzelm@22573: (fn (axiom, imports, elems, params, lparams, decls, regs, intros) => wenzelm@22573: (axiom, imports, elems @ [(Notes args', stamp ())], wenzelm@21499: params, lparams, decls, regs, intros)) wenzelm@21441: #> note_thmss_registrations loc args'); wenzelm@21582: in ctxt'' end; wenzelm@15703: wenzelm@11896: wenzelm@21665: (* declarations *) wenzelm@21665: wenzelm@21665: local wenzelm@21665: wenzelm@21665: val dummy_thm = Drule.mk_term (Thm.cterm_of ProtoPure.thy (Term.dummy_pattern propT)); wenzelm@21665: fun decl_attrib decl phi = Thm.declaration_attribute (K (decl phi)); wenzelm@21665: wenzelm@21665: fun add_decls add loc decl = wenzelm@21665: ProofContext.theory (change_locale loc wenzelm@22573: (fn (axiom, imports, elems, params, lparams, decls, regs, intros) => wenzelm@22573: (axiom, imports, elems, params, lparams, add (decl, stamp ()) decls, regs, intros))) #> wenzelm@21701: add_thmss loc Thm.internalK [(("", [Attrib.internal (decl_attrib decl)]), [([dummy_thm], [])])]; wenzelm@21665: wenzelm@23418: in wenzelm@21665: wenzelm@21665: val add_type_syntax = add_decls (apfst o cons); wenzelm@21665: val add_term_syntax = add_decls (apsnd o cons); wenzelm@21665: val add_declaration = add_decls (K I); wenzelm@21665: wenzelm@21665: end; wenzelm@21665: wenzelm@21665: wenzelm@18137: wenzelm@18137: (** define locales **) wenzelm@18137: wenzelm@13336: (* predicate text *) ballarin@15596: (* CB: generate locale predicates and delta predicates *) wenzelm@13336: wenzelm@13375: local wenzelm@13375: ballarin@15206: (* introN: name of theorems for introduction rules of locale and ballarin@15206: delta predicates; ballarin@15206: axiomsN: name of theorem set with destruct rules for locale predicates, ballarin@15206: also name suffix of delta predicates. *) ballarin@15206: wenzelm@13375: val introN = "intro"; ballarin@15206: val axiomsN = "axioms"; wenzelm@13375: wenzelm@16458: fun atomize_spec thy ts = wenzelm@13375: let wenzelm@23418: val t = Logic.mk_conjunction_balanced ts; wenzelm@16458: val body = ObjectLogic.atomize_term thy t; wenzelm@13375: val bodyT = Term.fastype_of body; wenzelm@13375: in wenzelm@16458: if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of thy t)) wenzelm@23591: else (body, bodyT, ObjectLogic.atomize (Thm.cterm_of thy t)) wenzelm@13375: end; wenzelm@13375: wenzelm@13394: fun aprop_tr' n c = (c, fn args => wenzelm@13394: if length args = n then Syntax.const "_aprop" $ Term.list_comb (Syntax.free c, args) wenzelm@13394: else raise Match); wenzelm@13336: ballarin@15104: (* CB: define one predicate including its intro rule and axioms ballarin@15104: - bname: predicate name ballarin@15104: - parms: locale parameters ballarin@15104: - defs: thms representing substitutions from defines elements ballarin@15104: - ts: terms representing locale assumptions (not normalised wrt. defs) ballarin@15104: - norm_ts: terms representing locale assumptions (normalised wrt. defs) ballarin@15104: - thy: the theory ballarin@15104: *) ballarin@15104: wenzelm@13420: fun def_pred bname parms defs ts norm_ts thy = wenzelm@13375: let wenzelm@16458: val name = Sign.full_name thy bname; wenzelm@13375: wenzelm@16458: val (body, bodyT, body_eq) = atomize_spec thy norm_ts; wenzelm@13394: val env = Term.add_term_free_names (body, []); wenzelm@20664: val xs = filter (member (op =) env o #1) parms; wenzelm@13394: val Ts = map #2 xs; wenzelm@23178: val extraTs = (Term.term_tfrees body \\ List.foldr Term.add_typ_tfrees [] Ts) wenzelm@13394: |> Library.sort_wrt #1 |> map TFree; wenzelm@13399: val predT = map Term.itselfT extraTs ---> Ts ---> bodyT; wenzelm@13336: wenzelm@13394: val args = map Logic.mk_type extraTs @ map Free xs; wenzelm@13394: val head = Term.list_comb (Const (name, predT), args); wenzelm@18123: val statement = ObjectLogic.ensure_propT thy head; wenzelm@13375: haftmann@18358: val ([pred_def], defs_thy) = wenzelm@13375: thy wenzelm@13394: |> (if bodyT <> propT then I else wenzelm@22796: Sign.add_trfuns ([], [], map (aprop_tr' (length args)) (NameSpace.accesses' name), [])) wenzelm@22796: |> Sign.add_consts_i [(bname, predT, NoSyn)] wenzelm@13375: |> PureThy.add_defs_i false [((Thm.def_name bname, Logic.mk_equals (head, body)), [])]; wenzelm@20059: val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head; wenzelm@13394: wenzelm@16458: val cert = Thm.cterm_of defs_thy; wenzelm@13375: wenzelm@20059: val intro = Goal.prove_global defs_thy [] norm_ts statement (fn _ => wenzelm@21708: MetaSimplifier.rewrite_goals_tac [pred_def] THEN wenzelm@13375: Tactic.compose_tac (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN wenzelm@23418: Tactic.compose_tac (false, wenzelm@23418: Conjunction.intr_balanced (map (Thm.assume o cert) norm_ts), 0) 1); wenzelm@13375: wenzelm@13375: val conjuncts = wenzelm@19423: (Drule.equal_elim_rule2 OF [body_eq, wenzelm@21708: MetaSimplifier.rewrite_rule [pred_def] (Thm.assume (cert statement))]) wenzelm@23418: |> Conjunction.elim_balanced (length ts); haftmann@17257: val axioms = ts ~~ conjuncts |> map (fn (t, ax) => wenzelm@20059: Element.prove_witness defs_ctxt t wenzelm@21708: (MetaSimplifier.rewrite_goals_tac defs THEN wenzelm@13375: Tactic.compose_tac (false, ax, 0) 1)); haftmann@18550: in ((statement, intro, axioms), defs_thy) end; wenzelm@13375: haftmann@18550: fun assumes_to_notes (Assumes asms) axms = wenzelm@21441: fold_map (fn (a, spec) => fn axs => wenzelm@21441: let val (ps, qs) = chop (length spec) axs wenzelm@21441: in ((a, [(ps, [])]), qs) end) asms axms wenzelm@21441: |> apfst (curry Notes Thm.assumptionK) haftmann@18550: | assumes_to_notes e axms = (e, axms); wenzelm@13394: ballarin@19931: (* CB: the following two change only "new" elems, these have identifier ("", _). *) ballarin@19931: ballarin@19931: (* turn Assumes into Notes elements *) ballarin@15206: ballarin@19931: fun change_assumes_elemss axioms elemss = haftmann@18550: let wenzelm@21483: val satisfy = Element.morph_ctxt (Element.satisfy_morphism axioms); ballarin@19931: fun change (id as ("", _), es) = wenzelm@21483: fold_map assumes_to_notes (map satisfy es) haftmann@18550: #-> (fn es' => pair (id, es')) haftmann@18550: | change e = pair e; haftmann@18550: in wenzelm@19780: fst (fold_map change elemss (map Element.conclude_witness axioms)) haftmann@18550: end; wenzelm@13394: ballarin@19931: (* adjust hyps of Notes elements *) ballarin@19931: ballarin@19931: fun change_elemss_hyps axioms elemss = ballarin@19931: let wenzelm@21483: val satisfy = Element.morph_ctxt (Element.satisfy_morphism axioms); wenzelm@21483: fun change (id as ("", _), es) = (id, map (fn e as Notes _ => satisfy e | e => e) es) ballarin@19931: | change e = e; ballarin@19931: in map change elemss end; ballarin@19931: wenzelm@13394: in wenzelm@13375: ballarin@15104: (* CB: main predicate definition function *) ballarin@15104: haftmann@22351: fun define_preds pname (parms, ((exts, exts'), (ints, ints')), defs) elemss thy = wenzelm@13394: let haftmann@22351: val ((elemss', more_ts), a_elem, a_intro, thy'') = ballarin@19931: if null exts then ((elemss, []), [], [], thy) wenzelm@13394: else wenzelm@13394: let haftmann@22351: val aname = if null ints then pname else pname ^ "_" ^ axiomsN; haftmann@22351: val ((statement, intro, axioms), thy') = haftmann@22351: thy haftmann@22351: |> def_pred aname parms defs exts exts'; ballarin@19931: val elemss' = change_assumes_elemss axioms elemss; haftmann@22351: val a_elem = [(("", []), [Assumes [((pname ^ "_" ^ axiomsN, []), [(statement, [])])]])]; haftmann@22351: val (_, thy'') = haftmann@22351: thy' haftmann@22351: |> PureThy.note_thmss_qualified Thm.internalK aname [((introN, []), [([intro], [])])]; haftmann@22351: in ((elemss', [statement]), a_elem, [intro], thy'') end; haftmann@22351: val (predicate, stmt', elemss'', b_intro, thy'''') = haftmann@22351: if null ints then (([], []), more_ts, elemss' @ a_elem, [], thy'') wenzelm@13394: else wenzelm@13394: let haftmann@22351: val ((statement, intro, axioms), thy''') = haftmann@22351: thy'' haftmann@22351: |> def_pred pname parms defs (ints @ more_ts) (ints' @ more_ts); haftmann@22351: val cstatement = Thm.cterm_of thy''' statement; ballarin@19931: val elemss'' = change_elemss_hyps axioms elemss'; ballarin@19931: val b_elem = [(("", []), haftmann@22351: [Assumes [((pname ^ "_" ^ axiomsN, []), [(statement, [])])]])]; haftmann@22351: val (_, thy'''') = haftmann@22351: thy''' haftmann@22351: |> PureThy.note_thmss_qualified Thm.internalK pname haftmann@22351: [((introN, []), [([intro], [])]), haftmann@22351: ((axiomsN, []), [(map (Drule.standard o Element.conclude_witness) axioms, [])])]; haftmann@22351: in (([cstatement], axioms), [statement], elemss'' @ b_elem, [intro], thy'''') end; haftmann@22351: in (((elemss'', predicate, stmt'), (a_intro, b_intro)), thy'''') end; wenzelm@13375: wenzelm@13375: end; wenzelm@13336: wenzelm@13336: wenzelm@13297: (* add_locale(_i) *) wenzelm@13297: wenzelm@13297: local wenzelm@13297: ballarin@19931: (* turn Defines into Notes elements, accumulate definition terms *) ballarin@19931: ballarin@19942: fun defines_to_notes is_ext thy (Defines defs) defns = ballarin@19942: let ballarin@19942: val defs' = map (fn (_, (def, _)) => (("", []), (def, []))) defs ballarin@19942: val notes = map (fn (a, (def, _)) => ballarin@19942: (a, [([assume (cterm_of thy def)], [])])) defs wenzelm@21441: in wenzelm@21441: (if is_ext then SOME (Notes (Thm.definitionK, notes)) else NONE, defns @ [Defines defs']) wenzelm@21441: end ballarin@19942: | defines_to_notes _ _ e defns = (SOME e, defns); ballarin@19931: ballarin@19942: fun change_defines_elemss thy elemss defns = ballarin@19931: let ballarin@19942: fun change (id as (n, _), es) defns = ballarin@19931: let ballarin@19942: val (es', defns') = fold_map (defines_to_notes (n="") thy) es defns ballarin@19942: in ((id, map_filter I es'), defns') end ballarin@19942: in fold_map change elemss defns end; ballarin@19931: wenzelm@18343: fun gen_add_locale prep_ctxt prep_expr wenzelm@22573: predicate_name bname raw_imports raw_body thy = haftmann@22351: (* predicate_name: NONE - open locale without predicate haftmann@22351: SOME "" - locale with predicate named as locale haftmann@22351: SOME "name" - locale with predicate named "name" *) wenzelm@13297: let wenzelm@16458: val name = Sign.full_name thy bname; wenzelm@21962: val _ = is_some (get_locale thy name) andalso wenzelm@21962: error ("Duplicate definition of locale " ^ quote name); wenzelm@13297: wenzelm@13297: val thy_ctxt = ProofContext.init thy; ballarin@17228: val (((import_ctxt, import_elemss), (body_ctxt, body_elemss, syn)), ballarin@19931: text as (parms, ((_, exts'), _), defs)) = wenzelm@22573: prep_ctxt raw_imports raw_body thy_ctxt; ballarin@19931: val elemss = import_elemss @ body_elemss |> ballarin@19931: map_filter (fn ((id, Assumed axs), elems) => SOME (id, elems) | _ => NONE); wenzelm@22573: val imports = prep_expr thy raw_imports; wenzelm@13297: wenzelm@23178: val extraTs = List.foldr Term.add_term_tfrees [] exts' \\ wenzelm@23178: List.foldr Term.add_typ_tfrees [] (map snd parms); ballarin@17228: val _ = if null extraTs then () ballarin@17437: else warning ("Additional type variable(s) in locale specification " ^ quote bname); ballarin@17228: ballarin@19931: val ((((elemss', predicate as (predicate_statement, predicate_axioms), stmt'), intros), wenzelm@22573: pred_thy), (imports, regs)) = haftmann@22351: case predicate_name haftmann@22351: of SOME predicate_name => haftmann@22351: let haftmann@22351: val predicate_name' = case predicate_name of "" => bname | _ => predicate_name; haftmann@22351: val (elemss', defns) = change_defines_elemss thy elemss []; haftmann@22351: val elemss'' = elemss' @ [(("", []), defns)]; haftmann@22351: val (((elemss''', predicate as (_, axioms), stmt'), intros), thy') = haftmann@22351: define_preds predicate_name' text elemss'' thy; haftmann@22351: fun mk_regs elemss wits = haftmann@22351: fold_map (fn (id, elems) => fn wts => let haftmann@22351: val ts = List.concat (List.mapPartial (fn (Assumes asms) => haftmann@22351: SOME (List.concat (map (map #1 o #2) asms)) | _ => NONE) elems); haftmann@22351: val (wts1, wts2) = chop (length ts) wts haftmann@22351: in ((apsnd (map fst) id, wts1), wts2) end) elemss wits |> fst; haftmann@22351: val regs = mk_regs elemss''' axioms |> haftmann@22351: map_filter (fn (("", _), _) => NONE | e => SOME e); haftmann@22351: in ((((elemss''', predicate, stmt'), intros), thy'), (empty, regs)) end wenzelm@22573: | NONE => ((((elemss, ([], []), []), ([], [])), thy), (imports, [])); wenzelm@13420: wenzelm@18137: fun axiomify axioms elemss = ballarin@15206: (axioms, elemss) |> foldl_map (fn (axs, (id, elems)) => let wenzelm@19482: val ts = flat (map_filter (fn (Assumes asms) => wenzelm@19482: SOME (maps (map #1 o #2) asms) | _ => NONE) elems); wenzelm@19018: val (axs1, axs2) = chop (length ts) axs; ballarin@17000: in (axs2, ((id, Assumed axs1), elems)) end) ballarin@15206: |> snd; ballarin@19931: val (ctxt, (_, facts)) = activate_facts true (K I) wenzelm@20965: (ProofContext.init pred_thy, axiomify predicate_axioms elemss'); wenzelm@21393: val view_ctxt = Assumption.add_view thy_ctxt predicate_statement ctxt; wenzelm@21602: val export = Goal.close_result o Goal.norm_result o wenzelm@21602: singleton (ProofContext.export view_ctxt thy_ctxt); wenzelm@13420: val facts' = facts |> map (fn (a, ths) => ((a, []), [(map export ths, [])])); wenzelm@19482: val elems' = maps #2 (filter (equal "" o #1 o #1) elemss'); ballarin@19783: val elems'' = map_filter (fn (Fixes _) => NONE | e => SOME e) elems'; ballarin@19931: val axs' = map (Element.assume_witness pred_thy) stmt'; wenzelm@21393: val loc_ctxt = pred_thy wenzelm@21441: |> PureThy.note_thmss_qualified Thm.assumptionK bname facts' |> snd wenzelm@19061: |> declare_locale name wenzelm@19061: |> put_locale name ballarin@19931: {axiom = axs', wenzelm@22573: imports = imports, ballarin@19783: elems = map (fn e => (e, stamp ())) elems'', ballarin@19278: params = params_of elemss' |> map (fn (x, SOME T) => ((x, T), the (Symtab.lookup syn x))), ballarin@19931: lparams = map #1 (params_of' body_elemss), wenzelm@21665: decls = ([], []), ballarin@19931: regs = regs, wenzelm@21393: intros = intros} wenzelm@21393: |> init name; wenzelm@21393: in (name, loc_ctxt) end; wenzelm@13297: wenzelm@13297: in wenzelm@13297: haftmann@18917: val add_locale = gen_add_locale read_context intern_expr; haftmann@18917: val add_locale_i = gen_add_locale cert_context (K I); wenzelm@13297: wenzelm@13297: end; wenzelm@13297: wenzelm@19018: val _ = Context.add_setup haftmann@22351: (add_locale_i (SOME "") "var" empty [Fixes [(Name.internal "x", NONE, NoSyn)]] #> wenzelm@20965: snd #> ProofContext.theory_of #> haftmann@22351: add_locale_i (SOME "") "struct" empty [Fixes [(Name.internal "S", NONE, Structure)]] #> wenzelm@20965: snd #> ProofContext.theory_of); wenzelm@15801: wenzelm@13297: wenzelm@12730: wenzelm@17355: ballarin@19931: (** Normalisation of locale statements --- ballarin@19931: discharges goals implied by interpretations **) ballarin@19931: ballarin@19931: local ballarin@19931: ballarin@19931: fun locale_assm_intros thy = ballarin@19931: Symtab.fold (fn (_, {intros = (a, _), ...}) => fn intros => (a @ intros)) ballarin@19931: (#2 (GlobalLocalesData.get thy)) []; ballarin@19931: fun locale_base_intros thy = ballarin@19931: Symtab.fold (fn (_, {intros = (_, b), ...}) => fn intros => (b @ intros)) ballarin@19931: (#2 (GlobalLocalesData.get thy)) []; ballarin@19931: ballarin@19931: fun all_witnesses ctxt = ballarin@19931: let ballarin@19931: val thy = ProofContext.theory_of ctxt; ballarin@19931: fun get registrations = Symtab.fold (fn (_, regs) => fn thms => ballarin@22658: (Registrations.dest thy regs |> map (fn (_, (_, wits, _)) => ballarin@19931: map Element.conclude_witness wits) |> flat) @ thms) ballarin@19931: registrations []; ballarin@19931: val globals = get (#3 (GlobalLocalesData.get thy)); ballarin@19931: val locals = get (LocalLocalesData.get ctxt); ballarin@19931: in globals @ locals end; ballarin@19931: (* FIXME: proper varification *) ballarin@19931: ballarin@19931: in ballarin@19931: ballarin@19984: fun intro_locales_tac eager ctxt facts st = ballarin@19931: let ballarin@19931: val wits = all_witnesses ctxt |> map Thm.varifyT; ballarin@19931: val thy = ProofContext.theory_of ctxt; ballarin@19931: val intros = locale_base_intros thy @ (if eager then locale_assm_intros thy else []); ballarin@19931: in ballarin@19931: (ALLGOALS (Method.insert_tac facts THEN' ballarin@19931: REPEAT_ALL_NEW (resolve_tac (wits @ intros))) ballarin@19931: THEN Tactic.distinct_subgoals_tac) st ballarin@19931: end; ballarin@19931: ballarin@19931: val _ = Context.add_setup (Method.add_methods ballarin@19931: [("intro_locales", ballarin@19984: Method.ctxt_args (fn ctxt => Method.METHOD (intro_locales_tac false ctxt)), ballarin@19984: "back-chain introduction rules of locales without unfolding predicates"), ballarin@19984: ("unfold_locales", ballarin@19984: Method.ctxt_args (fn ctxt => Method.METHOD (intro_locales_tac true ctxt)), ballarin@19984: "back-chain all introduction rules of locales")]); ballarin@19931: ballarin@19931: end; ballarin@19931: wenzelm@19780: ballarin@15598: (** Interpretation commands **) ballarin@15596: ballarin@15596: local ballarin@15596: ballarin@15596: (* extract proof obligations (assms and defs) from elements *) ballarin@15596: wenzelm@19780: fun extract_asms_elems ((id, Assumed _), elems) = (id, maps Element.prems_of elems) ballarin@17138: | extract_asms_elems ((id, Derived _), _) = (id, []); ballarin@15596: ballarin@15596: ballarin@15624: (* activate instantiated facts in theory or context *) ballarin@15596: ballarin@22658: structure Idtab = ballarin@22658: TableFun(type key = string * term list ballarin@22658: val ord = prod_ord string_ord (list_ord Term.fast_term_ord)); ballarin@22658: ballarin@22658: fun gen_activate_facts_elemss mk_ctxt get_reg note attrib std put_reg add_wit add_eqn wenzelm@18123: attn all_elemss new_elemss propss thmss thy_ctxt = wenzelm@21499: let ballarin@22658: val ctxt = mk_ctxt thy_ctxt; ballarin@22658: val (propss, eq_props) = chop (length new_elemss) propss; ballarin@22658: val (thmss, eq_thms) = chop (length new_elemss) thmss; ballarin@22658: ballarin@22658: fun activate_elem eqns disch ((fully_qualified, prfx), atts) (Notes (kind, facts)) thy_ctxt = wenzelm@21499: let ballarin@22658: val ctxt = mk_ctxt thy_ctxt; wenzelm@22241: val fact_morphism = haftmann@22300: Morphism.name_morphism (NameSpace.qualified prfx) ballarin@22756: $> Morphism.term_morphism (MetaSimplifier.rewrite_term ballarin@22756: (ProofContext.theory_of ctxt) eqns []) ballarin@22756: $> Morphism.thm_morphism (disch #> MetaSimplifier.rewrite_rule eqns); wenzelm@21499: val facts' = facts wenzelm@21499: (* discharge hyps in attributes *) haftmann@22300: |> Attrib.map_facts haftmann@22300: (attrib thy_ctxt o Args.morph_values fact_morphism) wenzelm@21499: (* append interpretation attributes *) wenzelm@21499: |> map (apfst (apsnd (fn a => a @ map (attrib thy_ctxt) atts))) wenzelm@21499: (* discharge hyps *) ballarin@22658: |> map (apsnd (map (apfst (map disch)))) ballarin@22658: (* unfold eqns *) ballarin@22756: |> map (apsnd (map (apfst (map (MetaSimplifier.rewrite_rule eqns))))) haftmann@22351: in snd (note kind (fully_qualified, prfx) facts' thy_ctxt) end ballarin@22658: | activate_elem _ _ _ _ thy_ctxt = thy_ctxt; ballarin@22658: ballarin@22658: fun activate_elems eqns disch ((id, _), elems) thy_ctxt = wenzelm@21499: let ballarin@22658: val (prfx_atts, _, _) = case get_reg thy_ctxt id haftmann@22351: of SOME x => x haftmann@22351: | NONE => sys_error ("Unknown registration of " ^ quote (fst id) haftmann@22351: ^ " while activating facts."); ballarin@22658: in fold (activate_elem (the (Idtab.lookup eqns id)) disch prfx_atts) elems thy_ctxt end; ballarin@17033: wenzelm@21499: val thy_ctxt' = thy_ctxt wenzelm@21499: (* add registrations *) wenzelm@21499: |> fold (fn ((id, _), _) => put_reg id attn) new_elemss ballarin@22658: (* add witnesses of Assumed elements (only those generate proof obligations) *) ballarin@22658: |> fold (fn (id, thms) => fold (add_wit id) thms) (map fst propss ~~ thmss) ballarin@22658: (* add equations *) ballarin@22658: |> fold (fn (id, thms) => fold (add_eqn id) thms) (map fst eq_props ~~ eq_thms); ballarin@15596: wenzelm@21499: val prems = flat (map_filter ballarin@22658: (fn ((id, Assumed _), _) => Option.map #2 (get_reg thy_ctxt' id) wenzelm@21499: | ((_, Derived _), _) => NONE) all_elemss); wenzelm@21499: val satisfy = Element.satisfy_morphism prems; wenzelm@21499: val thy_ctxt'' = thy_ctxt' wenzelm@21499: (* add witnesses of Derived elements *) wenzelm@21499: |> fold (fn (id, thms) => fold (add_wit id o Element.morph_witness satisfy) thms) wenzelm@21499: (map_filter (fn ((_, Assumed _), _) => NONE wenzelm@21499: | ((id, Derived thms), _) => SOME (id, thms)) new_elemss) wenzelm@19780: ballarin@22658: (* Accumulate equations *) ballarin@22658: fun add_eqns ((id, _), _) eqns = ballarin@22658: let ballarin@22658: val eqns' = case get_reg thy_ctxt'' id ballarin@22658: of NONE => eqns ballarin@22658: | SOME (_, _, eqns') => Termtab.join (fn t => fn (_, e) => e) (eqns, eqns') wenzelm@23655: (* handle Termtab.DUP t => wenzelm@23655: error (implode ("Conflicting equations for term " :: wenzelm@23655: quote (ProofContext.string_of_term ctxt t))) ballarin@22658: *) ballarin@22658: in ((id, eqns'), eqns') end; ballarin@22658: val all_eqns = fold_map add_eqns all_elemss Termtab.empty |> fst ballarin@22658: |> map (apsnd (map (Element.conclude_witness o snd) o Termtab.dest)) ballarin@22658: |> (fn xs => fold Idtab.update xs Idtab.empty) ballarin@22658: (* Idtab.make doesn't work: can have conflicting duplicates, ballarin@22658: because instantiation may equate ids and equations are accumulated! *) ballarin@22658: ballarin@22658: (* TODO: can use dest_witness here? (more efficient) *) ballarin@22658: wenzelm@21499: val disch' = std o Morphism.thm satisfy; (* FIXME *) wenzelm@21499: in wenzelm@21499: thy_ctxt'' wenzelm@21499: (* add facts to theory or context *) ballarin@22658: |> fold (activate_elems all_eqns disch') new_elemss wenzelm@21499: end; ballarin@15596: wenzelm@17355: fun global_activate_facts_elemss x = gen_activate_facts_elemss ballarin@22658: ProofContext.init ballarin@15696: (fn thy => fn (name, ps) => ballarin@15696: get_global_registration thy (name, map Logic.varify ps)) wenzelm@21441: global_note_prefix_i wenzelm@18728: Attrib.attribute_i Drule.standard ballarin@17033: (fn (name, ps) => put_global_registration (name, map Logic.varify ps)) wenzelm@19780: (fn (n, ps) => add_global_witness (n, map Logic.varify ps) o ballarin@19931: Element.map_witness (fn (t, th) => (Logic.legacy_unvarify t, legacy_unvarify th)) ballarin@22658: (* FIXME *)) ballarin@22658: (fn (n, ps) => add_global_equation (n, map Logic.varify ps)) ballarin@22658: x; wenzelm@17355: wenzelm@17355: fun local_activate_facts_elemss x = gen_activate_facts_elemss ballarin@22658: I ballarin@15696: get_local_registration wenzelm@19018: local_note_prefix_i wenzelm@18728: (Attrib.attribute_i o ProofContext.theory_of) I ballarin@17033: put_local_registration ballarin@22658: add_local_witness ballarin@22658: add_local_equation ballarin@22658: x; ballarin@22658: ballarin@22756: fun read_instantiations ctxt parms (insts, eqns) = ballarin@15596: let wenzelm@22772: val thy = ProofContext.theory_of ctxt; wenzelm@22772: wenzelm@22772: (* parameter instantiations *) wenzelm@22772: val parms' = map (apsnd Logic.varifyT) parms; wenzelm@22772: val d = length parms - length insts; wenzelm@22772: val insts = wenzelm@22772: if d < 0 then error "More arguments than parameters in instantiation." wenzelm@22772: else insts @ replicate d NONE; wenzelm@22772: wenzelm@22772: val given = (parms' ~~ insts) |> map_filter wenzelm@22772: (fn (_, NONE) => NONE wenzelm@22772: | ((x, T), SOME inst) => SOME (x, (inst, T))); ballarin@22658: val (given_ps, given_insts) = split_list given; ballarin@22658: wenzelm@22772: (* equations *) ballarin@22658: val (lefts, rights) = split_list eqns; ballarin@22658: val max_eqs = length eqns; wenzelm@22772: val eqTs = map (fn i => TypeInfer.param i ("'a", [])) (0 upto max_eqs - 1); wenzelm@22772: wenzelm@22772: (* read given insts / eqns *) wenzelm@22772: val all_vs = ProofContext.read_termTs ctxt (given_insts @ (lefts ~~ eqTs) @ (rights ~~ eqTs)); wenzelm@22772: val ctxt' = ctxt |> fold Variable.declare_term all_vs; wenzelm@22772: val (vs, (lefts', rights')) = all_vs |> chop (length given_insts) ||> chop max_eqs; wenzelm@22772: wenzelm@22772: (* infer parameter types *) wenzelm@22772: val tyenv = fold (fn ((_, T), t) => Sign.typ_match thy (T, Term.fastype_of t)) wenzelm@22772: (given_insts ~~ vs) Vartab.empty; wenzelm@22772: val looseTs = fold (Term.add_tvarsT o Envir.typ_subst_TVars tyenv o #2) parms' []; wenzelm@22772: val (fixedTs, _) = Variable.invent_types (map #2 looseTs) ctxt'; wenzelm@22772: val tyenv' = wenzelm@22772: fold (fn ((xi, S), v) => Vartab.update_new (xi, (S, TFree v))) (looseTs ~~ fixedTs) tyenv; wenzelm@22772: wenzelm@22772: (*results*) wenzelm@22772: val instT = Vartab.fold (fn ((a, 0), (S, T)) => wenzelm@22772: if T = TFree (a, S) then I else Symtab.update (a, T)) tyenv' Symtab.empty; wenzelm@22772: val inst = Symtab.make (given_ps ~~ vs); ballarin@22658: in ((instT, inst), lefts' ~~ rights') end; ballarin@22658: ballarin@15598: wenzelm@22772: fun gen_prep_registration mk_ctxt test_reg activate ballarin@22658: prep_attr prep_expr prep_insts haftmann@22300: thy_ctxt raw_attn raw_expr raw_insts = haftmann@22300: let haftmann@22300: val ctxt = mk_ctxt thy_ctxt; haftmann@22300: val thy = ProofContext.theory_of ctxt; haftmann@22300: val ctxt' = ProofContext.init thy; haftmann@22300: haftmann@22300: val attn = (apsnd o map) haftmann@22300: (Attrib.crude_closure ctxt o Args.assignable o prep_attr thy) raw_attn; haftmann@22300: val expr = prep_expr thy raw_expr; ballarin@15596: haftmann@22300: val pts = params_of_expr ctxt' [] expr ([], Symtab.empty, Symtab.empty); haftmann@22300: val params_ids = make_params_ids (#1 pts); haftmann@22300: val raw_params_elemss = make_raw_params_elemss pts; haftmann@22300: val ((ids, _), raw_elemss) = flatten (ctxt', I) (([], Symtab.empty), Expr expr); haftmann@22300: val ((parms, all_elemss, _), (_, (_, defs, _))) = haftmann@22300: read_elemss false ctxt' [] (raw_params_elemss @ raw_elemss) []; haftmann@22300: haftmann@22300: (** compute instantiation **) haftmann@22300: ballarin@22658: (* consistency check: equations need to be stored in a particular locale, ballarin@22658: therefore if equations are present locale expression must be a name *) ballarin@22658: ballarin@22658: val _ = case (expr, snd raw_insts) of ballarin@22658: (Locale _, _) => () | (_, []) => () ballarin@22658: | (_, _) => error "Interpretations with `where' only permitted if locale expression is a name."; ballarin@22658: ballarin@22658: (* read or certify instantiation *) ballarin@22756: val ((instT, inst1), eqns') = prep_insts ctxt parms raw_insts; haftmann@22300: haftmann@22300: (* defined params without given instantiation *) haftmann@22300: val not_given = filter_out (Symtab.defined inst1 o fst) parms; wenzelm@18137: fun add_def (p, pT) inst = ballarin@15596: let ballarin@15596: val (t, T) = case find_first (fn (Free (a, _), _) => a = p) defs of ballarin@15596: NONE => error ("Instance missing for parameter " ^ quote p) ballarin@15596: | SOME (Free (_, T), t) => (t, T); haftmann@22300: val d = Element.inst_term (instT, inst) t; wenzelm@18137: in Symtab.update_new (p, d) inst end; haftmann@22300: val inst2 = fold add_def not_given inst1; haftmann@22300: val inst_morphism = Element.inst_morphism thy (instT, inst2); wenzelm@18137: (* Note: insts contain no vars. *) ballarin@15596: ballarin@15596: (** compute proof obligations **) ballarin@15596: ballarin@15598: (* restore "small" ids *) ballarin@17000: val ids' = map (fn ((n, ps), (_, mode)) => ballarin@19783: ((n, map (fn p => Free (p, (the o AList.lookup (op =) parms) p)) ps), mode)) ballarin@19931: ids; ballarin@19931: val (_, all_elemss') = chop (length raw_params_elemss) all_elemss ballarin@15596: (* instantiate ids and elements *) ballarin@19931: val inst_elemss = (ids' ~~ all_elemss') |> map (fn (((n, ps), _), ((_, mode), elems)) => wenzelm@21483: ((n, map (Morphism.term inst_morphism) ps), wenzelm@21483: map (fn Int e => Element.morph_ctxt inst_morphism e) elems) wenzelm@21483: |> apfst (fn id => (id, map_mode (map (Element.morph_witness inst_morphism)) mode))); ballarin@15596: ballarin@15624: (* remove fragments already registered with theory or context *) ballarin@22658: val new_inst_elemss = filter_out (fn ((id, _), _) => ballarin@22658: test_reg thy_ctxt id) inst_elemss; ballarin@22658: (* val new_ids = map #1 new_inst_elemss; *) ballarin@22658: ballarin@22658: (* equations *) ballarin@22658: val eqn_elems = if null eqns' then [] else [(Library.last_elem inst_elemss |> fst |> fst, map Logic.mk_equals eqns')]; ballarin@22658: ballarin@22658: val propss = map extract_asms_elems new_inst_elemss @ eqn_elems; ballarin@15596: haftmann@22300: in (propss, activate attn inst_elemss new_inst_elemss propss) end; ballarin@15596: ballarin@22756: fun gen_prep_global_registration mk_ctxt = gen_prep_registration ProofContext.init haftmann@22300: (fn thy => fn (name, ps) => test_global_registration thy (name, map Logic.varify ps)) haftmann@22300: global_activate_facts_elemss mk_ctxt; ballarin@15624: ballarin@22756: fun gen_prep_local_registration mk_ctxt = gen_prep_registration I haftmann@22300: smart_test_registration haftmann@22300: local_activate_facts_elemss mk_ctxt; ballarin@15624: haftmann@22300: val prep_global_registration = gen_prep_global_registration haftmann@22300: Attrib.intern_src intern_expr read_instantiations; haftmann@22300: val prep_global_registration_i = gen_prep_global_registration ballarin@22756: (K I) (K I) ((K o K) I); haftmann@22300: haftmann@22300: val prep_local_registration = gen_prep_local_registration haftmann@22300: Attrib.intern_src intern_expr read_instantiations; haftmann@22300: val prep_local_registration_i = gen_prep_local_registration ballarin@22756: (K I) (K I) ((K o K) I); ballarin@15596: ballarin@17000: fun prep_registration_in_locale target expr thy = ballarin@17000: (* target already in internal form *) ballarin@17000: let ballarin@17000: val ctxt = ProofContext.init thy; ballarin@17138: val ((raw_target_ids, target_syn), _) = flatten (ctxt, I) ballarin@17000: (([], Symtab.empty), Expr (Locale target)); ballarin@19278: val fixed = the_locale thy target |> #params |> map #1; ballarin@17000: val ((all_ids, syn), raw_elemss) = flatten (ctxt, intern_expr thy) ballarin@17138: ((raw_target_ids, target_syn), Expr expr); wenzelm@19018: val (target_ids, ids) = chop (length raw_target_ids) all_ids; ballarin@17138: val ((parms, elemss, _), _) = read_elemss false ctxt fixed raw_elemss []; ballarin@17000: ballarin@17000: (** compute proof obligations **) ballarin@17000: ballarin@17000: (* restore "small" ids, with mode *) ballarin@17000: val ids' = map (apsnd snd) ids; ballarin@17000: (* remove Int markers *) ballarin@17000: val elemss' = map (fn (_, es) => ballarin@17000: map (fn Int e => e) es) elemss ballarin@17000: (* extract assumptions and defs *) ballarin@17138: val ids_elemss = ids' ~~ elemss'; wenzelm@19780: val propss = map extract_asms_elems ids_elemss; ballarin@17000: ballarin@17138: (** activation function: ballarin@17138: - add registrations to the target locale ballarin@17138: - add induced registrations for all global registrations of ballarin@17138: the target, unless already present ballarin@17138: - add facts of induced registrations to theory **) ballarin@17138: ballarin@22658: (* val t_ids = map_filter ballarin@22658: (fn (id, (_, Assumed _)) => SOME id | _ => NONE) target_ids; *) ballarin@17138: wenzelm@18123: fun activate thmss thy = let wenzelm@19780: val satisfy = Element.satisfy_thm (flat thmss); wenzelm@18123: val ids_elemss_thmss = ids_elemss ~~ thmss; ballarin@17138: val regs = get_global_registrations thy target; ballarin@17138: ballarin@17138: fun activate_id (((id, Assumed _), _), thms) thy = ballarin@17033: thy |> put_registration_in_locale target id ballarin@17138: |> fold (add_witness_in_locale target id) thms ballarin@17138: | activate_id _ thy = thy; ballarin@17138: ballarin@22658: fun activate_reg (vts, (((fully_qualified, prfx), atts2), _, _)) thy = wenzelm@21483: let ballarin@22658: val (insts, wits, _) = collect_global_witnesses thy fixed target_ids vts; wenzelm@18137: fun inst_parms ps = map haftmann@17485: (the o AList.lookup (op =) (map #1 fixed ~~ vts)) ps; wenzelm@19780: val disch = Element.satisfy_thm wits; wenzelm@19482: val new_elemss = filter (fn (((name, ps), _), _) => ballarin@17138: not (test_global_registration thy (name, inst_parms ps))) (ids_elemss); ballarin@17138: fun activate_assumed_id (((_, Derived _), _), _) thy = thy ballarin@17138: | activate_assumed_id ((((name, ps), Assumed _), _), thms) thy = let ballarin@17138: val ps' = inst_parms ps; ballarin@17138: in ballarin@17138: if test_global_registration thy (name, ps') ballarin@17138: then thy ballarin@17138: else thy haftmann@22351: |> put_global_registration (name, ps') ((fully_qualified, prfx), atts2) wenzelm@19780: |> fold (fn witn => fn thy => add_global_witness (name, ps') wenzelm@21483: (Element.morph_witness (Element.inst_morphism thy insts) witn) thy) thms ballarin@17138: end; ballarin@17138: ballarin@17138: fun activate_derived_id ((_, Assumed _), _) thy = thy ballarin@17138: | activate_derived_id (((name, ps), Derived ths), _) thy = let ballarin@17138: val ps' = inst_parms ps; ballarin@17138: in ballarin@17138: if test_global_registration thy (name, ps') ballarin@17138: then thy ballarin@17138: else thy haftmann@22351: |> put_global_registration (name, ps') ((fully_qualified, prfx), atts2) wenzelm@19780: |> fold (fn witn => fn thy => add_global_witness (name, ps') wenzelm@19780: (witn |> Element.map_witness (fn (t, th) => (* FIXME *) wenzelm@18137: (Element.inst_term insts t, wenzelm@19780: disch (Element.inst_thm thy insts (satisfy th))))) thy) ths ballarin@17138: end; ballarin@17138: wenzelm@21441: fun activate_elem (Notes (kind, facts)) thy = ballarin@17138: let wenzelm@21523: val att_morphism = wenzelm@22241: Morphism.name_morphism (NameSpace.qualified prfx) $> wenzelm@21523: Morphism.thm_morphism satisfy $> wenzelm@21523: Element.inst_morphism thy insts $> wenzelm@21523: Morphism.thm_morphism disch; ballarin@17138: val facts' = facts wenzelm@21523: |> Attrib.map_facts (Attrib.attribute_i thy o Args.morph_values att_morphism) wenzelm@21523: |> map (apfst (apsnd (fn a => a @ map (Attrib.attribute thy) atts2))) wenzelm@21523: |> map (apsnd (map (apfst (map (disch o Element.inst_thm thy insts o satisfy))))) ballarin@17138: in haftmann@18377: thy haftmann@22351: |> global_note_prefix_i kind (fully_qualified, prfx) facts' haftmann@18377: |> snd ballarin@17138: end ballarin@17138: | activate_elem _ thy = thy; ballarin@17138: ballarin@17138: fun activate_elems (_, elems) thy = fold activate_elem elems thy; ballarin@17138: ballarin@17138: in thy |> fold activate_assumed_id ids_elemss_thmss ballarin@17138: |> fold activate_derived_id ids_elemss ballarin@17138: |> fold activate_elems new_elemss end; ballarin@17033: in ballarin@17138: thy |> fold activate_id ids_elemss_thmss ballarin@17138: |> fold activate_reg regs ballarin@17033: end; ballarin@17000: ballarin@17033: in (propss, activate) end; ballarin@17000: wenzelm@21005: fun prep_propp propss = propss |> map (fn (_, props) => wenzelm@21361: map (rpair [] o Element.mark_witness) props); wenzelm@18123: wenzelm@18123: fun prep_result propps thmss = wenzelm@19780: ListPair.map (fn ((_, props), thms) => map2 Element.make_witness props thms) (propps, thmss); ballarin@17437: haftmann@22300: fun gen_interpretation prep_registration after_qed (prfx, raw_atts) raw_expr raw_insts thy = haftmann@22351: (* prfx = (flag indicating full qualification, name prefix) *) wenzelm@17355: let haftmann@22300: val (propss, activate) = prep_registration thy (prfx, raw_atts) raw_expr raw_insts; wenzelm@20366: fun after_qed' results = wenzelm@20366: ProofContext.theory (activate (prep_result propss results)) wenzelm@20366: #> after_qed; wenzelm@18123: in haftmann@22300: thy wenzelm@23418: |> ProofContext.init wenzelm@21441: |> Proof.theorem_i NONE after_qed' (prep_propp propss) haftmann@22300: |> Element.refine_witness haftmann@22300: |> Seq.hd haftmann@22300: end; haftmann@22300: haftmann@22300: fun gen_interpret prep_registration after_qed (prfx, atts) expr insts int state = haftmann@22351: (* prfx = (flag indicating full qualification, name prefix) *) haftmann@22300: let haftmann@22300: val _ = Proof.assert_forward_or_chain state; haftmann@22300: val ctxt = Proof.context_of state; haftmann@22300: val (propss, activate) = prep_registration ctxt (prfx, atts) expr insts; haftmann@22300: fun after_qed' results = haftmann@22300: Proof.map_context (K (ctxt |> activate (prep_result propss results))) haftmann@22300: #> Proof.put_facts NONE haftmann@22300: #> after_qed; haftmann@22300: in haftmann@22300: state haftmann@22300: |> Proof.local_goal (ProofDisplay.print_results int) (K I) ProofContext.bind_propp_i haftmann@22300: "interpret" NONE after_qed' (map (pair ("", [])) (prep_propp propss)) wenzelm@19810: |> Element.refine_witness |> Seq.hd wenzelm@18123: end; wenzelm@17355: haftmann@22300: in haftmann@22300: haftmann@22351: val interpretation_i = gen_interpretation prep_global_registration_i; haftmann@22300: val interpretation = gen_interpretation prep_global_registration; haftmann@22351: haftmann@22300: wenzelm@20366: fun interpretation_in_locale after_qed (raw_target, expr) thy = wenzelm@17355: let wenzelm@17355: val target = intern thy raw_target; wenzelm@17355: val (propss, activate) = prep_registration_in_locale target expr thy; wenzelm@21361: val raw_propp = prep_propp propss; wenzelm@21005: wenzelm@21005: val (_, _, goal_ctxt, propp) = thy wenzelm@21665: |> ProofContext.init wenzelm@21361: |> cert_context_statement (SOME target) [] raw_propp; wenzelm@21005: wenzelm@21005: fun after_qed' results = wenzelm@20366: ProofContext.theory (activate (prep_result propss results)) wenzelm@20366: #> after_qed; wenzelm@18123: in wenzelm@21005: goal_ctxt wenzelm@21441: |> Proof.theorem_i NONE after_qed' propp wenzelm@19810: |> Element.refine_witness |> Seq.hd wenzelm@18123: end; wenzelm@17449: haftmann@22351: val interpret_i = gen_interpret prep_local_registration_i; haftmann@22300: val interpret = gen_interpret prep_local_registration; ballarin@15596: wenzelm@11896: end; wenzelm@17355: wenzelm@17355: end;