src/Pure/Isar/isar_thy.ML
changeset 12876 a70df1e5bf10
parent 12848 ac191fb20ff1
child 12929 dbac8831d954
--- a/src/Pure/Isar/isar_thy.ML	Tue Feb 12 20:25:58 2002 +0100
+++ b/src/Pure/Isar/isar_thy.ML	Tue Feb 12 20:28:27 2002 +0100
@@ -8,206 +8,154 @@
 
 signature ISAR_THY =
 sig
-  val add_header: Comment.text -> Toplevel.transition -> Toplevel.transition
-  val add_chapter: Comment.text -> theory -> theory
-  val add_section: Comment.text -> theory -> theory
-  val add_subsection: Comment.text -> theory -> theory
-  val add_subsubsection: Comment.text -> theory -> theory
-  val add_text: Comment.text -> theory -> theory
-  val add_text_raw: Comment.text -> theory -> theory
-  val add_sect: Comment.text -> ProofHistory.T -> ProofHistory.T
-  val add_subsect: Comment.text -> ProofHistory.T -> ProofHistory.T
-  val add_subsubsect: Comment.text -> ProofHistory.T -> ProofHistory.T
-  val add_txt: Comment.text -> ProofHistory.T -> ProofHistory.T
-  val add_txt_raw: Comment.text -> ProofHistory.T -> ProofHistory.T
-  val global_path: Comment.text -> theory -> theory
-  val local_path: Comment.text -> theory -> theory
-  val hide_space: (string * xstring list) * Comment.text -> theory -> theory
-  val hide_space_i: (string * string list) * Comment.text -> theory -> theory
-  val add_classes: (bclass * xclass list) list * Comment.text -> theory -> theory
-  val add_classes_i: (bclass * class list) list * Comment.text  -> theory -> theory
-  val add_classrel: (xclass * xclass) * Comment.text -> theory -> theory
-  val add_classrel_i: (class * class) * Comment.text -> theory -> theory
-  val add_defsort: string * Comment.text -> theory -> theory
-  val add_defsort_i: sort * Comment.text -> theory -> theory
-  val add_nonterminals: (bstring * Comment.text) list -> theory -> theory
-  val add_tyabbrs: ((bstring * string list * string * mixfix) * Comment.text) list
-    -> theory -> theory
-  val add_tyabbrs_i: ((bstring * string list * typ * mixfix) * Comment.text) list
-    -> theory -> theory
-  val add_arities: ((xstring * string list * string) * Comment.text) list -> theory -> theory
-  val add_arities_i: ((string * sort list * sort) * Comment.text) list -> theory -> theory
-  val add_typedecl: (bstring * string list * mixfix) * Comment.text -> theory -> theory
-  val add_consts: ((bstring * string * mixfix) * Comment.text) list -> theory -> theory
-  val add_consts_i: ((bstring * typ * mixfix) * Comment.text) list -> theory -> theory
-  val add_modesyntax: string * bool -> ((bstring * string * mixfix) * Comment.text) list
-    -> theory -> theory
-  val add_modesyntax_i: string * bool -> ((bstring * typ * mixfix) * Comment.text) list
-    -> theory -> theory
-  val add_trrules: ((xstring * string) Syntax.trrule * Comment.text) list -> theory -> theory
-  val add_trrules_i: (Syntax.ast Syntax.trrule * Comment.text) list -> theory -> theory
-  val add_judgment: (bstring * string * mixfix) * Comment.text -> theory -> theory
-  val add_judgment_i: (bstring * typ * mixfix) * Comment.text -> theory -> theory
-  val add_axioms: (((bstring * string) * Args.src list) * Comment.text) list -> theory -> theory
-  val add_axioms_i: (((bstring * term) * theory attribute list) * Comment.text) list
-    -> theory -> theory
-  val add_defs: bool * (((bstring * string) * Args.src list) * Comment.text) list
-    -> theory -> theory
-  val add_defs_i: bool * (((bstring * term) * theory attribute list) * Comment.text) list
-    -> theory -> theory
-  val add_constdefs: (((bstring * string * mixfix) * Comment.text) * (string * Comment.text)) list
-    -> theory -> theory
-  val add_constdefs_i: (((bstring * typ * mixfix) * Comment.text) * (term * Comment.text)) list
-    -> theory -> theory
-  val theorems: string ->
-    (((bstring * Args.src list) * (xstring * Args.src list) list) * Comment.text) list
+  val add_header: string -> Toplevel.transition -> Toplevel.transition
+  val add_chapter: string -> theory -> theory
+  val add_section: string -> theory -> theory
+  val add_subsection: string -> theory -> theory
+  val add_subsubsection: string -> theory -> theory
+  val add_text: string -> theory -> theory
+  val add_text_raw: string -> theory -> theory
+  val add_sect: string -> ProofHistory.T -> ProofHistory.T
+  val add_subsect: string -> ProofHistory.T -> ProofHistory.T
+  val add_subsubsect: string -> ProofHistory.T -> ProofHistory.T
+  val add_txt: string -> ProofHistory.T -> ProofHistory.T
+  val add_txt_raw: string -> ProofHistory.T -> ProofHistory.T
+  val hide_space: string * xstring list -> theory -> theory
+  val hide_space_i: string * string list -> theory -> theory
+  val add_axioms: ((bstring * string) * Args.src list) list -> theory -> theory
+  val add_axioms_i: ((bstring * term) * theory attribute list) list -> theory -> theory
+  val add_defs: bool * ((bstring * string) * Args.src list) list -> theory -> theory
+  val add_defs_i: bool * ((bstring * term) * theory attribute list) list -> theory -> theory
+  val add_constdefs: ((bstring * string * mixfix) * string) list -> theory -> theory
+  val add_constdefs_i: ((bstring * typ * mixfix) * term) list -> theory -> theory
+  val theorems: string -> ((bstring * Args.src list) * (xstring * Args.src list) list) list
     -> theory -> theory * (string * thm list) list
   val theorems_i: string ->
