# HG changeset patch # User wenzelm # Date 1124356663 -7200 # Node ID 3962f74bbb8e9eb67e45e85bdba166d2cdd26d84 # Parent 2c35e00ee2abf2925739e350a0446416a45296f9 moved translation functions to Pure/sign.ML; moved attribute preparation to actual operations in proof.ML etc.; removed various trivial interfaces; diff -r 2c35e00ee2ab -r 3962f74bbb8e src/Pure/Isar/isar_thy.ML --- a/src/Pure/Isar/isar_thy.ML Thu Aug 18 11:17:42 2005 +0200 +++ b/src/Pure/Isar/isar_thy.ML Thu Aug 18 11:17:43 2005 +0200 @@ -21,39 +21,17 @@ -> theory -> theory * (string * thm list) list val locale_theorems: string -> xstring -> ((bstring * Attrib.src list) * (thmref * Attrib.src list) list) list - -> theory -> theory * (bstring * thm list) list + -> theory -> (theory * Proof.context) * (bstring * thm list) list val locale_theorems_i: string -> string -> ((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list -> - theory -> theory * (string * thm list) list + theory -> (theory * Proof.context) * (string * thm list) list val smart_theorems: string -> xstring option -> ((bstring * Attrib.src list) * (thmref * Attrib.src list) list) list - -> theory -> theory * (string * thm list) list + -> theory -> (theory * Proof.context) * (string * thm list) list val declare_theorems: xstring option -> - (thmref * Attrib.src list) list -> theory -> theory + (thmref * Attrib.src list) list -> theory -> theory * Proof.context val apply_theorems: (thmref * Attrib.src list) list -> theory -> theory * thm list val apply_theorems_i: (thm list * theory attribute list) list -> theory -> theory * thm list - val have_facts: ((string * Attrib.src list) * (thmref * Attrib.src list) list) list - -> ProofHistory.T -> ProofHistory.T - val have_facts_i: ((string * Proof.context attribute list) * - (thm * Proof.context attribute list) list) list -> ProofHistory.T -> ProofHistory.T - val from_facts: (thmref * Attrib.src list) list list -> ProofHistory.T -> ProofHistory.T - val from_facts_i: (thm * Proof.context attribute list) list list - -> ProofHistory.T -> ProofHistory.T - val with_facts: (thmref * Attrib.src list) list list -> ProofHistory.T -> ProofHistory.T - val with_facts_i: (thm * Proof.context attribute list) list list - -> ProofHistory.T -> ProofHistory.T - val using_facts: (thmref * Attrib.src list) list list - -> ProofHistory.T -> ProofHistory.T - val using_facts_i: (thm * Proof.context attribute list) list list - -> ProofHistory.T -> ProofHistory.T - val chain: ProofHistory.T -> ProofHistory.T - val fix: (string list * string option) list -> ProofHistory.T -> ProofHistory.T - val fix_i: (string list * typ option) list -> ProofHistory.T -> ProofHistory.T - val let_bind: (string list * string) list -> ProofHistory.T -> ProofHistory.T - val let_bind_i: (term list * term) list -> ProofHistory.T -> ProofHistory.T - val invoke_case: string * string option list * Attrib.src list -> ProofHistory.T -> ProofHistory.T - val invoke_case_i: string * string option list * Proof.context attribute list - -> ProofHistory.T -> ProofHistory.T val theorem: string -> (bstring * Attrib.src list) * (string * (string list * string list)) -> bool -> theory -> ProofHistory.T val theorem_i: string -> (bstring * theory attribute list) * @@ -87,16 +65,6 @@ -> bstring * Attrib.src list -> Locale.element Locale.elem_expr list -> ((bstring * Attrib.src list) * (string * (string list * string list)) list) list -> bool -> theory -> ProofHistory.T - val assume: ((string * Attrib.src list) * (string * (string list * string list)) list) list - -> ProofHistory.T -> ProofHistory.T - val assume_i: - ((string * Proof.context attribute list) * (term * (term list * term list)) list) list - -> ProofHistory.T -> ProofHistory.T - val presume: ((string * Attrib.src list) * (string * (string list * string list)) list) list - -> ProofHistory.T -> ProofHistory.T - val presume_i: - ((string * Proof.context attribute list) * (term * (term list * term list)) list) list - -> ProofHistory.T -> ProofHistory.T val have: ((string * Attrib.src list) * (string * (string list * string list)) list) list -> ProofHistory.T -> ProofHistory.T val have_i: ((string * Proof.context attribute list) * @@ -117,24 +85,12 @@ val thus_i: ((string * Proof.context attribute list) * (term * (term list * term list)) list) list -> bool -> ProofHistory.T -> ProofHistory.T - val local_def: (string * Attrib.src list) * (string * (string * string list)) - -> ProofHistory.T -> ProofHistory.T - val local_def_i: (string * Attrib.src list) * (string * (term * term list)) - -> ProofHistory.T -> ProofHistory.T val obtain: (string list * string option) list * ((string * Attrib.src list) * (string * (string list * string list)) list) list -> Toplevel.transition -> Toplevel.transition val obtain_i: (string list * typ option) list * ((string * Proof.context attribute list) * (term * (term list * term list)) list) list -> Toplevel.transition -> Toplevel.transition - val begin_block: ProofHistory.T -> ProofHistory.T - val next_block: ProofHistory.T -> ProofHistory.T - val end_block: ProofHistory.T -> ProofHistory.T - val defer: int option -> ProofHistory.T -> ProofHistory.T - val prefer: int -> ProofHistory.T -> ProofHistory.T - val apply: Method.text -> ProofHistory.T -> ProofHistory.T - val apply_end: Method.text -> ProofHistory.T -> ProofHistory.T - val proof: Method.text option -> ProofHistory.T -> ProofHistory.T val three_buffersN: string val cond_print: Toplevel.transition -> Toplevel.transition val qed: Method.text option -> Toplevel.transition -> Toplevel.transition @@ -152,16 +108,9 @@ val finally_i: thm list option -> Toplevel.transition -> Toplevel.transition val moreover: Toplevel.transition -> Toplevel.transition val ultimately: Toplevel.transition -> Toplevel.transition - val parse_ast_translation: bool * string -> theory -> theory - val parse_translation: bool * string -> theory -> theory - val print_translation: bool * string -> theory -> theory - val typed_print_translation: bool * string -> theory -> theory - val print_ast_translation: bool * string -> theory -> theory - val token_translation: string -> theory -> theory val generic_setup: string -> theory -> theory val method_setup: bstring * string * string -> theory -> theory val add_oracle: bstring * string * string -> theory -> theory - val add_locale: bool * bstring * Locale.expr * Locale.element list -> theory -> theory val register_globally: ((string * Attrib.src list) * Locale.expr) * string option list -> bool -> theory -> ProofHistory.T @@ -224,48 +173,41 @@ fun add_defs_i (x, args) = (#1 oo PureThy.add_defs_i x) args; -(* attributes *) - -local - -fun prep_attribs f = map (fn ((name, more_srcs), th_srcs) => - ((name, map f more_srcs), map (apsnd (map f)) th_srcs)); - -in - -fun global_attribs thy = prep_attribs (Attrib.global_attribute thy); -fun local_attribs thy = prep_attribs (Attrib.local_attribute thy); - -end; - - (* theorems *) local fun present_thmss kind (thy, named_thmss) = - (conditional (kind <> "" andalso kind <> Drule.internalK) (fn () => + conditional (kind <> "" andalso kind <> Drule.internalK) (fn () => Context.setmp (SOME thy) (Present.results (kind ^ "s")) named_thmss); - (thy, named_thmss)); + +fun present_thmss' kind ((thy, _), named_thmss) = present_thmss kind (thy, named_thmss); + +fun global_attribs thy = Attrib.map_facts (Attrib.global_attribute thy); in -fun theorems_i k = present_thmss k oo PureThy.note_thmss_i (Drule.kind k); -fun locale_theorems_i k loc = present_thmss k oo Locale.note_thmss_i k loc; +fun theorems_i k = tap (present_thmss k) oo PureThy.note_thmss_i (Drule.kind k); +fun locale_theorems_i k loc = tap (present_thmss' k) oo Locale.note_thmss_i k loc; fun theorems k args thy = thy |> PureThy.note_thmss (Drule.kind k) (global_attribs thy args) - |> present_thmss k; + |> tap (present_thmss k); fun locale_theorems k loc args thy = thy |> Locale.note_thmss k loc args - |> present_thmss k; + |> tap (present_thmss' k); -fun smart_theorems k opt_loc args thy = thy - |> (case opt_loc of - NONE => PureThy.note_thmss (Drule.kind k) (global_attribs thy args) - | SOME loc => Locale.note_thmss k loc args) - |> present_thmss k; +fun smart_theorems k NONE args thy = + let val (thy', res) = + thy |> PureThy.note_thmss (Drule.kind k) (global_attribs thy args) + in + present_thmss k (thy', res); + ((thy', ProofContext.init thy'), res) + end + | smart_theorems k (SOME loc) args thy = thy + |> Locale.note_thmss k loc args + |> tap (present_thmss' k); fun declare_theorems opt_loc args = #1 o smart_theorems "" opt_loc [(("", []), args)]; @@ -275,44 +217,6 @@ end; -(* facts and forward chaining *) - -fun local_thmss f args state = f (local_attribs (Proof.theory_of state) args) state; -fun local_thmss_i f args = f (map (fn ((name, more_atts), th_atts) => - ((name, more_atts), map (apfst single) th_atts)) args); - -fun chain_thmss f args = ProofHistory.apply (Proof.chain o f (map (pair ("", [])) args)); - -val have_facts = ProofHistory.apply o local_thmss Proof.note_thmss; -val have_facts_i = ProofHistory.apply o local_thmss_i Proof.note_thmss_i; -val from_facts = chain_thmss (local_thmss Proof.note_thmss); -val from_facts_i = chain_thmss (local_thmss_i Proof.note_thmss_i); -val with_facts = chain_thmss (local_thmss Proof.with_thmss); -val with_facts_i = chain_thmss (local_thmss_i Proof.with_thmss_i); - -fun using_facts args = ProofHistory.apply (fn state => - Proof.using_thmss (map (map (apsnd (map - (Attrib.local_attribute (Proof.theory_of state))))) args) state); - -val using_facts_i = ProofHistory.apply o Proof.using_thmss_i o map (map (apfst single)); - -val chain = ProofHistory.apply Proof.chain; - - -(* context *) - -val fix = ProofHistory.apply o Proof.fix; -val fix_i = ProofHistory.apply o Proof.fix_i; -val let_bind = ProofHistory.apply o Proof.let_bind; -val let_bind_i = ProofHistory.apply o Proof.let_bind_i; - -fun invoke_case (name, xs, src) = ProofHistory.apply (fn state => - Proof.invoke_case (name, xs, - map (Attrib.local_attribute (Proof.theory_of state)) src) state); - -val invoke_case_i = ProofHistory.apply o Proof.invoke_case; - - (* local results *) local @@ -369,29 +273,18 @@ local -fun global_attributes thy ((name, src), s) = ((name, map (Attrib.global_attribute thy) src), s); -fun local_attributes thy ((name, src), s) = ((name, map (Attrib.local_attribute thy) src), s); -fun local_attributes' state = local_attributes (Proof.theory_of state); - -fun global_statement f args int thy = - ProofHistory.init (Toplevel.undo_limit int) (f (map (global_attributes thy) args) thy); -fun global_statement_i f args int thy = ProofHistory.init (Toplevel.undo_limit int) (f args thy); - -fun local_statement' f g args int = ProofHistory.apply (fn state => - f (map (local_attributes' state) args) int (g state)); -fun local_statement_i' f g args int = ProofHistory.apply (f args int o g); +fun global_statement f args int thy = ProofHistory.init (Toplevel.undo_limit int) (f args thy); +fun local_statement' f g args int = ProofHistory.apply (f args int o g); fun local_statement f g args = local_statement' (K o f) g args (); -fun local_statement_i f g args = local_statement_i' (K o f) g args (); in fun multi_theorem k after_qed export a elems concl int thy = - global_statement (Proof.multi_theorem k after_qed export NONE - (apsnd (map (Attrib.global_attribute thy)) a) + global_statement (Proof.multi_theorem k after_qed export NONE a (map (Locale.intern_attrib_elem_expr thy) elems)) concl int thy; fun multi_theorem_i k after_qed export a = - global_statement_i o Proof.multi_theorem_i k after_qed export NONE a; + global_statement o Proof.multi_theorem_i k after_qed export NONE a; fun locale_multi_theorem k after_qed export locale (name, atts) elems concl int thy = global_statement (Proof.multi_theorem k after_qed export @@ -403,7 +296,7 @@ (map (apfst (apsnd (K []))) concl) int thy; fun locale_multi_theorem_i k after_qed export locale (name, atts) elems concl = - global_statement_i (Proof.multi_theorem_i k after_qed export + global_statement (Proof.multi_theorem_i k after_qed export (SOME (locale, (atts, map (snd o fst) concl))) (name, []) elems) (map (apfst (apsnd (K []))) concl); @@ -418,68 +311,33 @@ | smart_multi_theorem k (SOME locale) a elems = locale_multi_theorem k (K I) ProofContext.export_standard locale a elems; -val assume = local_statement Proof.assume I; -val assume_i = local_statement_i Proof.assume_i I; -val presume = local_statement Proof.presume I; -val presume_i = local_statement_i Proof.presume_i I; val have = local_statement (Proof.have (K Seq.single) true) I; -val have_i = local_statement_i (Proof.have_i (K Seq.single) true) I; +val have_i = local_statement (Proof.have_i (K Seq.single) true) I; val hence = local_statement (Proof.have (K Seq.single) true) Proof.chain; -val hence_i = local_statement_i (Proof.have_i (K Seq.single) true) Proof.chain; +val hence_i = local_statement (Proof.have_i (K Seq.single) true) Proof.chain; val show = local_statement' (Proof.show check_goal (K Seq.single) true) I; -val show_i = local_statement_i' (Proof.show_i check_goal (K Seq.single) true) I; +val show_i = local_statement' (Proof.show_i check_goal (K Seq.single) true) I; val thus = local_statement' (Proof.show check_goal (K Seq.single) true) Proof.chain; -val thus_i = local_statement_i' (Proof.show_i check_goal (K Seq.single) true) Proof.chain; - -fun gen_def f ((name, srcs), def) = ProofHistory.apply (fn state => - f name (map (Attrib.local_attribute (Proof.theory_of state)) srcs) def state); +val thus_i = local_statement' (Proof.show_i check_goal (K Seq.single) true) Proof.chain; -val local_def = gen_def Proof.def; -val local_def_i = gen_def Proof.def_i; - -fun obtain (xs, asms) = proof'' (ProofHistory.applys o (fn print => fn state => - Obtain.obtain xs (map (local_attributes' state) asms) print state)); - +fun obtain (xs, asms) = proof'' (ProofHistory.applys o Obtain.obtain xs asms); fun obtain_i (xs, asms) = proof'' (ProofHistory.applys o Obtain.obtain_i xs asms); end; -(* blocks *) - -val begin_block = ProofHistory.apply Proof.begin_block; -val next_block = ProofHistory.apply Proof.next_block; -val end_block = ProofHistory.applys Proof.end_block; - - -(* shuffle state *) - -fun shuffle meth i = Method.refine (Method.Basic (K (meth i))) o Proof.assert_no_chain; - -fun defer i = ProofHistory.applys (shuffle Method.defer i); -fun prefer i = ProofHistory.applys (shuffle Method.prefer i); - - -(* backward steps *) - -fun apply m = ProofHistory.applys - (Seq.map (Proof.goal_facts (K [])) o Method.refine m o Proof.assert_backward); -fun apply_end m = ProofHistory.applys (Method.refine_end m o Proof.assert_forward); -val proof = ProofHistory.applys o Method.proof; - - (* local endings *) val three_buffersN = "three_buffers"; val cond_print = Toplevel.print' three_buffersN; -val local_qed = proof'' o (ProofHistory.applys oo Method.local_qed true); +val local_qed = proof'' o (ProofHistory.applys oo Proof.local_qed true); val skip_local_qed = Toplevel.skip_proof (History.apply (fn i => if i > 1 then i - 1 else raise Toplevel.UNDEF)); -val local_terminal_proof = proof'' o (ProofHistory.applys oo Method.local_terminal_proof); -val local_default_proof = proof'' (ProofHistory.applys o Method.local_default_proof); -val local_immediate_proof = proof'' (ProofHistory.applys o Method.local_immediate_proof); -val local_done_proof = proof'' (ProofHistory.applys o Method.local_done_proof); +val local_terminal_proof = proof'' o (ProofHistory.applys oo Proof.local_terminal_proof); +val local_default_proof = proof'' (ProofHistory.applys o Proof.local_default_proof); +val local_immediate_proof = proof'' (ProofHistory.applys o Proof.local_immediate_proof); +val local_done_proof = proof'' (ProofHistory.applys o Proof.local_done_proof); val local_skip_proof = Toplevel.proof' (fn int => ProofHistory.applys (SkipProof.local_skip_proof (cond_print_result_rule int) int)); @@ -503,25 +361,25 @@ end; in #2 (foldl_map name_res (1, res)) end; -fun global_result finish = Toplevel.proof_to_theory' (fn int => fn prf => +fun global_result finish = Toplevel.proof_to_theory_context (fn int => fn prf => let val state = ProofHistory.current prf; val _ = if Proof.at_bottom state then () else raise Toplevel.UNDEF; - val (thy, ((kind, name), res)) = finish int state; + val ((thy, ctxt), ((kind, name), res)) = finish int state; in if kind = "" orelse kind = Drule.internalK then () else (print_result (Proof.context_of state) ((kind, name), res); Context.setmp (SOME thy) (Present.results kind) (name_results name res)); - thy + (thy, ctxt) end); -fun global_qed m = global_result (K (Method.global_qed true m)); +fun global_qed m = global_result (K (Proof.global_qed true m)); val skip_global_qed = Toplevel.skip_proof_to_theory (equal 1 o History.current); -fun global_terminal_proof m = global_result (K (Method.global_terminal_proof m)); -val global_default_proof = global_result (K Method.global_default_proof); -val global_immediate_proof = global_result (K Method.global_immediate_proof); +fun global_terminal_proof m = global_result (K (Proof.global_terminal_proof m)); +val global_default_proof = global_result (K Proof.global_default_proof); +val global_immediate_proof = global_result (K Proof.global_immediate_proof); val global_skip_proof = global_result SkipProof.global_skip_proof; -val global_done_proof = global_result (K Method.global_done_proof); +val global_done_proof = global_result (K Proof.global_done_proof); (* common endings *) @@ -552,7 +410,7 @@ in -fun get_thmss srcs = Proof.the_facts o local_thmss Proof.note_thmss [(("", []), srcs)]; +fun get_thmss srcs = Proof.the_facts o Proof.note_thmss [(("", []), srcs)]; fun get_thmss_i thms _ = thms; fun also x = proof''' (ProofHistory.applys o gen_calc get_thmss Calculation.also x); @@ -565,44 +423,6 @@ end; -(* translation functions *) - -fun advancedT false = "" - | advancedT true = "theory -> "; - -fun advancedN false = "" - | advancedN true = "advanced_"; - -fun parse_ast_translation (a, txt) = - txt |> Context.use_let ("val parse_ast_translation: (string * (" ^ advancedT a ^ - "Syntax.ast list -> Syntax.ast)) list") - ("Theory.add_" ^ advancedN a ^ "trfuns (parse_ast_translation, [], [], [])"); - -fun parse_translation (a, txt) = - txt |> Context.use_let ("val parse_translation: (string * (" ^ advancedT a ^ - "term list -> term)) list") - ("Theory.add_" ^ advancedN a ^ "trfuns ([], parse_translation, [], [])"); - -fun print_translation (a, txt) = - txt |> Context.use_let ("val print_translation: (string * (" ^ advancedT a ^ - "term list -> term)) list") - ("Theory.add_" ^ advancedN a ^ "trfuns ([], [], print_translation, [])"); - -fun print_ast_translation (a, txt) = - txt |> Context.use_let ("val print_ast_translation: (string * (" ^ advancedT a ^ - "Syntax.ast list -> Syntax.ast)) list") - ("Theory.add_" ^ advancedN a ^ "trfuns ([], [], [], print_ast_translation)"); - -fun typed_print_translation (a, txt) = - txt |> Context.use_let ("val typed_print_translation: (string * (" ^ advancedT a ^ - "bool -> typ -> term list -> term)) list") - ("Theory.add_" ^ advancedN a ^ "trfunsT typed_print_translation"); - -val token_translation = - Context.use_let "val token_translation: (string * string * (string -> string * real)) list" - "Theory.add_tokentrfuns token_translation"; - - (* generic_setup *) val generic_setup = @@ -638,12 +458,6 @@ in Context.use_mltext_theory txt false end; -(* add_locale *) - -fun add_locale (do_pred, name, import, body) thy = - Locale.add_locale do_pred name import body thy; - - (* registration of locale interpretations *) fun register_globally (((prfx, atts), expr), insts) int thy =