wenzelm@18744: (* Title: Pure/Isar/local_theory.ML wenzelm@18744: ID: $Id$ wenzelm@18744: Author: Makarius wenzelm@18744: wenzelm@20888: Local theory operations, with abstract target context. wenzelm@18744: *) wenzelm@18744: wenzelm@18951: type local_theory = Proof.context; wenzelm@18951: wenzelm@18744: signature LOCAL_THEORY = wenzelm@18744: sig wenzelm@20888: type operations wenzelm@20888: val target_of: local_theory -> Proof.context wenzelm@20960: val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory wenzelm@20960: val raw_theory: (theory -> theory) -> local_theory -> local_theory wenzelm@22240: val full_naming: local_theory -> NameSpace.naming wenzelm@21614: val full_name: local_theory -> bstring -> string wenzelm@20960: val theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory wenzelm@18951: val theory: (theory -> theory) -> local_theory -> local_theory wenzelm@20960: val target_result: (Proof.context -> 'a * Proof.context) -> local_theory -> 'a * local_theory wenzelm@20888: val target: (Proof.context -> Proof.context) -> local_theory -> local_theory wenzelm@21860: val affirm: local_theory -> local_theory wenzelm@20960: val pretty: local_theory -> Pretty.T list wenzelm@20888: val consts: (string * typ -> bool) -> wenzelm@20888: ((bstring * typ) * mixfix) list -> local_theory -> (term * thm) list * local_theory wenzelm@21433: val axioms: string -> ((bstring * Attrib.src list) * term list) list -> local_theory -> wenzelm@18951: (bstring * thm list) list * local_theory wenzelm@21802: val abbrev: Syntax.mode -> (bstring * mixfix) * term -> local_theory -> wenzelm@21802: (term * term) * local_theory wenzelm@21433: val defs: string -> ((bstring * mixfix) * ((bstring * Attrib.src list) * term)) list -> wenzelm@20888: local_theory -> (term * (bstring * thm)) list * local_theory wenzelm@21433: val notes: string -> ((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list -> wenzelm@18951: local_theory -> (bstring * thm list) list * local_theory wenzelm@21498: val type_syntax: (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory wenzelm@21498: val term_syntax: (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory wenzelm@21498: val declaration: (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory wenzelm@21433: val def: string -> (bstring * mixfix) * ((bstring * Attrib.src list) * term) -> wenzelm@18951: local_theory -> (term * (bstring * thm)) * local_theory wenzelm@21433: val note: string -> (bstring * Attrib.src list) * thm list -> wenzelm@20888: local_theory -> (bstring * thm list) * Proof.context wenzelm@21743: val notation: Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory wenzelm@21529: val target_morphism: local_theory -> morphism wenzelm@22203: val target_naming: local_theory -> NameSpace.naming wenzelm@21529: val target_name: local_theory -> bstring -> string wenzelm@21529: val init: string -> operations -> Proof.context -> local_theory wenzelm@21273: val restore: local_theory -> local_theory wenzelm@21292: val reinit: local_theory -> theory -> local_theory wenzelm@21292: val exit: local_theory -> Proof.context wenzelm@18744: end; wenzelm@18744: wenzelm@18744: structure LocalTheory: LOCAL_THEORY = wenzelm@18744: struct wenzelm@18744: wenzelm@20888: (** local theory data **) wenzelm@19059: wenzelm@20888: (* type lthy *) wenzelm@20888: wenzelm@20888: type operations = wenzelm@20960: {pretty: local_theory -> Pretty.T list, wenzelm@20888: consts: (string * typ -> bool) -> ((bstring * typ) * mixfix) list -> local_theory -> wenzelm@20888: (term * thm) list * local_theory, wenzelm@21433: axioms: string -> ((bstring * Attrib.src list) * term list) list -> local_theory -> wenzelm@20888: (bstring * thm list) list * local_theory, wenzelm@21802: abbrev: Syntax.mode -> (bstring * mixfix) * term -> local_theory -> (term * term) * local_theory, wenzelm@21433: defs: string -> ((bstring * mixfix) * ((bstring * Attrib.src list) * term)) list -> wenzelm@21433: local_theory -> (term * (bstring * thm)) list * local_theory, wenzelm@21433: notes: string -> wenzelm@21433: ((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list -> wenzelm@21433: local_theory -> (bstring * thm list) list * local_theory, wenzelm@21498: type_syntax: (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory, wenzelm@21498: term_syntax: (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory, wenzelm@21498: declaration: (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory, wenzelm@21529: target_morphism: local_theory -> morphism, wenzelm@22203: target_naming: local_theory -> NameSpace.naming, wenzelm@21292: reinit: local_theory -> theory -> local_theory, wenzelm@21292: exit: local_theory -> Proof.context}; wenzelm@20888: wenzelm@20888: datatype lthy = LThy of wenzelm@21529: {theory_prefix: string, wenzelm@20888: operations: operations, wenzelm@20888: target: Proof.context}; wenzelm@20888: wenzelm@20888: fun make_lthy (theory_prefix, operations, target) = wenzelm@20888: LThy {theory_prefix = theory_prefix, operations = operations, target = target}; wenzelm@20888: wenzelm@20888: wenzelm@20888: (* context data *) wenzelm@18744: wenzelm@18744: structure Data = ProofDataFun wenzelm@18744: ( wenzelm@20888: type T = lthy option; wenzelm@20888: fun init _ = NONE; wenzelm@18744: ); wenzelm@18744: wenzelm@20888: fun get_lthy lthy = wenzelm@20888: (case Data.get lthy of wenzelm@20888: NONE => error "No local theory context" wenzelm@20888: | SOME (LThy data) => data); wenzelm@18951: wenzelm@20888: fun map_lthy f lthy = wenzelm@20888: let val {theory_prefix, operations, target} = get_lthy lthy wenzelm@20888: in Data.put (SOME (make_lthy (f (theory_prefix, operations, target)))) lthy end; wenzelm@18767: wenzelm@20888: val target_of = #target o get_lthy; wenzelm@21860: val affirm = tap target_of; wenzelm@18767: wenzelm@20888: fun map_target f = map_lthy (fn (theory_prefix, operations, target) => wenzelm@20888: (theory_prefix, operations, f target)); wenzelm@18767: wenzelm@18767: wenzelm@20888: (* substructure mappings *) wenzelm@19661: wenzelm@20960: fun raw_theory_result f lthy = wenzelm@20960: let wenzelm@20960: val (res, thy') = f (ProofContext.theory_of lthy); wenzelm@20960: val lthy' = lthy wenzelm@20960: |> map_target (ProofContext.transfer thy') wenzelm@20960: |> ProofContext.transfer thy'; wenzelm@20960: in (res, lthy') end; wenzelm@20960: wenzelm@20960: fun raw_theory f = #2 o raw_theory_result (f #> pair ()); wenzelm@20960: wenzelm@22240: fun full_naming lthy = wenzelm@22240: Sign.naming_of (ProofContext.theory_of lthy) wenzelm@22240: |> NameSpace.sticky_prefix (#theory_prefix (get_lthy lthy)) wenzelm@22240: |> NameSpace.qualified_names; wenzelm@22240: wenzelm@22240: val full_name = NameSpace.full o full_naming; wenzelm@21614: wenzelm@21529: fun theory_result f lthy = lthy |> raw_theory_result (fn thy => thy wenzelm@21529: |> Sign.sticky_prefix (#theory_prefix (get_lthy lthy)) wenzelm@21529: |> Sign.qualified_names wenzelm@21529: |> f wenzelm@21529: ||> Sign.restore_naming thy); wenzelm@19661: wenzelm@20888: fun theory f = #2 o theory_result (f #> pair ()); wenzelm@19661: wenzelm@20888: fun target_result f lthy = wenzelm@18744: let wenzelm@20888: val (res, ctxt') = f (#target (get_lthy lthy)); wenzelm@20888: val thy' = ProofContext.theory_of ctxt'; wenzelm@20888: val lthy' = lthy wenzelm@20888: |> map_target (K ctxt') wenzelm@20888: |> ProofContext.transfer thy'; wenzelm@20888: in (res, lthy') end; wenzelm@18876: wenzelm@20888: fun target f = #2 o target_result (f #> pair ()); wenzelm@18767: wenzelm@18767: wenzelm@21664: (* basic operations *) wenzelm@18781: wenzelm@20888: fun operation f lthy = f (#operations (get_lthy lthy)) lthy; wenzelm@20888: fun operation1 f x = operation (fn ops => f ops x); wenzelm@20888: fun operation2 f x y = operation (fn ops => f ops x y); wenzelm@18781: wenzelm@20888: val pretty = operation #pretty; wenzelm@20888: val consts = operation2 #consts; wenzelm@21433: val axioms = operation2 #axioms; wenzelm@21743: val abbrev = operation2 #abbrev; wenzelm@21433: val defs = operation2 #defs; wenzelm@21433: val notes = operation2 #notes; wenzelm@21498: val type_syntax = operation1 #type_syntax; wenzelm@20888: val term_syntax = operation1 #term_syntax; wenzelm@20888: val declaration = operation1 #declaration; wenzelm@21529: val target_morphism = operation #target_morphism; wenzelm@22203: val target_naming = operation #target_naming; wenzelm@18781: wenzelm@21433: fun def kind arg lthy = lthy |> defs kind [arg] |>> hd; wenzelm@21433: fun note kind (a, ths) lthy = lthy |> notes kind [(a, [(ths, [])])] |>> hd; wenzelm@18781: wenzelm@21743: fun notation mode raw_args lthy = wenzelm@21743: let val args = map (apfst (Morphism.term (target_morphism lthy))) raw_args wenzelm@21743: in term_syntax (ProofContext.target_notation mode args) lthy end; wenzelm@21743: wenzelm@22203: val target_name = NameSpace.full o target_naming; wenzelm@22203: wenzelm@21664: wenzelm@20910: (* init -- exit *) wenzelm@20910: wenzelm@20910: fun init theory_prefix operations target = target |> Data.map wenzelm@20910: (fn NONE => SOME (make_lthy (theory_prefix, operations, target)) wenzelm@20910: | SOME _ => error "Local theory already initialized"); wenzelm@20910: wenzelm@21273: fun restore lthy = wenzelm@20910: let val {theory_prefix, operations, target} = get_lthy lthy wenzelm@20910: in init theory_prefix operations target end; wenzelm@20910: wenzelm@21292: val reinit = operation #reinit; wenzelm@21292: val exit = operation #exit; wenzelm@20910: wenzelm@18781: end;