-    (((bstring * theory attribute list)
-      * (thm list * theory attribute list) list) * Comment.text) list
+    ((bstring * theory attribute list) * (thm list * theory attribute list) list) list
     -> theory -> theory * (string * thm list) list
   val locale_theorems: string -> xstring ->
-    (((bstring * Args.src list) * (xstring * Args.src list) list) * Comment.text) list
+    ((bstring * Args.src list) * (xstring * Args.src list) list) list
     -> theory -> theory * (bstring * thm list) list
   val locale_theorems_i: string -> string ->
-    (((bstring * Proof.context attribute list)
-      * (thm list * Proof.context attribute list) list) * Comment.text) list
+    ((bstring * Proof.context attribute list)
+      * (thm list * Proof.context attribute list) list) list
     -> theory -> theory * (string * thm list) list
   val smart_theorems: string -> xstring option ->
-    (((bstring * Args.src list) * (xstring * Args.src list) list) * Comment.text) list
+    ((bstring * Args.src list) * (xstring * Args.src list) list) list
     -> theory -> theory * (string * thm list) list
-  val declare_theorems: xstring option -> (xstring * Args.src list) list * Comment.text
-    -> theory -> theory
+  val declare_theorems: xstring option -> (xstring * Args.src list) list -> theory -> theory
   val apply_theorems: (xstring * Args.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 * Args.src list) * (string * Args.src list) list) * Comment.text) list
-    -> ProofHistory.T -> ProofHistory.T
-  val have_facts_i: (((string * Proof.context attribute list) *
-    (thm * Proof.context attribute list) list) * Comment.text) list
+  val have_facts: ((string * Args.src list) * (string * Args.src list) list) list
     -> ProofHistory.T -> ProofHistory.T
