paulson@1526: (* Title: Pure/theory.ML paulson@1526: ID: $Id$ paulson@1526: Author: Lawrence C Paulson and Markus Wenzel paulson@1526: wenzelm@28290: Logical theory content: axioms, definitions, and begin/end wrappers. paulson@1526: *) wenzelm@16291: wenzelm@26668: signature THEORY = wenzelm@3767: sig wenzelm@16443: val eq_thy: theory * theory -> bool wenzelm@3996: val subthy: theory * theory -> bool wenzelm@24666: val assert_super: theory -> theory -> theory wenzelm@22684: val parents_of: theory -> theory list wenzelm@22684: val ancestors_of: theory -> theory list wenzelm@24666: val check_thy: theory -> theory_ref wenzelm@24666: val deref: theory_ref -> theory wenzelm@24666: val merge: theory * theory -> theory wenzelm@24666: val merge_refs: theory_ref * theory_ref -> theory_ref wenzelm@24666: val merge_list: theory list -> theory wenzelm@16495: val checkpoint: theory -> theory wenzelm@16495: val copy: theory -> theory wenzelm@24666: val requires: theory -> string -> string -> unit wenzelm@16443: val axiom_space: theory -> NameSpace.T wenzelm@22684: val axiom_table: theory -> term Symtab.table wenzelm@16339: val axioms_of: theory -> (string * term) list wenzelm@16339: val all_axioms_of: theory -> (string * term) list wenzelm@24666: val defs_of: theory -> Defs.T wenzelm@24666: val at_begin: (theory -> theory option) -> theory -> theory wenzelm@24666: val at_end: (theory -> theory option) -> theory -> theory wenzelm@24666: val begin_theory: string -> theory list -> theory wenzelm@24666: val end_theory: theory -> theory wenzelm@3996: val add_axioms: (bstring * string) list -> theory -> theory wenzelm@3996: val add_axioms_i: (bstring * term) list -> theory -> theory wenzelm@19708: val add_deps: string -> string * typ -> (string * typ) list -> theory -> theory wenzelm@19630: val add_defs: bool -> bool -> (bstring * string) list -> theory -> theory wenzelm@19630: val add_defs_i: bool -> bool -> (bstring * term) list -> theory -> theory wenzelm@16443: val add_finals: bool -> string list -> theory -> theory wenzelm@16443: val add_finals_i: bool -> term list -> theory -> theory haftmann@28965: val specify_const: Properties.T -> (Binding.T * typ) * mixfix -> theory -> term * theory wenzelm@16495: end paulson@1526: wenzelm@24666: structure Theory: THEORY = wenzelm@16443: struct wenzelm@16443: wenzelm@19708: wenzelm@24666: (** theory context operations **) wenzelm@16443: wenzelm@16443: val eq_thy = Context.eq_thy; wenzelm@16443: val subthy = Context.subthy; paulson@1526: haftmann@24626: fun assert_super thy1 thy2 = haftmann@24626: if subthy (thy1, thy2) then thy2 haftmann@24626: else raise THEORY ("Not a super theory", [thy1, thy2]); haftmann@24626: wenzelm@16443: val parents_of = Context.parents_of; wenzelm@16443: val ancestors_of = Context.ancestors_of; wenzelm@16443: wenzelm@24137: val check_thy = Context.check_thy; wenzelm@16443: val deref = Context.deref; wenzelm@24666: wenzelm@16443: val merge = Context.merge; wenzelm@16443: val merge_refs = Context.merge_refs; wenzelm@16443: wenzelm@23600: fun merge_list [] = raise THEORY ("Empty merge of theories", []) wenzelm@21608: | merge_list (thy :: thys) = Library.foldl merge (thy, thys); wenzelm@21608: wenzelm@16495: val checkpoint = Context.checkpoint_thy; wenzelm@16495: val copy = Context.copy_thy; wenzelm@16495: wenzelm@24666: fun requires thy name what = wenzelm@24666: if Context.exists_name name thy then () wenzelm@24666: else error ("Require theory " ^ quote name ^ " as an ancestor for " ^ what); wenzelm@24666: wenzelm@24666: wenzelm@24666: wenzelm@25059: (** datatype thy **) wenzelm@24666: wenzelm@24666: type wrapper = (theory -> theory option) * stamp; wenzelm@24666: wenzelm@24666: fun apply_wrappers (wrappers: wrapper list) = wenzelm@25059: perhaps (perhaps_loop (perhaps_apply (map fst wrappers))); wenzelm@24666: wenzelm@24666: datatype thy = Thy of wenzelm@24666: {axioms: term NameSpace.table, wenzelm@24666: defs: Defs.T, wenzelm@24666: wrappers: wrapper list * wrapper list}; wenzelm@24666: wenzelm@28290: fun make_thy (axioms, defs, wrappers) = Thy {axioms = axioms, defs = defs, wrappers = wrappers}; wenzelm@24666: wenzelm@24666: fun err_dup_axm dup = error ("Duplicate axiom: " ^ quote dup); wenzelm@24666: wenzelm@24666: structure ThyData = TheoryDataFun wenzelm@24666: ( wenzelm@24666: type T = thy; wenzelm@28290: val empty = make_thy (NameSpace.empty_table, Defs.empty, ([], [])); wenzelm@24666: val copy = I; wenzelm@24666: wenzelm@28290: fun extend (Thy {axioms, defs, wrappers}) = make_thy (NameSpace.empty_table, defs, wrappers); wenzelm@24666: wenzelm@24666: fun merge pp (thy1, thy2) = wenzelm@24666: let wenzelm@28290: val Thy {axioms = _, defs = defs1, wrappers = (bgs1, ens1)} = thy1; wenzelm@28290: val Thy {axioms = _, defs = defs2, wrappers = (bgs2, ens2)} = thy2; wenzelm@24666: wenzelm@24666: val axioms' = NameSpace.empty_table; wenzelm@24666: val defs' = Defs.merge pp (defs1, defs2); wenzelm@24666: val bgs' = Library.merge (eq_snd op =) (bgs1, bgs2); wenzelm@24666: val ens' = Library.merge (eq_snd op =) (ens1, ens2); wenzelm@28290: in make_thy (axioms', defs', (bgs', ens')) end; wenzelm@24666: ); wenzelm@24666: wenzelm@24666: fun rep_theory thy = ThyData.get thy |> (fn Thy args => args); wenzelm@24666: wenzelm@28290: fun map_thy f = ThyData.map (fn (Thy {axioms, defs, wrappers}) => wenzelm@28290: make_thy (f (axioms, defs, wrappers))); wenzelm@24666: wenzelm@24666: wenzelm@28290: fun map_axioms f = map_thy (fn (axioms, defs, wrappers) => (f axioms, defs, wrappers)); wenzelm@28290: fun map_defs f = map_thy (fn (axioms, defs, wrappers) => (axioms, f defs, wrappers)); wenzelm@28290: fun map_wrappers f = map_thy (fn (axioms, defs, wrappers) => (axioms, defs, f wrappers)); wenzelm@24666: wenzelm@24666: wenzelm@24666: (* basic operations *) wenzelm@24666: wenzelm@24666: val axiom_space = #1 o #axioms o rep_theory; wenzelm@24666: val axiom_table = #2 o #axioms o rep_theory; wenzelm@24666: wenzelm@24666: val axioms_of = Symtab.dest o #2 o #axioms o rep_theory; wenzelm@24666: fun all_axioms_of thy = maps axioms_of (thy :: ancestors_of thy); wenzelm@24666: wenzelm@24666: val defs_of = #defs o rep_theory; wenzelm@24666: wenzelm@24666: wenzelm@24666: (* begin/end theory *) wenzelm@24666: wenzelm@24666: val begin_wrappers = rev o #1 o #wrappers o rep_theory; wenzelm@24666: val end_wrappers = rev o #2 o #wrappers o rep_theory; wenzelm@24666: wenzelm@24666: fun at_begin f = map_wrappers (apfst (cons (f, stamp ()))); wenzelm@24666: fun at_end f = map_wrappers (apsnd (cons (f, stamp ()))); wenzelm@24666: wenzelm@24666: fun begin_theory name imports = wenzelm@24666: let wenzelm@26939: val thy = Context.begin_thy Syntax.pp_global name imports; wenzelm@24666: val wrappers = begin_wrappers thy; wenzelm@24666: in thy |> Sign.local_path |> apply_wrappers wrappers end; wenzelm@24666: wenzelm@24666: fun end_theory thy = wenzelm@24666: thy |> apply_wrappers (end_wrappers thy) |> Context.finish_thy; wenzelm@24666: wenzelm@16443: wenzelm@3996: wenzelm@24666: (** add axioms **) wenzelm@3814: paulson@1526: (* prepare axioms *) paulson@1526: wenzelm@18678: fun err_in_axm msg name = wenzelm@18678: cat_error msg ("The error(s) above occurred in axiom " ^ quote name); paulson@1526: wenzelm@16443: fun cert_axm thy (name, raw_tm) = paulson@1526: let wenzelm@18968: val (t, T, _) = Sign.certify_prop thy raw_tm wenzelm@2979: handle TYPE (msg, _, _) => error msg wenzelm@16291: | TERM (msg, _) => error msg; paulson@1526: in wenzelm@9537: Term.no_dummy_patterns t handle TERM (msg, _) => error msg; wenzelm@26939: (name, Sign.no_vars (Syntax.pp_global thy) t) wenzelm@9629: end; paulson@1526: wenzelm@22684: fun read_axm thy (name, str) = wenzelm@24708: cert_axm thy (name, Syntax.read_prop_global thy str) wenzelm@22689: handle ERROR msg => err_in_axm msg name; paulson@1526: paulson@1526: wenzelm@16443: (* add_axioms(_i) *) paulson@1526: wenzelm@16291: local wenzelm@16291: wenzelm@16443: fun gen_add_axioms prep_axm raw_axms thy = thy |> map_axioms (fn axioms => paulson@1526: let haftmann@29004: val axms = map (apfst Binding.name o apsnd Logic.varify o prep_axm thy) raw_axms; haftmann@29004: val axioms' = fold (snd oo NameSpace.bind (Sign.naming_of thy)) axms axioms wenzelm@23655: handle Symtab.DUP dup => err_dup_axm dup; wenzelm@16443: in axioms' end); paulson@1526: wenzelm@16291: in wenzelm@16291: wenzelm@16291: val add_axioms = gen_add_axioms read_axm; wenzelm@16291: val add_axioms_i = gen_add_axioms cert_axm; wenzelm@16291: wenzelm@16291: end; paulson@1526: paulson@1526: wenzelm@3767: wenzelm@3767: (** add constant definitions **) wenzelm@3767: wenzelm@19708: (* dependencies *) wenzelm@19708: wenzelm@19708: fun dependencies thy unchecked is_def name lhs rhs = wenzelm@19708: let wenzelm@26939: val pp = Syntax.pp_global thy; wenzelm@19708: val consts = Sign.consts_of thy; wenzelm@19727: fun prep const = wenzelm@19727: let val Const (c, T) = Sign.no_vars pp (Const const) wenzelm@26631: in (c, Consts.typargs consts (c, Logic.varifyT T)) end; wenzelm@19708: wenzelm@19708: val lhs_vars = Term.add_tfreesT (#2 lhs) []; wenzelm@19708: val rhs_extras = fold (#2 #> Term.fold_atyps (fn TFree v => wenzelm@19708: if member (op =) lhs_vars v then I else insert (op =) v | _ => I)) rhs []; wenzelm@19708: val _ = wenzelm@19708: if null rhs_extras then () wenzelm@19708: else error ("Specification depends on extra type variables: " ^ wenzelm@19708: commas_quote (map (Pretty.string_of_typ pp o TFree) rhs_extras) ^ wenzelm@19708: "\nThe error(s) above occurred in " ^ quote name); haftmann@24199: in Defs.define pp unchecked is_def name (prep lhs) (map prep rhs) end; wenzelm@19708: wenzelm@19708: fun add_deps a raw_lhs raw_rhs thy = wenzelm@19708: let wenzelm@19708: val lhs :: rhs = map (dest_Const o Sign.cert_term thy o Const) (raw_lhs :: raw_rhs); wenzelm@19708: val name = if a = "" then (#1 lhs ^ " axiom") else a; wenzelm@19708: in thy |> map_defs (dependencies thy false false name lhs rhs) end; wenzelm@17706: wenzelm@28112: fun specify_const tags decl thy = wenzelm@25017: let val (t as Const const, thy') = Sign.declare_const tags decl thy wenzelm@28112: in (t, add_deps "" const [] thy') end; wenzelm@25017: wenzelm@17706: wenzelm@16944: (* check_overloading *) nipkow@9280: wenzelm@16944: fun check_overloading thy overloaded (c, T) = wenzelm@16291: let wenzelm@24763: val declT = Sign.the_const_constraint thy c wenzelm@24763: handle TYPE (msg, _, _) => error msg; wenzelm@19806: val T' = Logic.varifyT T; wenzelm@16944: wenzelm@16944: fun message txt = wenzelm@16944: [Pretty.block [Pretty.str "Specification of constant ", wenzelm@26939: Pretty.str c, Pretty.str " ::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ_global thy T)], wenzelm@16944: Pretty.str txt] |> Pretty.chunks |> Pretty.string_of; wenzelm@16291: in wenzelm@16944: if Sign.typ_instance thy (declT, T') then () wenzelm@16944: else if Type.raw_instance (declT, T') then wenzelm@16944: error (Library.setmp show_sorts true wenzelm@16944: message "imposes additional sort constraints on the constant declaration") wenzelm@16944: else if overloaded then () wenzelm@16944: else warning (message "is strictly less general than the declared type"); wenzelm@16944: (c, T) nipkow@9280: end; nipkow@9280: wenzelm@3767: wenzelm@16291: (* check_def *) wenzelm@16291: wenzelm@19630: fun check_def thy unchecked overloaded (bname, tm) defs = wenzelm@16291: let wenzelm@24981: val ctxt = ProofContext.init thy; haftmann@28965: val name = Sign.full_bname thy bname; wenzelm@24981: val (lhs_const, rhs) = Sign.cert_def ctxt tm; wenzelm@16944: val rhs_consts = fold_aterms (fn Const const => insert (op =) const | _ => I) rhs []; wenzelm@18943: val _ = check_overloading thy overloaded lhs_const; wenzelm@19708: in defs |> dependencies thy unchecked true name lhs_const rhs_consts end wenzelm@18678: handle ERROR msg => cat_error msg (Pretty.string_of (Pretty.block wenzelm@16883: [Pretty.str ("The error(s) above occurred in definition " ^ quote bname ^ ":"), wenzelm@26939: Pretty.fbrk, Pretty.quote (Syntax.pretty_term_global thy tm)])); wenzelm@3767: wenzelm@3767: wenzelm@16443: (* add_defs(_i) *) wenzelm@3767: wenzelm@16291: local wenzelm@9320: wenzelm@19630: fun gen_add_defs prep_axm unchecked overloaded raw_axms thy = wenzelm@16443: let val axms = map (prep_axm thy) raw_axms in wenzelm@16443: thy wenzelm@19630: |> map_defs (fold (check_def thy unchecked overloaded) axms) wenzelm@9320: |> add_axioms_i axms wenzelm@3767: end; wenzelm@3767: wenzelm@16291: in wenzelm@16291: wenzelm@16291: val add_defs_i = gen_add_defs cert_axm; wenzelm@16291: val add_defs = gen_add_defs read_axm; wenzelm@16291: wenzelm@16291: end; wenzelm@3767: wenzelm@3767: wenzelm@16443: (* add_finals(_i) *) skalberg@14223: wenzelm@16291: local wenzelm@16291: wenzelm@17706: fun gen_add_finals prep_term overloaded args thy = skalberg@14223: let wenzelm@17706: fun const_of (Const const) = const wenzelm@17706: | const_of (Free _) = error "Attempt to finalize variable (or undeclared constant)" wenzelm@17706: | const_of _ = error "Attempt to finalize non-constant term"; wenzelm@19708: fun specify (c, T) = dependencies thy false false (c ^ " axiom") (c, T) []; wenzelm@24708: val finalize = specify o check_overloading thy overloaded o const_of o wenzelm@24708: Sign.cert_term thy o prep_term thy; wenzelm@17706: in thy |> map_defs (fold finalize args) end; wenzelm@16291: skalberg@14223: in wenzelm@16291: wenzelm@24708: val add_finals = gen_add_finals Syntax.read_term_global; wenzelm@24708: val add_finals_i = gen_add_finals (K I); wenzelm@16291: skalberg@14223: end; skalberg@14223: paulson@1526: end;