--- a/doc-src/IsarImplementation/Thy/document/logic.tex Wed Dec 13 15:47:37 2006 +0100
+++ b/doc-src/IsarImplementation/Thy/document/logic.tex Wed Dec 13 16:26:45 2006 +0100
@@ -329,8 +329,7 @@
\indexml{lambda}\verb|lambda: term -> term -> term| \\
\indexml{betapply}\verb|betapply: term * term -> term| \\
\indexml{Sign.add-consts-i}\verb|Sign.add_consts_i: (string * typ * mixfix) list -> theory -> theory| \\
- \indexml{Sign.add-abbrevs}\verb|Sign.add_abbrevs: string * bool ->|\isasep\isanewline%
-\verb| ((string * mixfix) * term) list -> theory -> (term * term) list * theory| \\
+ \indexml{Sign.add-abbrev}\verb|Sign.add_abbrev: string -> bstring * term -> theory -> (term * term) * theory| \\
\indexml{Sign.const-typargs}\verb|Sign.const_typargs: theory -> string * typ -> typ list| \\
\indexml{Sign.const-instance}\verb|Sign.const_instance: theory -> string * typ list -> typ| \\
\end{mldecls}
@@ -370,9 +369,8 @@
\item \verb|Sign.add_consts_i|~\isa{{\isacharbrackleft}{\isacharparenleft}c{\isacharcomma}\ {\isasymsigma}{\isacharcomma}\ mx{\isacharparenright}{\isacharcomma}\ {\isasymdots}{\isacharbrackright}} declares a
new constant \isa{c\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} with optional mixfix syntax.
- \item \verb|Sign.add_abbrevs|~\isa{print{\isacharunderscore}mode\ {\isacharbrackleft}{\isacharparenleft}{\isacharparenleft}c{\isacharcomma}\ t{\isacharparenright}{\isacharcomma}\ mx{\isacharparenright}{\isacharcomma}\ {\isasymdots}{\isacharbrackright}}
- declares a new term abbreviation \isa{c\ {\isasymequiv}\ t} with optional
- mixfix syntax.
+ \item \verb|Sign.add_abbrev|~\isa{print{\isacharunderscore}mode\ {\isacharparenleft}c{\isacharcomma}\ t{\isacharparenright}}
+ introduces a new term abbreviation \isa{c\ {\isasymequiv}\ t}.
\item \verb|Sign.const_typargs|~\isa{thy\ {\isacharparenleft}c{\isacharcomma}\ {\isasymtau}{\isacharparenright}} and \verb|Sign.const_instance|~\isa{thy\ {\isacharparenleft}c{\isacharcomma}\ {\isacharbrackleft}{\isasymtau}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymtau}\isactrlisub n{\isacharbrackright}{\isacharparenright}}
convert between two representations of polymorphic constants: full
--- a/doc-src/IsarImplementation/Thy/document/proof.tex Wed Dec 13 15:47:37 2006 +0100
+++ b/doc-src/IsarImplementation/Thy/document/proof.tex Wed Dec 13 16:26:45 2006 +0100
@@ -356,8 +356,8 @@
\item \verb|Goal.prove_multi| is simular to \verb|Goal.prove|, but
states several conclusions simultaneously. The goal is encoded by
- means of Pure conjunction; \verb|Tactic.conjunction_tac| will turn
- this into a collection of individual subgoals.
+ means of Pure conjunction; \verb|Goal.conjunction_tac| will turn this
+ into a collection of individual subgoals.
\item \verb|Obtain.result|~\isa{tac\ thms\ ctxt} eliminates the
given facts using a tactic, which results in additional fixed
--- a/doc-src/IsarImplementation/Thy/logic.thy Wed Dec 13 15:47:37 2006 +0100
+++ b/doc-src/IsarImplementation/Thy/logic.thy Wed Dec 13 16:26:45 2006 +0100
@@ -327,8 +327,7 @@
@{index_ML lambda: "term -> term -> term"} \\
@{index_ML betapply: "term * term -> term"} \\
@{index_ML Sign.add_consts_i: "(string * typ * mixfix) list -> theory -> theory"} \\
- @{index_ML Sign.add_abbrevs: "string * bool ->
- ((string * mixfix) * term) list -> theory -> (term * term) list * theory"} \\
+ @{index_ML Sign.add_abbrev: "string -> bstring * term -> theory -> (term * term) * theory"} \\
@{index_ML Sign.const_typargs: "theory -> string * typ -> typ list"} \\
@{index_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\
\end{mldecls}
@@ -376,9 +375,8 @@
\item @{ML Sign.add_consts_i}~@{text "[(c, \<sigma>, mx), \<dots>]"} declares a
new constant @{text "c :: \<sigma>"} with optional mixfix syntax.
- \item @{ML Sign.add_abbrevs}~@{text "print_mode [((c, t), mx), \<dots>]"}
- declares a new term abbreviation @{text "c \<equiv> t"} with optional
- mixfix syntax.
+ \item @{ML Sign.add_abbrev}~@{text "print_mode (c, t)"}
+ introduces a new term abbreviation @{text "c \<equiv> t"}.
\item @{ML Sign.const_typargs}~@{text "thy (c, \<tau>)"} and @{ML
Sign.const_instance}~@{text "thy (c, [\<tau>\<^isub>1, \<dots>, \<tau>\<^isub>n])"}
--- a/doc-src/IsarImplementation/Thy/proof.thy Wed Dec 13 15:47:37 2006 +0100
+++ b/doc-src/IsarImplementation/Thy/proof.thy Wed Dec 13 16:26:45 2006 +0100
@@ -318,8 +318,8 @@
\item @{ML Goal.prove_multi} is simular to @{ML Goal.prove}, but
states several conclusions simultaneously. The goal is encoded by
- means of Pure conjunction; @{ML Tactic.conjunction_tac} will turn
- this into a collection of individual subgoals.
+ means of Pure conjunction; @{ML Goal.conjunction_tac} will turn this
+ into a collection of individual subgoals.
\item @{ML Obtain.result}~@{text "tac thms ctxt"} eliminates the
given facts using a tactic, which results in additional fixed