-  val from_facts: ((string * Args.src list) list * Comment.text) list
-    -> ProofHistory.T -> ProofHistory.T
-  val from_facts_i: ((thm * Proof.context attribute list) list * Comment.text) list
+  val have_facts_i: ((string * Proof.context attribute list) *
+    (thm * Proof.context attribute list) list) list -> ProofHistory.T -> ProofHistory.T
+  val from_facts: (string * Args.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: ((string * Args.src list) list * Comment.text) list
-    -> ProofHistory.T -> ProofHistory.T
-  val with_facts_i: ((thm * Proof.context attribute list) list * Comment.text) list
+  val with_facts: (string * Args.src list) list list -> ProofHistory.T -> ProofHistory.T
+  val with_facts_i: (thm * Proof.context attribute list) list list
     -> ProofHistory.T -> ProofHistory.T
-  val chain: Comment.text -> ProofHistory.T -> ProofHistory.T
-  val fix: ((string list * string option) * Comment.text) list -> ProofHistory.T -> ProofHistory.T
-  val fix_i: ((string list * typ option) * Comment.text) list -> ProofHistory.T -> ProofHistory.T
-  val let_bind: ((string list * string) * Comment.text) list -> ProofHistory.T -> ProofHistory.T
-  val let_bind_i: ((term list * term) * Comment.text) list -> ProofHistory.T -> ProofHistory.T
-  val invoke_case: (string * string option list * Args.src list) * Comment.text
+  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 * Args.src list -> ProofHistory.T -> ProofHistory.T
+  val invoke_case_i: string * string option list * Proof.context attribute list
     -> ProofHistory.T -> ProofHistory.T
-  val invoke_case_i: (string * string option list * Proof.context attribute list) * Comment.text
-    -> ProofHistory.T -> ProofHistory.T
-  val theorem: string
-    -> ((bstring * Args.src list) * (string * (string list * string list))) * Comment.text
+  val theorem: string -> (bstring * Args.src list) * (string * (string list * string list))
     -> bool -> theory -> ProofHistory.T
-  val theorem_i: string -> ((bstring * theory attribute list) *
-    (term * (term list * term list))) * Comment.text -> bool -> theory -> ProofHistory.T
+  val theorem_i: string -> (bstring * theory attribute list) *
+    (term * (term list * term list)) -> bool -> theory -> ProofHistory.T
   val multi_theorem: string -> bstring * Args.src list ->
     Args.src Locale.element list ->
-    (((bstring * Args.src list) * (string * (string list * string list)) list)
-      * Comment.text) list -> bool -> theory -> ProofHistory.T
+    ((bstring * Args.src list) * (string * (string list * string list)) list) list
+    -> bool -> theory -> ProofHistory.T
   val multi_theorem_i: string -> bstring * theory attribute list ->
     Proof.context attribute Locale.element_i list ->
-    (((bstring * theory attribute list) * (term * (term list * term list)) list)
-      * Comment.text) list -> bool -> theory -> ProofHistory.T
+    ((bstring * theory attribute list) * (term * (term list * term list)) list) list
+    -> bool -> theory -> ProofHistory.T
   val locale_multi_theorem: string -> bstring * Args.src list ->
     xstring -> Args.src Locale.element list ->
-    (((bstring * Args.src list) * (string * (string list * string list)) list)
-      * Comment.text) list -> bool -> theory -> ProofHistory.T
+    ((bstring * Args.src list) * (string * (string list * string list)) list) list
+    -> bool -> theory -> ProofHistory.T
   val locale_multi_theorem_i: string -> bstring * theory attribute list ->
     string -> Proof.context attribute Locale.element_i list ->
-    (((bstring * Proof.context attribute list) * (term * (term list * term list)) list)
-      * Comment.text) list -> bool -> theory -> ProofHistory.T
+    ((bstring * Proof.context attribute list) * (term * (term list * term list)) list) list
+    -> bool -> theory -> ProofHistory.T
   val smart_multi_theorem: string -> bstring * Args.src list ->
     xstring Library.option * Args.src Locale.element list ->
-    (((bstring * Args.src list) * (string * (string list * string list)) list)
-      * Comment.text) list -> bool -> theory -> ProofHistory.T
-  val assume: (((string * Args.src list) * (string * (string list * string list)) list)
-    * Comment.text) list -> ProofHistory.T -> ProofHistory.T
-  val assume_i: (((string * Proof.context attribute list) * (term * (term list * term list)) list)
-    * Comment.text) list -> ProofHistory.T -> ProofHistory.T
-  val presume: (((string * Args.src list) * (string * (string list * string list)) list)
-    * Comment.text) list -> ProofHistory.T -> ProofHistory.T
-  val presume_i: (((string * Proof.context attribute list) * (term * (term list * term list)) list)
-    * Comment.text) list -> ProofHistory.T -> ProofHistory.T
-  val have: (((string * Args.src list) *
-    (string * (string list * string list)) list) * Comment.text) list
+    ((bstring * Args.src list) * (string * (string list * string list)) list) list
+    -> bool -> theory -> ProofHistory.T
+  val assume: ((string * Args.src list) * (string * (string list * string list)) list) list
     -> ProofHistory.T -> ProofHistory.T
-  val have_i: (((string * Proof.context attribute list) *
-    (term * (term list * term list)) list) * Comment.text) list
+  val assume_i:
+    ((string * Proof.context attribute list) * (term * (term list * term list)) list) list
+    -> ProofHistory.T -> ProofHistory.T
+  val presume: ((string * Args.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 hence: (((string * Args.src list) *
-    (string * (string list * string list)) list) * Comment.text) list
+  val have: ((string * Args.src list) * (string * (string list * string list)) list) list
     -> ProofHistory.T -> ProofHistory.T
-  val hence_i: (((string * Proof.context attribute list) *
-    (term * (term list * term list)) list) * Comment.text) list
+  val have_i: ((string * Proof.context attribute list) *
+    (term * (term list * term list)) list) list -> ProofHistory.T -> ProofHistory.T
+  val hence: ((string * Args.src list) * (string * (string list * string list)) list) list
     -> ProofHistory.T -> ProofHistory.T
-  val show: (((string * Args.src list) *
-    (string * (string list * string list)) list) * Comment.text) list
+  val hence_i: ((string * Proof.context attribute list) *
+    (term * (term list * term list)) list) list -> ProofHistory.T -> ProofHistory.T
+  val show: ((string * Args.src list) * (string * (string list * string list)) list) list
     -> bool -> ProofHistory.T -> ProofHistory.T
-  val show_i: (((string * Proof.context attribute list) *
-    (term * (term list * term list)) list) * Comment.text) list
-    -> bool -> ProofHistory.T -> ProofHistory.T
-  val thus: (((string * Args.src list) *
-    (string * (string list * string list)) list) * Comment.text) list
-    -> bool -> ProofHistory.T -> ProofHistory.T
-  val thus_i: (((string * Proof.context attribute list) *
-    (term * (term list * term list)) list) * Comment.text) list
+  val show_i: ((string * Proof.context attribute list) *
+    (term * (term list * term list)) list) list -> bool -> ProofHistory.T -> ProofHistory.T
+  val thus: ((string * Args.src list) * (string * (string list * string list)) list) list
     -> bool -> ProofHistory.T -> ProofHistory.T
-  val local_def: ((string * Args.src list) * (string * (string * string list)))
-    * Comment.text -> ProofHistory.T -> ProofHistory.T
-  val local_def_i: ((string * Args.src list) * (string * (term * term list)))
-    * Comment.text -> ProofHistory.T -> ProofHistory.T
-  val obtain: ((string list * string option) * Comment.text) list
-    * (((string * Args.src list) * (string * (string list * string list)) list)
-      * Comment.text) list -> ProofHistory.T -> ProofHistory.T
-  val obtain_i: ((string list * typ option) * Comment.text) list
-    * (((string * Proof.context attribute list) * (term * (term list * term list)) list)
-      * Comment.text) list -> ProofHistory.T -> ProofHistory.T
-  val begin_block: Comment.text -> ProofHistory.T -> ProofHistory.T
-  val next_block: Comment.text -> ProofHistory.T -> ProofHistory.T
-  val end_block: Comment.text -> ProofHistory.T -> ProofHistory.T
-  val defer: int option * Comment.text -> ProofHistory.T -> ProofHistory.T
-  val prefer: int * Comment.text -> ProofHistory.T -> ProofHistory.T
-  val apply: Method.text * Comment.text -> ProofHistory.T -> ProofHistory.T
-  val apply_end: Method.text * Comment.text -> ProofHistory.T -> ProofHistory.T
-  val proof: (Comment.interest * (Method.text * Comment.interest) option) * Comment.text
+  val thus_i: ((string * Proof.context attribute list)
+    * (term * (term list * term list)) list) list -> bool -> ProofHistory.T -> ProofHistory.T
+  val local_def: (string * Args.src list) * (string * (string * string list))
+    -> ProofHistory.T -> ProofHistory.T
+  val local_def_i: (string * Args.src list) * (string * (term * term list))
+    -> ProofHistory.T -> ProofHistory.T
+  val obtain: (string list * string option) list
+    * ((string * Args.src list) * (string * (string list * string list)) list) list
     -> ProofHistory.T -> ProofHistory.T
-  val qed: (Method.text * Comment.interest) option * Comment.text
+  val obtain_i: (string list * typ option) list
+    * ((string * Proof.context attribute list) * (term * (term list * term list)) list) list
+    -> ProofHistory.T -> ProofHistory.T
+  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 qed: Method.text option -> Toplevel.transition -> Toplevel.transition
+  val terminal_proof: Method.text * Method.text option
     -> Toplevel.transition -> Toplevel.transition
-  val terminal_proof: ((Method.text * Comment.interest) * (Method.text * Comment.interest) option)
-    * Comment.text -> Toplevel.transition -> Toplevel.transition
-  val default_proof: Comment.text -> Toplevel.transition -> Toplevel.transition
-  val immediate_proof: Comment.text -> Toplevel.transition -> Toplevel.transition
-  val done_proof: Comment.text -> Toplevel.transition -> Toplevel.transition
-  val skip_proof: Comment.text -> Toplevel.transition -> Toplevel.transition
-  val forget_proof: Comment.text -> Toplevel.transition -> Toplevel.transition
+  val default_proof: Toplevel.transition -> Toplevel.transition
+  val immediate_proof: Toplevel.transition -> Toplevel.transition
+  val done_proof: Toplevel.transition -> Toplevel.transition
+  val skip_proof: Toplevel.transition -> Toplevel.transition
+  val forget_proof: Toplevel.transition -> Toplevel.transition
   val get_thmss: (string * Args.src list) list -> Proof.state -> thm list
-  val also: ((string * Args.src list) list * Comment.interest) option * Comment.text
-    -> Toplevel.transition -> Toplevel.transition
-  val also_i: (thm list * Comment.interest) option * Comment.text
-    -> Toplevel.transition -> Toplevel.transition
-  val finally: ((string * Args.src list) list * Comment.interest) option * Comment.text
-    -> Toplevel.transition -> Toplevel.transition
-  val finally_i: (thm list * Comment.interest) option * Comment.text
-    -> Toplevel.transition -> Toplevel.transition
-  val moreover: Comment.text -> Toplevel.transition -> Toplevel.transition
-  val ultimately: Comment.text -> Toplevel.transition -> Toplevel.transition
-  val parse_ast_translation: string * Comment.text -> theory -> theory
-  val parse_translation: string * Comment.text -> theory -> theory
-  val print_translation: string * Comment.text -> theory -> theory
-  val typed_print_translation: string * Comment.text -> theory -> theory
-  val print_ast_translation: string * Comment.text -> theory -> theory
-  val token_translation: string * Comment.text -> theory -> theory
-  val method_setup: (bstring * string * string) * Comment.text -> theory -> theory
-  val add_oracle: (bstring * string) * Comment.text -> theory -> theory
-  val add_locale: bstring * Locale.expr * (Args.src Locale.element * Comment.text) list
-    -> theory -> theory
+  val also: (string * Args.src list) list option -> Toplevel.transition -> Toplevel.transition
+  val also_i: thm list option -> Toplevel.transition -> Toplevel.transition
+  val finally: (string * Args.src list) list option -> Toplevel.transition -> Toplevel.transition
+  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: string -> theory -> theory
+  val parse_translation: string -> theory -> theory
+  val print_translation: string -> theory -> theory
+  val typed_print_translation: string -> theory -> theory
+  val print_ast_translation: string -> theory -> theory
+  val token_translation: string -> theory -> theory
+  val method_setup: bstring * string * string -> theory -> theory
+  val add_oracle: bstring * string -> theory -> theory
+  val add_locale: bstring * Locale.expr * Args.src Locale.element list -> theory -> theory
   val begin_theory: string -> string list -> (string * bool) list -> theory
   val end_theory: theory -> theory
   val kill_theory: string -> unit
@@ -225,21 +173,20 @@
 
 (* markup commands *)
 
-fun add_header comment =
+fun add_header x =
   Toplevel.keep (fn state => if Toplevel.is_toplevel state then () else raise Toplevel.UNDEF);
 
 fun add_text _ x = x;
 fun add_text_raw _ x = x;
 
-fun add_heading add present txt thy =
-  (Context.setmp (Some thy) present (Comment.string_of txt); add txt thy);
+fun add_heading add present txt thy = (Context.setmp (Some thy) present txt; add txt thy);
 
 val add_chapter = add_heading add_text Present.section;
 val add_section = add_heading add_text Present.section;
 val add_subsection = add_heading add_text Present.subsection;
 val add_subsubsection = add_heading add_text Present.subsubsection;
 
-fun add_txt (_: Comment.text) = ProofHistory.apply I;
+fun add_txt (_: string) = ProofHistory.apply I;
 val add_txt_raw = add_txt;
 val add_sect = add_txt;
 val add_subsect = add_txt;
@@ -248,9 +195,6 @@
 
 (* name spaces *)
 
-fun global_path comment_ignore = PureThy.global_path;
-fun local_path comment_ignore = PureThy.local_path;
-
 local
 
 val kinds =
@@ -259,7 +203,7 @@
     can (Sign.certify_tycon sg) c orelse can (Sign.certify_tyabbr sg) c),
   (Sign.constK, can o Sign.certify_const)];
 
-fun gen_hide intern ((kind, xnames), comment_ignore) thy =
+fun gen_hide intern (kind, xnames) thy =
   (case assoc (kinds, kind) of
     Some check =>
       let
@@ -280,51 +224,27 @@
 end;
 
 
-(* signature and syntax *)
-
-val add_classes = Theory.add_classes o Comment.ignore;
-val add_classes_i = Theory.add_classes_i o Comment.ignore;
-val add_classrel = Theory.add_classrel o single o Comment.ignore;
-val add_classrel_i = Theory.add_classrel_i o single o Comment.ignore;
-val add_defsort = Theory.add_defsort o Comment.ignore;
-val add_defsort_i = Theory.add_defsort_i o Comment.ignore;
-val add_nonterminals = Theory.add_nonterminals o map Comment.ignore;
-val add_tyabbrs = Theory.add_tyabbrs o map Comment.ignore;
-val add_tyabbrs_i = Theory.add_tyabbrs_i o map Comment.ignore;
-val add_arities = Theory.add_arities o map Comment.ignore;
-val add_arities_i = Theory.add_arities_i o map Comment.ignore;
-val add_typedecl = PureThy.add_typedecls o single o Comment.ignore;
-val add_consts = Theory.add_consts o map Comment.ignore;
-val add_consts_i = Theory.add_consts_i o map Comment.ignore;
-fun add_modesyntax mode = Theory.add_modesyntax mode o map Comment.ignore;
-fun add_modesyntax_i mode = Theory.add_modesyntax_i mode o map Comment.ignore;
-val add_trrules = Theory.add_trrules o map Comment.ignore;
-val add_trrules_i = Theory.add_trrules_i o map Comment.ignore;
-val add_judgment = ObjectLogic.add_judgment o Comment.ignore;
-val add_judgment_i = ObjectLogic.add_judgment_i o Comment.ignore;
-
-
 (* axioms and defs *)
 
 fun add_axms f args thy =
   f (map (fn (x, srcs) => (x, map (Attrib.global_attribute thy) srcs)) args) thy;
 
-val add_axioms = add_axms (#1 oo PureThy.add_axioms) o map Comment.ignore;
-val add_axioms_i = (#1 oo PureThy.add_axioms_i) o map Comment.ignore;
-fun add_defs (x, args) = add_axms (#1 oo PureThy.add_defs x) (map Comment.ignore args);
-fun add_defs_i (x, args) = (#1 oo PureThy.add_defs_i x) (map Comment.ignore args);
+val add_axioms = add_axms (#1 oo PureThy.add_axioms);
+val add_axioms_i = #1 oo PureThy.add_axioms_i;
+fun add_defs (x, args) = add_axms (#1 oo PureThy.add_defs x) args;
+fun add_defs_i (x, args) = (#1 oo PureThy.add_defs_i x) args;
 
 
 (* constdefs *)
 
 fun gen_add_constdefs consts defs args thy =
   thy
-  |> consts (map (Comment.ignore o fst) args)
-  |> defs (false, map (fn (((c, _, mx), _), (s, _)) =>
-    (((Thm.def_name (Syntax.const_name c mx), s), []), Comment.none)) args);
+  |> consts (map fst args)
+  |> defs (false, map (fn ((c, _, mx), s) =>
+    ((Thm.def_name (Syntax.const_name c mx), s), [])) args);
 
-fun add_constdefs args = gen_add_constdefs Theory.add_consts add_defs args;
-fun add_constdefs_i args = gen_add_constdefs Theory.add_consts_i add_defs_i args;
+val add_constdefs = gen_add_constdefs Theory.add_consts add_defs;
+val add_constdefs_i = gen_add_constdefs Theory.add_consts_i add_defs_i;
 
 
 (* attributes *)
@@ -355,30 +275,27 @@
 
 in
 
-fun theorems_i k = (present_thmss k oo PureThy.have_thmss_i (Drule.kind k)) o map Comment.ignore;
-fun locale_theorems_i k loc = (present_thmss k oo Locale.have_thmss_i k loc) o map Comment.ignore;
+fun theorems_i k = present_thmss k oo PureThy.have_thmss_i (Drule.kind k);
+fun locale_theorems_i k loc = present_thmss k oo Locale.have_thmss_i k loc;
 
 fun theorems k args thy = thy
-  |> PureThy.have_thmss (Drule.kind k) (global_attribs thy (map Comment.ignore args))
+  |> PureThy.have_thmss (Drule.kind k) (global_attribs thy args)
   |> present_thmss k;
 
 fun locale_theorems k loc args thy = thy
-  |> Locale.have_thmss k loc (local_attribs thy (map Comment.ignore args))
+  |> Locale.have_thmss k loc (local_attribs thy args)
   |> present_thmss k;
 
 fun smart_theorems k opt_loc args thy = thy
   |> (case opt_loc of
-    None => PureThy.have_thmss (Drule.kind k) (global_attribs thy (map Comment.ignore args))
-  | Some loc => Locale.have_thmss k loc (local_attribs thy (map Comment.ignore args)))
+    None => PureThy.have_thmss (Drule.kind k) (global_attribs thy args)
+  | Some loc => Locale.have_thmss k loc (local_attribs thy args))
   |> present_thmss k;
 
-fun declare_theorems opt_loc (args, comment) =
-  #1 o smart_theorems "" opt_loc [((("", []), args), comment)];
+fun declare_theorems opt_loc args = #1 o smart_theorems "" opt_loc [(("", []), args)];
 
-fun apply_theorems args =
-  apsnd (flat o map snd) o theorems "" [((("", []), args), Comment.none)];
-fun apply_theorems_i args =
-  apsnd (flat o map snd) o theorems_i "" [((("", []), args), Comment.none)];
+fun apply_theorems args = apsnd (flat o map snd) o theorems "" [(("", []), args)];
+fun apply_theorems_i args = apsnd (flat o map snd) o theorems_i "" [(("", []), args)];
 
 end;
 
@@ -391,25 +308,25 @@
 
 fun chain_thmss f args = ProofHistory.apply (Proof.chain o f (map (pair ("", [])) args));
 
-val have_facts = ProofHistory.apply o local_thmss Proof.have_thmss o map Comment.ignore;
-val have_facts_i = ProofHistory.apply o local_thmss_i Proof.have_thmss_i o map Comment.ignore;
-val from_facts = chain_thmss (local_thmss Proof.have_thmss) o map Comment.ignore;
-val from_facts_i = chain_thmss (local_thmss_i Proof.have_thmss_i) o map Comment.ignore;
-val with_facts = chain_thmss (local_thmss Proof.with_thmss) o map Comment.ignore;
-val with_facts_i = chain_thmss (local_thmss_i Proof.with_thmss_i) o map Comment.ignore;
-fun chain comment_ignore = ProofHistory.apply Proof.chain;
+val have_facts = ProofHistory.apply o local_thmss Proof.have_thmss;
+val have_facts_i = ProofHistory.apply o local_thmss_i Proof.have_thmss_i;
+val from_facts = chain_thmss (local_thmss Proof.have_thmss);
+val from_facts_i = chain_thmss (local_thmss_i Proof.have_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);
+val chain = ProofHistory.apply Proof.chain;
 
 
 (* context *)
 
-val fix = ProofHistory.apply o Proof.fix o map Comment.ignore;
-val fix_i = ProofHistory.apply o Proof.fix_i o map Comment.ignore;
-val let_bind = ProofHistory.apply o Proof.let_bind o map Comment.ignore;
-val let_bind_i = ProofHistory.apply o Proof.let_bind_i o map Comment.ignore;
+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), comment_ignore) = ProofHistory.apply (fn state =>
+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 o Comment.ignore;
+val invoke_case_i = ProofHistory.apply o Proof.invoke_case;
 
 
 (* local results *)
@@ -487,94 +404,80 @@
 
 fun multi_theorem k a elems args int thy =
   global_statement (Proof.multi_theorem k (apsnd (map (Attrib.global_attribute thy)) a) None
-    (map (Locale.attribute (Attrib.local_attribute thy)) elems)) (map Comment.ignore args) int thy;
+    (map (Locale.attribute (Attrib.local_attribute thy)) elems)) args int thy;
 
-fun multi_theorem_i k a elems =
-  global_statement_i (Proof.multi_theorem_i k a None elems) o map Comment.ignore;
+fun multi_theorem_i k a elems = global_statement_i (Proof.multi_theorem_i k a None elems);
 
 fun locale_multi_theorem k a locale elems args int thy =
   global_statement (Proof.multi_theorem k (apsnd (map (Attrib.global_attribute thy)) a)
-      (Some (locale, map (map (Attrib.local_attribute thy) o snd o fst o Comment.ignore) args))
+      (Some (locale, map (map (Attrib.local_attribute thy) o snd o fst) args))
       (map (Locale.attribute (Attrib.local_attribute thy)) elems))
-    (map (apfst (apsnd (K [])) o Comment.ignore) args) int thy;
+    (map (apfst (apsnd (K []))) args) int thy;
 
 fun locale_multi_theorem_i k a locale elems args =
-  global_statement_i (Proof.multi_theorem_i k a
-      (Some (locale, map (snd o fst o Comment.ignore) args)) elems)
-    (map (apfst (apsnd (K [])) o Comment.ignore) args);
+  global_statement_i (Proof.multi_theorem_i k a (Some (locale, map (snd o fst) args)) elems)
+    (map (apfst (apsnd (K []))) args);
 
-fun theorem k ((a, t), cmt) = multi_theorem k a [] [((("", []), [t]), cmt)];
-fun theorem_i k ((a, t), cmt) = multi_theorem_i k a [] [((("", []), [t]), cmt)];
+fun theorem k (a, t) = multi_theorem k a [] [(("", []), [t])];
+fun theorem_i k (a, t) = multi_theorem_i k a [] [(("", []), [t])];
 
 fun smart_multi_theorem k a (None, elems) = multi_theorem k a elems
   | smart_multi_theorem k a (Some locale, elems) = locale_multi_theorem k a locale elems;
 
-val assume      = local_statement Proof.assume I o map Comment.ignore;
-val assume_i    = local_statement_i Proof.assume_i I o map Comment.ignore;
-val presume     = local_statement Proof.presume I o map Comment.ignore;
-val presume_i   = local_statement_i Proof.presume_i I o map Comment.ignore;
-val have        = local_statement (Proof.have Seq.single) I o map Comment.ignore;
-val have_i      = local_statement_i (Proof.have_i Seq.single) I o map Comment.ignore;
-val hence       = local_statement (Proof.have Seq.single) Proof.chain o map Comment.ignore;
-val hence_i     = local_statement_i (Proof.have_i Seq.single) Proof.chain o map Comment.ignore;
-val show        = local_statement' (Proof.show check_goal Seq.single) I o map Comment.ignore;
-val show_i      = local_statement_i' (Proof.show_i check_goal Seq.single) I o map Comment.ignore;
-val thus = local_statement' (Proof.show check_goal Seq.single) Proof.chain o map Comment.ignore;
-val thus_i = local_statement_i' (Proof.show_i check_goal Seq.single) Proof.chain o map Comment.ignore;
+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 Seq.single) I;
+val have_i      = local_statement_i (Proof.have_i Seq.single) I;
+val hence       = local_statement (Proof.have Seq.single) Proof.chain;
+val hence_i     = local_statement_i (Proof.have_i Seq.single) Proof.chain;
+val show        = local_statement' (Proof.show check_goal Seq.single) I;
+val show_i      = local_statement_i' (Proof.show_i check_goal Seq.single) I;
+val thus        = local_statement' (Proof.show check_goal Seq.single) Proof.chain;
+val thus_i      = local_statement_i' (Proof.show_i check_goal Seq.single) 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 local_def = gen_def Proof.def o Comment.ignore;
-val local_def_i = gen_def Proof.def_i o Comment.ignore;
+val local_def = gen_def Proof.def;
+val local_def_i = gen_def Proof.def_i;
 
 fun obtain (xs, asms) = ProofHistory.applys (fn state =>
-  Obtain.obtain (map Comment.ignore xs)
-    (map (local_attributes' state) (map Comment.ignore asms)) state);
+  Obtain.obtain xs (map (local_attributes' state) asms) state);
 
-fun obtain_i (xs, asms) = ProofHistory.applys (fn state =>
-  Obtain.obtain_i (map Comment.ignore xs) (map Comment.ignore asms) state);
+fun obtain_i (xs, asms) = ProofHistory.applys (Obtain.obtain_i xs asms);
 
 end;
 
 
 (* blocks *)
 
-fun begin_block comment_ignore = ProofHistory.apply Proof.begin_block;
-fun next_block comment_ignore = ProofHistory.apply Proof.next_block;
-fun end_block comment_ignore = ProofHistory.applys Proof.end_block;
+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, comment_ignore) = ProofHistory.applys (shuffle Method.defer i);
-fun prefer (i, comment_ignore) = ProofHistory.applys (shuffle Method.prefer i);
+fun defer i = ProofHistory.applys (shuffle Method.defer i);
+fun prefer i = ProofHistory.applys (shuffle Method.prefer i);
 
 
 (* backward steps *)
 
-fun apply (m, comment_ignore) = ProofHistory.applys
+fun apply m = ProofHistory.applys
   (Seq.map (Proof.goal_facts (K [])) o Method.refine m o Proof.assert_backward);
-
-fun apply_end (m, comment_ignore) = ProofHistory.applys
-  (Method.refine_end m o Proof.assert_forward);
-
-val proof = ProofHistory.applys o Method.proof o
-  apsome Comment.ignore_interest o Comment.ignore_interest' o Comment.ignore;
+fun apply_end m = ProofHistory.applys (Method.refine_end m o Proof.assert_forward);
+val proof = ProofHistory.applys o Method.proof;
 
 
 (* local endings *)
 
-val local_qed =
-  proof'' o (ProofHistory.applys oo Method.local_qed true) o apsome Comment.ignore_interest;
-
-fun ignore_interests (x, y) = (Comment.ignore_interest x, apsome Comment.ignore_interest y);
-
-val local_terminal_proof =
-  proof'' o (ProofHistory.applys oo Method.local_terminal_proof) o ignore_interests;
-
+val local_qed = proof'' o (ProofHistory.applys oo Method.local_qed true);
+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);
@@ -596,9 +499,8 @@
     thy
   end);
 
-fun global_qed m = global_result (K (Method.global_qed true (apsome Comment.ignore_interest m)));
-fun global_terminal_proof m =
-  global_result (K (Method.global_terminal_proof (ignore_interests m)));
+fun global_qed m = global_result (K (Method.global_qed true m));
+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);
 val global_skip_proof = global_result SkipProof.global_skip_proof;
@@ -607,14 +509,13 @@
 
 (* common endings *)
 
-fun qed (meth, comment) = local_qed meth o global_qed meth;
-fun terminal_proof (meths, comment) = local_terminal_proof meths o global_terminal_proof meths;
-fun default_proof comment = local_default_proof o global_default_proof;
-fun immediate_proof comment = local_immediate_proof o global_immediate_proof;
-fun done_proof comment = local_done_proof o global_done_proof;
-fun skip_proof comment = local_skip_proof o global_skip_proof;
-
-fun forget_proof comment = Toplevel.proof_to_theory (Proof.theory_of o ProofHistory.current);
+fun qed meth = local_qed meth o global_qed meth;
+fun terminal_proof meths = local_terminal_proof meths o global_terminal_proof meths;
+val default_proof = local_default_proof o global_default_proof;
+val immediate_proof = local_immediate_proof o global_immediate_proof;
+val done_proof = local_done_proof o global_done_proof;
+val skip_proof = local_skip_proof o global_skip_proof;
+val forget_proof = Toplevel.proof_to_theory (Proof.theory_of o ProofHistory.current);
 
 
 (* calculational proof commands *)
@@ -628,8 +529,8 @@
 
 fun proof''' f = Toplevel.proof' (f o cond_print_calc);
 
-fun gen_calc get f (args, _) prt state =
-  f (apsome (fn (srcs, _) => get srcs state) args) prt state;
+fun gen_calc get f args prt state =
+  f (apsome (fn srcs => get srcs state) args) prt state;
 
 in
 
@@ -640,8 +541,8 @@
 fun also_i x = proof''' (ProofHistory.applys o gen_calc get_thmss_i Calculation.also x);
 fun finally x = proof''' (ProofHistory.applys o gen_calc get_thmss Calculation.finally x);
 fun finally_i x = proof''' (ProofHistory.applys o gen_calc get_thmss_i Calculation.finally x);
-fun moreover comment = proof''' (ProofHistory.apply o Calculation.moreover);
-fun ultimately comment = proof''' (ProofHistory.apply o Calculation.ultimately);
+val moreover = proof''' (ProofHistory.apply o Calculation.moreover);
+val ultimately = proof''' (ProofHistory.apply o Calculation.ultimately);
 
 end;
 
@@ -650,32 +551,32 @@
 
 val parse_ast_translation =
   Context.use_let "val parse_ast_translation: (string * (Syntax.ast list -> Syntax.ast)) list"
-    "Theory.add_trfuns (parse_ast_translation, [], [], [])" o Comment.ignore;
+    "Theory.add_trfuns (parse_ast_translation, [], [], [])";
 
 val parse_translation =
   Context.use_let "val parse_translation: (string * (term list -> term)) list"
-    "Theory.add_trfuns ([], parse_translation, [], [])" o Comment.ignore;
+    "Theory.add_trfuns ([], parse_translation, [], [])";
 
 val print_translation =
   Context.use_let "val print_translation: (string * (term list -> term)) list"
-    "Theory.add_trfuns ([], [], print_translation, [])" o Comment.ignore;
+    "Theory.add_trfuns ([], [], print_translation, [])";
 
 val print_ast_translation =
   Context.use_let "val print_ast_translation: (string * (Syntax.ast list -> Syntax.ast)) list"
-    "Theory.add_trfuns ([], [], [], print_ast_translation)" o Comment.ignore;
+    "Theory.add_trfuns ([], [], [], print_ast_translation)";
 
 val typed_print_translation =
   Context.use_let "val typed_print_translation: (string * (bool -> typ -> term list -> term)) list"
-    "Theory.add_trfunsT typed_print_translation" o Comment.ignore;
+    "Theory.add_trfunsT typed_print_translation";
 
 val token_translation =
   Context.use_let "val token_translation: (string * string * (string -> string * real)) list"
-    "Theory.add_tokentrfuns token_translation" o Comment.ignore;
+    "Theory.add_tokentrfuns token_translation";
 
 
 (* add method *)
 
-fun method_setup ((name, txt, cmt), comment_ignore) =
+fun method_setup (name, txt, cmt) =
   Context.use_let
     "val thm = PureThy.get_thm_closure (Context.the_context ());\n\
     \val thms = PureThy.get_thms_closure (Context.the_context ());\n\
@@ -686,7 +587,7 @@
 
 (* add_oracle *)
 
-fun add_oracle ((name, txt), comment_ignore) =
+fun add_oracle (name, txt) =
   Context.use_let
     "val oracle: bstring * (Sign.sg * Object.T -> term)"
     "Theory.add_oracle oracle"
@@ -696,8 +597,7 @@
 (* add_locale *)
 
 fun add_locale (name, import, body) thy =
-  Locale.add_locale name import
-    (map (Locale.attribute (Attrib.local_attribute thy) o Comment.ignore) body) thy;
+  Locale.add_locale name import (map (Locale.attribute (Attrib.local_attribute thy)) body) thy;
 
 
 (* theory init and exit *)