# HG changeset patch # User bulwahn # Date 1287986888 -7200 # Node ID 0d579da1902a015d7117a32ec3341446041803d3 # Parent 82873a6f2b813a1334bf214e3f38f7474b023c53# Parent 0fb7891f0c7c8458043d1b6fbb6fac1257cb974f merged diff -r 82873a6f2b81 -r 0d579da1902a NEWS --- a/NEWS Fri Oct 22 18:38:59 2010 +0200 +++ b/NEWS Mon Oct 25 08:08:08 2010 +0200 @@ -94,6 +94,8 @@ set_ext -> set_eqI INCOMPATIBILITY. +* Renamed lemma list: nat_number -> eval_nat_numeral + * Renamed class eq and constant eq (for code generation) to class equal and constant equal, plus renaming of related facts and various tuning. INCOMPATIBILITY. diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/IsaMakefile --- a/doc-src/IsarImplementation/IsaMakefile Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/IsaMakefile Mon Oct 25 08:08:08 2010 +0200 @@ -19,13 +19,13 @@ ## sessions -Thy: $(LOG)/Pure-Thy.gz +Thy: $(LOG)/HOL-Thy.gz -$(LOG)/Pure-Thy.gz: Thy/ROOT.ML Thy/Base.thy Thy/Integration.thy \ +$(LOG)/HOL-Thy.gz: Thy/ROOT.ML Thy/Base.thy Thy/Integration.thy \ Thy/Isar.thy Thy/Local_Theory.thy Thy/Logic.thy Thy/Prelim.thy \ Thy/Proof.thy Thy/Syntax.thy Thy/Tactic.thy Thy/ML.thy \ ../antiquote_setup.ML - @$(USEDIR) Pure Thy + @$(USEDIR) HOL Thy @rm -f Thy/document/isabelle.sty Thy/document/isabellesym.sty \ Thy/document/pdfsetup.sty Thy/document/session.tex @@ -33,4 +33,4 @@ ## clean clean: - @rm -f $(LOG)/Pure-Thy.gz + @rm -f $(LOG)/HOL-Thy.gz diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Makefile --- a/doc-src/IsarImplementation/Makefile Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Makefile Mon Oct 25 08:08:08 2010 +0200 @@ -11,8 +11,8 @@ NAME = implementation FILES = ../extra.sty ../iman.sty ../isabelle.sty ../isabellesym.sty \ - ../isar.sty ../manual.bib ../pdfsetup.sty ../proof.sty \ - Thy/document/Integration.tex Thy/document/Isar.tex \ + ../isar.sty ../manual.bib ../pdfsetup.sty ../proof.sty ../rail.sty \ + ../railsetup.sty Thy/document/Integration.tex Thy/document/Isar.tex \ Thy/document/Local_Theory.tex Thy/document/Logic.tex \ Thy/document/Prelim.tex Thy/document/Proof.tex \ Thy/document/Syntax.tex Thy/document/Tactic.tex implementation.tex \ @@ -22,6 +22,7 @@ $(NAME).dvi: $(FILES) isabelle_isar.eps $(LATEX) $(NAME) + $(RAIL) $(NAME) $(BIBTEX) $(NAME) $(LATEX) $(NAME) $(LATEX) $(NAME) @@ -33,6 +34,7 @@ $(NAME).pdf: $(FILES) isabelle_isar.pdf $(PDFLATEX) $(NAME) + $(RAIL) $(NAME) $(BIBTEX) $(NAME) $(PDFLATEX) $(NAME) $(PDFLATEX) $(NAME) diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/Base.thy --- a/doc-src/IsarImplementation/Thy/Base.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/Base.thy Mon Oct 25 08:08:08 2010 +0200 @@ -1,6 +1,12 @@ theory Base -imports Pure +imports Main uses "../../antiquote_setup.ML" begin +(* FIXME move to src/Pure/ML/ml_antiquote.ML *) +ML {* + ML_Antiquote.inline "assert" + (Scan.succeed "(fn b => if b then () else raise General.Fail \"Assertion failed\")") +*} + end diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/Integration.thy --- a/doc-src/IsarImplementation/Thy/Integration.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/Integration.thy Mon Oct 25 08:08:08 2010 +0200 @@ -8,13 +8,12 @@ text {* The Isar toplevel may be considered the centeral hub of the Isabelle/Isar system, where all key components and sub-systems are - integrated into a single read-eval-print loop of Isar commands. We - shall even incorporate the existing {\ML} toplevel of the compiler - and run-time system (cf.\ \secref{sec:ML-toplevel}). + integrated into a single read-eval-print loop of Isar commands, + which also incorporates the underlying ML compiler. Isabelle/Isar departs from the original ``LCF system architecture'' - where {\ML} was really The Meta Language for defining theories and - conducting proofs. Instead, {\ML} now only serves as the + where ML was really The Meta Language for defining theories and + conducting proofs. Instead, ML now only serves as the implementation language for the system (and user extensions), while the specific Isar toplevel supports the concepts of theory and proof development natively. This includes the graph structure of theories @@ -66,11 +65,11 @@ \begin{description} - \item @{ML_type Toplevel.state} represents Isar toplevel states, - which are normally manipulated through the concept of toplevel - transitions only (\secref{sec:toplevel-transition}). Also note that - a raw toplevel state is subject to the same linearity restrictions - as a theory context (cf.~\secref{sec:context-theory}). + \item Type @{ML_type Toplevel.state} represents Isar toplevel + states, which are normally manipulated through the concept of + toplevel transitions only (\secref{sec:toplevel-transition}). Also + note that a raw toplevel state is subject to the same linearity + restrictions as a theory context (cf.~\secref{sec:context-theory}). \item @{ML Toplevel.UNDEF} is raised for undefined toplevel operations. Many operations work only partially for certain cases, @@ -94,13 +93,29 @@ information for each Isar command being executed. \item @{ML Toplevel.profiling}~@{verbatim ":="}~@{text "n"} controls - low-level profiling of the underlying {\ML} runtime system. For + low-level profiling of the underlying ML runtime system. For Poly/ML, @{text "n = 1"} means time and @{text "n = 2"} space profiling. \end{description} *} +text %mlantiq {* + \begin{matharray}{rcl} + @{ML_antiquotation_def "Isar.state"} & : & @{text ML_antiquotation} \\ + \end{matharray} + + \begin{description} + + \item @{text "@{Isar.state}"} refers to Isar toplevel state at that + point --- as abstract value. + + This only works for diagnostic ML commands, such as @{command + ML_val} or @{command ML_command}. + + \end{description} +*} + subsection {* Toplevel transitions \label{sec:toplevel-transition} *} @@ -182,123 +197,6 @@ *} -subsection {* Toplevel control *} - -text {* - There are a few special control commands that modify the behavior - the toplevel itself, and only make sense in interactive mode. Under - normal circumstances, the user encounters these only implicitly as - part of the protocol between the Isabelle/Isar system and a - user-interface such as Proof~General. - - \begin{description} - - \item \isacommand{undo} follows the three-level hierarchy of empty - toplevel vs.\ theory vs.\ proof: undo within a proof reverts to the - previous proof context, undo after a proof reverts to the theory - before the initial goal statement, undo of a theory command reverts - to the previous theory value, undo of a theory header discontinues - the current theory development and removes it from the theory - database (\secref{sec:theory-database}). - - \item \isacommand{kill} aborts the current level of development: - kill in a proof context reverts to the theory before the initial - goal statement, kill in a theory context aborts the current theory - development, removing it from the database. - - \item \isacommand{exit} drops out of the Isar toplevel into the - underlying {\ML} toplevel (\secref{sec:ML-toplevel}). The Isar - toplevel state is preserved and may be continued later. - - \item \isacommand{quit} terminates the Isabelle/Isar process without - saving. - - \end{description} -*} - - -section {* ML toplevel \label{sec:ML-toplevel} *} - -text {* - The {\ML} toplevel provides a read-compile-eval-print loop for {\ML} - values, types, structures, and functors. {\ML} declarations operate - on the global system state, which consists of the compiler - environment plus the values of {\ML} reference variables. There is - no clean way to undo {\ML} declarations, except for reverting to a - previously saved state of the whole Isabelle process. {\ML} input - is either read interactively from a TTY, or from a string (usually - within a theory text), or from a source file (usually loaded from a - theory). - - Whenever the {\ML} toplevel is active, the current Isabelle theory - context is passed as an internal reference variable. Thus {\ML} - code may access the theory context during compilation, it may even - change the value of a theory being under construction --- while - observing the usual linearity restrictions - (cf.~\secref{sec:context-theory}). -*} - -text %mlref {* - \begin{mldecls} - @{index_ML ML_Context.the_generic_context: "unit -> Context.generic"} \\ - @{index_ML "Context.>> ": "(Context.generic -> Context.generic) -> unit"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML "ML_Context.the_generic_context ()"} refers to the theory - context of the {\ML} toplevel --- at compile time! {\ML} code needs - to take care to refer to @{ML "ML_Context.the_generic_context ()"} - correctly. Recall that evaluation of a function body is delayed - until actual runtime. Moreover, persistent {\ML} toplevel bindings - to an unfinished theory should be avoided: code should either - project out the desired information immediately, or produce an - explicit @{ML_type theory_ref} (cf.\ \secref{sec:context-theory}). - - \item @{ML "Context.>>"}~@{text f} applies context transformation - @{text f} to the implicit context of the {\ML} toplevel. - - \end{description} - - It is very important to note that the above functions are really - restricted to the compile time, even though the {\ML} compiler is - invoked at runtime! The majority of {\ML} code uses explicit - functional arguments of a theory or proof context instead. Thus it - may be invoked for an arbitrary context later on, without having to - worry about any operational details. - - \bigskip - - \begin{mldecls} - @{index_ML Isar.main: "unit -> unit"} \\ - @{index_ML Isar.loop: "unit -> unit"} \\ - @{index_ML Isar.state: "unit -> Toplevel.state"} \\ - @{index_ML Isar.exn: "unit -> (exn * string) option"} \\ - @{index_ML Isar.goal: "unit -> - {context: Proof.context, facts: thm list, goal: thm}"} \\ - \end{mldecls} - - \begin{description} - - \item @{ML "Isar.main ()"} invokes the Isar toplevel from {\ML}, - initializing an empty toplevel state. - - \item @{ML "Isar.loop ()"} continues the Isar toplevel with the - current state, after having dropped out of the Isar toplevel loop. - - \item @{ML "Isar.state ()"} and @{ML "Isar.exn ()"} get current - toplevel state and error condition, respectively. This only works - after having dropped out of the Isar toplevel loop. - - \item @{ML "Isar.goal ()"} produces the full Isar goal state, - consisting of proof context, facts that have been indicated for - immediate use, and the tactical goal according to - \secref{sec:tactical-goals}. - - \end{description} -*} - - section {* Theory database \label{sec:theory-database} *} text {* @@ -315,10 +213,10 @@ Theory @{text A} is associated with the main theory file @{text A}\verb,.thy,, which needs to be accessible through the theory - loader path. Any number of additional {\ML} source files may be + loader path. Any number of additional ML source files may be associated with each theory, by declaring these dependencies in the theory header as @{text \}, and loading them consecutively - within the theory context. The system keeps track of incoming {\ML} + within the theory context. The system keeps track of incoming ML sources and associates them with the current theory. The basic internal actions of the theory database are @{text diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/Isar.thy --- a/doc-src/IsarImplementation/Thy/Isar.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/Isar.thy Mon Oct 25 08:08:08 2010 +0200 @@ -5,16 +5,35 @@ chapter {* Isar language elements *} text {* - The primary Isar language consists of three main categories of - language elements: + The Isar proof language (see also \cite[\S2]{isabelle-isar-ref}) + consists of three main categories of language elements: \begin{enumerate} - \item Proof commands + \item Proof \emph{commands} define the primary language of + transactions of the underlying Isar/VM interpreter. Typical + examples are @{command "fix"}, @{command "assume"}, @{command + "show"}, @{command "proof"}, and @{command "qed"}. + + Composing proof commands according to the rules of the Isar/VM leads + to expressions of structured proof text, such that both the machine + and the human reader can give it a meaning as formal reasoning. - \item Proof methods + \item Proof \emph{methods} define a secondary language of mixed + forward-backward refinement steps involving facts and goals. + Typical examples are @{method rule}, @{method unfold}, and @{method + simp}. - \item Attributes + Methods can occur in certain well-defined parts of the Isar proof + language, say as arguments to @{command "proof"}, @{command "qed"}, + or @{command "by"}. + + \item \emph{Attributes} define a tertiary language of small + annotations to theorems being defined or referenced. Attributes can + modify both the context and the theorem. + + Typical examples are @{attribute intro} (which affects the context), + and @{attribute symmetric} (which affects the theorem). \end{enumerate} *} @@ -22,16 +41,523 @@ section {* Proof commands *} -text FIXME +text {* A \emph{proof command} is state transition of the Isar/VM + proof interpreter. + + In principle, Isar proof commands could be defined in + user-space as well. The system is built like that in the first + place: part of the commands are primitive, the other part is defined + as derived elements. Adding to the genuine structured proof + language requires profound understanding of the Isar/VM machinery, + though, so this is beyond the scope of this manual. + + What can be done realistically is to define some diagnostic commands + that inspect the general state of the Isar/VM, and report some + feedback to the user. Typically this involves checking of the + linguistic \emph{mode} of a proof state, or peeking at the pending + goals (if available). + + Another common application is to define a toplevel command that + poses a problem to the user as Isar proof state and stores the final + result in a suitable context data slot. Thus a proof can be + incorporated into the context of some user-space tool, without + modifying the Isar proof language itself. +*} + +text %mlref {* + \begin{mldecls} + @{index_ML_type Proof.state} \\ + @{index_ML Proof.assert_forward: "Proof.state -> Proof.state"} \\ + @{index_ML Proof.assert_chain: "Proof.state -> Proof.state"} \\ + @{index_ML Proof.assert_backward: "Proof.state -> Proof.state"} \\ + @{index_ML Proof.simple_goal: "Proof.state -> {context: Proof.context, goal: thm}"} \\ + @{index_ML Proof.goal: "Proof.state -> + {context: Proof.context, facts: thm list, goal: thm}"} \\ + @{index_ML Proof.raw_goal: "Proof.state -> + {context: Proof.context, facts: thm list, goal: thm}"} \\ + @{index_ML Proof.theorem: "Method.text option -> + (thm list list -> Proof.context -> Proof.context) -> + (term * term list) list list -> Proof.context -> Proof.state"} \\ + \end{mldecls} + + \begin{description} + + \item Type @{ML_type Proof.state} represents Isar proof states. + This is a block-structured configuration with proof context, + linguistic mode, and optional goal. The latter consists of goal + context, goal facts (``@{text "using"}''), and tactical goal state + (see \secref{sec:tactical-goals}). + + The general idea is that the facts shall contribute to the + refinement of some parts of the tactical goal --- how exactly is + defined by the proof method that is applied in that situation. + + \item @{ML Proof.assert_forward}, @{ML Proof.assert_chain}, @{ML + Proof.assert_backward} are partial identity functions that fail + unless a certain linguistic mode is active, namely ``@{text + "proof(state)"}'', ``@{text "proof(chain)"}'', ``@{text + "proof(prove)"}'', respectively (using the terminology of + \cite{isabelle-isar-ref}). + + It is advisable study the implementations of existing proof commands + for suitable modes to be asserted. + + \item @{ML Proof.simple_goal}~@{text "state"} returns the structured + Isar goal (if available) in the form seen by ``simple'' methods + (like @{method simp} or @{method blast}). The Isar goal facts are + already inserted as premises into the subgoals, which are presented + individually as in @{ML Proof.goal}. + + \item @{ML Proof.goal}~@{text "state"} returns the structured Isar + goal (if available) in the form seen by regular methods (like + @{method rule}). The auxiliary internal encoding of Pure + conjunctions is split into individual subgoals as usual. + + \item @{ML Proof.raw_goal}~@{text "state"} returns the structured + Isar goal (if available) in the raw internal form seen by ``raw'' + methods (like @{method induct}). This form is rarely appropriate + for dignostic tools; @{ML Proof.simple_goal} or @{ML Proof.goal} + should be used in most situations. + + \item @{ML Proof.theorem}~@{text "before_qed after_qed statement ctxt"} + initializes a toplevel Isar proof state within a given context. + + The optional @{text "before_qed"} method is applied at the end of + the proof, just before extracting the result (this feature is rarely + used). + + The @{text "after_qed"} continuation receives the extracted result + in order to apply it to the final context in a suitable way (e.g.\ + storing named facts). Note that at this generic level the target + context is specified as @{ML_type Proof.context}, but the usual + wrapping of toplevel proofs into command transactions will provide a + @{ML_type local_theory} here (see also \chref{ch:local-theory}). + This usually affects the way how results are stored. + + The @{text "statement"} is given as a nested list of terms, each + associated with optional @{keyword "is"} patterns as usual in the + Isar source language. The original list structure over terms is + turned into one over theorems when @{text "after_qed"} is invoked. + + \end{description} +*} + + +text %mlantiq {* + \begin{matharray}{rcl} + @{ML_antiquotation_def "Isar.goal"} & : & @{text ML_antiquotation} \\ + \end{matharray} + + \begin{description} + + \item @{text "@{Isar.goal}"} refers to the regular goal state (if + available) of the current proof state managed by the Isar toplevel + --- as abstract value. + + This only works for diagnostic ML commands, such as @{command + ML_val} or @{command ML_command}. + + \end{description} +*} + +text %mlex {* The following example peeks at a certain goal configuration. *} + +example_proof + have A and B and C + ML_val {* + val n = Thm.nprems_of (#goal @{Isar.goal}); + @{assert} (n = 3); + *} + oops + +text {* Here we see 3 individual subgoals in the same way as regular + proof methods would do. *} section {* Proof methods *} -text FIXME +text {* A @{text "method"} is a function @{text "context \ thm\<^sup>* \ goal + \ (cases \ goal)\<^sup>*\<^sup>*"} that operates on the full Isar goal + configuration with context, goal facts, and tactical goal state and + enumerates possible follow-up goal states, with the potential + addition of named extensions of the proof context (\emph{cases}). + The latter feature is rarely used. + + This means a proof method is like a structurally enhanced tactic + (cf.\ \secref{sec:tactics}). The well-formedness conditions for + tactics need to hold for methods accordingly, with the following + additions. + + \begin{itemize} + + \item Goal addressing is further limited either to operate either + uniformly on \emph{all} subgoals, or specifically on the + \emph{first} subgoal. + + Exception: old-style tactic emulations that are embedded into the + method space, e.g.\ @{method rule_tac}. + + \item A non-trivial method always needs to make progress: an + identical follow-up goal state has to be avoided.\footnote{This + enables the user to write method expressions like @{text "meth\<^sup>+"} + without looping, while the trivial do-nothing case can be recovered + via @{text "meth\<^sup>?"}.} + + Exception: trivial stuttering steps, such as ``@{method -}'' or + @{method succeed}. + + \item Goal facts passed to the method must not be ignored. If there + is no sensible use of facts outside the goal state, facts should be + inserted into the subgoals that are addressed by the method. + + \end{itemize} + + \medskip Syntactically, the language of proof methods is embedded + into Isar as arguments to certain commands, e.g.\ @{command "by"} or + @{command apply}. User-space additions are reasonably easy by + plugging suitable method-valued parser functions into the framework, + using the @{command method_setup} command, for example. + + To get a better idea about the range of possibilities, consider the + following Isar proof schemes. First some general form of structured + proof text: + + \medskip + \begin{tabular}{l} + @{command from}~@{text "facts\<^sub>1"}~@{command have}~@{text "props"}~@{command using}~@{text "facts\<^sub>2"} \\ + @{command proof}~@{text "(initial_method)"} \\ + \quad@{text "body"} \\ + @{command qed}~@{text "(terminal_method)"} \\ + \end{tabular} + \medskip + + The goal configuration consists of @{text "facts\<^sub>1"} and @{text + "facts\<^sub>2"} appended in that order, and various @{text "props"} being + claimed here. The @{text "initial_method"} is invoked with facts + and goals together and refines the problem to something that is + handled recursively in the proof @{text "body"}. The @{text + "terminal_method"} has another chance to finish-off any remaining + subgoals, but it does not see the facts of the initial step. + + \medskip The second pattern illustrates unstructured proof scripts: + + \medskip + \begin{tabular}{l} + @{command have}~@{text "props"} \\ + \quad@{command using}~@{text "facts\<^sub>1"}~@{command apply}~@{text "method\<^sub>1"} \\ + \quad@{command apply}~@{text "method\<^sub>2"} \\ + \quad@{command using}~@{text "facts\<^sub>3"}~@{command apply}~@{text "method\<^sub>3"} \\ + \quad@{command done} \\ + \end{tabular} + \medskip + + The @{text "method\<^sub>1"} operates on the original claim together while + using @{text "facts\<^sub>1"}. Since the @{command apply} command + structurally resets the facts, the @{text "method\<^sub>2"} will operate on + the remaining goal state without facts. The @{text "method\<^sub>3"} will + see again a collection of @{text "facts\<^sub>3"} that has been inserted + into the script explicitly. + + \medskip Empirically, Isar proof methods can be categorized as + follows. + + \begin{enumerate} + + \item \emph{Special method with cases} with named context additions + associated with the follow-up goal state. + + Example: @{method "induct"}, which is also a ``raw'' method since it + operates on the internal representation of simultaneous claims as + Pure conjunction, instead of separate subgoals. + + \item \emph{Structured method} with strong emphasis on facts outside + the goal state. + + Example: @{method "rule"}, which captures the key ideas behind + structured reasoning in Isar in purest form. + + \item \emph{Simple method} with weaker emphasis on facts, which are + inserted into subgoals to emulate old-style tactical as + ``premises''. + + Examples: @{method "simp"}, @{method "blast"}, @{method "auto"}. + + \item \emph{Old-style tactic emulation} with detailed numeric goal + addressing and explicit references to entities of the internal goal + state (which are otherwise invisible from proper Isar proof text). + To make the special non-standard status clear, the naming convention + @{text "foo_tac"} needs to be observed. + + Example: @{method "rule_tac"}. + + \end{enumerate} + + When implementing proof methods, it is advisable to study existing + implementations carefully and imitate the typical ``boiler plate'' + for context-sensitive parsing and further combinators to wrap-up + tactic expressions as methods.\footnote{Aliases or abbreviations of + the standard method combinators should be avoided. Note that from + Isabelle99 until Isabelle2009 the system did provide various odd + combinations of method wrappers that made user applications more + complicated than necessary.} +*} + +text %mlref {* + \begin{mldecls} + @{index_ML_type Proof.method} \\ + @{index_ML METHOD_CASES: "(thm list -> cases_tactic) -> Proof.method"} \\ + @{index_ML METHOD: "(thm list -> tactic) -> Proof.method"} \\ + @{index_ML SIMPLE_METHOD: "tactic -> Proof.method"} \\ + @{index_ML SIMPLE_METHOD': "(int -> tactic) -> Proof.method"} \\ + @{index_ML HEADGOAL: "(int -> tactic) -> tactic"} \\ + @{index_ML Method.insert_tac: "thm list -> int -> tactic"} \\ + @{index_ML Method.setup: "binding -> (Proof.context -> Proof.method) context_parser -> + string -> theory -> theory"} \\ + \end{mldecls} + + \begin{description} + + \item Type @{ML_type Proof.method} represents proof methods as + abstract type. + + \item @{ML METHOD_CASES}~@{text "(fn facts => cases_tactic)"} wraps + @{text cases_tactic} depending on goal facts as proof method with + cases; the goal context is passed via method syntax. + + \item @{ML METHOD}~@{text "(fn facts => tactic)"} wraps @{text + tactic} depending on goal facts as regular proof method; the goal + context is passed via method syntax. + + \item @{ML SIMPLE_METHOD}~@{text "tactic"} wraps a tactic that + addresses all subgoals uniformly as simple proof method. Goal facts + are already inserted into all subgoals before @{text "tactic"} is + applied. + + \item @{ML SIMPLE_METHOD'}~@{text "tactic"} wraps a tactic that + addresses a specific subgoal as simple proof method. Goal facts are + already inserted into the first subgoal before @{text "tactic"} is + applied to the same. + + \item @{ML HEADGOAL}~@{text "tactic"} applies @{text "tactic"} to + the first subgoal. This is convenient to reproduce part the @{ML + SIMPLE_METHOD'} wrapping within regular @{ML METHOD}, for example. + + \item @{ML Method.insert_tac}~@{text "facts i"} inserts @{text + "facts"} into subgoal @{text "i"}. This is convenient to reproduce + part of the @{ML SIMPLE_METHOD} or @{ML SIMPLE_METHOD'} wrapping + within regular @{ML METHOD}, for example. + + \item @{ML Method.setup}~@{text "name parser description"} provides + the functionality of the Isar command @{command method_setup} as ML + function. + + \end{description} +*} + +text %mlex {* See also @{command method_setup} in + \cite{isabelle-isar-ref} which includes some abstract examples. + + \medskip The following toy examples illustrate how the goal facts + and state are passed to proof methods. The pre-defined proof method + called ``@{method tactic}'' wraps ML source of type @{ML_type + tactic} (abstracted over @{verbatim facts}). This allows immediate + experimentation without parsing of concrete syntax. *} + +example_proof + assume a: A and b: B + + have "A \ B" + apply (tactic {* rtac @{thm conjI} 1 *}) + using a apply (tactic {* resolve_tac facts 1 *}) + using b apply (tactic {* resolve_tac facts 1 *}) + done + + have "A \ B" + using a and b + ML_val "@{Isar.goal}" + apply (tactic {* Method.insert_tac facts 1 *}) + apply (tactic {* (rtac @{thm conjI} THEN_ALL_NEW atac) 1 *}) + done +qed + +text {* \medskip The next example implements a method that simplifies + the first subgoal by rewrite rules given as arguments. *} + +method_setup my_simp = {* + Attrib.thms >> (fn thms => fn ctxt => + SIMPLE_METHOD' (fn i => + CHANGED (asm_full_simp_tac + (HOL_basic_ss addsimps thms) i))) +*} "rewrite subgoal by given rules" + +text {* The concrete syntax wrapping of @{command method_setup} always + passes-through the proof context at the end of parsing, but it is + not used in this example. + + The @{ML Attrib.thms} parser produces a list of theorems from the + usual Isar syntax involving attribute expressions etc.\ (syntax + category @{syntax thmrefs}). The resulting @{verbatim thms} are + added to @{ML HOL_basic_ss} which already contains the basic + Simplifier setup for HOL. + + The tactic @{ML asm_full_simp_tac} is the one that is also used in + method @{method simp} by default. The extra wrapping by the @{ML + CHANGED} tactical ensures progress of simplification: identical goal + states are filtered out explicitly to make the raw tactic conform to + standard Isar method behaviour. + + \medskip Method @{method my_simp} can be used in Isar proofs like + this: +*} + +example_proof + fix a b c + assume a: "a = b" + assume b: "b = c" + have "a = c" by (my_simp a b) +qed + +text {* Here is a similar method that operates on all subgoals, + instead of just the first one. *} + +method_setup my_simp_all = {* + Attrib.thms >> (fn thms => fn ctxt => + SIMPLE_METHOD + (CHANGED + (ALLGOALS (asm_full_simp_tac + (HOL_basic_ss addsimps thms))))) +*} "rewrite all subgoals by given rules" + +example_proof + fix a b c + assume a: "a = b" + assume b: "b = c" + have "a = c" and "c = b" by (my_simp_all a b) + +qed + +text {* \medskip Apart from explicit arguments, common proof methods + typically work with a default configuration provided by the context. + As a shortcut to rule management we use a cheap solution via functor + @{ML_functor Named_Thms} (see also @{"file" + "~~/src/Pure/Tools/named_thms.ML"}). *} + +ML {* + structure My_Simps = + Named_Thms + (val name = "my_simp" val description = "my_simp rule") +*} +setup My_Simps.setup + +text {* This provides ML access to a list of theorems in canonical + declaration order via @{ML My_Simps.get}. The user can add or + delete rules via the attribute @{attribute my_simp}. The actual + proof method is now defined as before, but we append the explicit + arguments and the rules from the context. *} + +method_setup my_simp' = {* + Attrib.thms >> (fn thms => fn ctxt => + SIMPLE_METHOD' (fn i => + CHANGED (asm_full_simp_tac + (HOL_basic_ss addsimps (thms @ My_Simps.get ctxt)) i))) +*} "rewrite subgoal by given rules and my_simp rules from the context" + +text {* + \medskip Method @{method my_simp'} can be used in Isar proofs + like this: +*} + +example_proof + fix a b c + assume [my_simp]: "a \ b" + assume [my_simp]: "b \ c" + have "a \ c" by my_simp' +qed + +text {* \medskip The @{method my_simp} variants defined above are + ``simple'' methods, i.e.\ the goal facts are merely inserted as goal + premises by the @{ML SIMPLE_METHOD'} or @{ML SIMPLE_METHOD} wrapper. + For proof methods that are similar to the standard collection of + @{method simp}, @{method blast}, @{method auto} little more can be + done here. + + Note that using the primary goal facts in the same manner as the + method arguments obtained via concrete syntax or the context does + not meet the requirement of ``strong emphasis on facts'' of regular + proof methods, because rewrite rules as used above can be easily + ignored. A proof text ``@{command using}~@{text "foo"}~@{command + "by"}~@{text "my_simp"}'' where @{text "foo"} is not used would + deceive the reader. + + \medskip The technical treatment of rules from the context requires + further attention. Above we rebuild a fresh @{ML_type simpset} from + the arguments and \emph{all} rules retrieved from the context on + every invocation of the method. This does not scale to really large + collections of rules, which easily emerges in the context of a big + theory library, for example. + + This is an inherent limitation of the simplistic rule management via + functor @{ML_functor Named_Thms}, because it lacks tool-specific + storage and retrieval. More realistic applications require + efficient index-structures that organize theorems in a customized + manner, such as a discrimination net that is indexed by the + left-hand sides of rewrite rules. For variations on the Simplifier, + re-use of the existing type @{ML_type simpset} is adequate, but + scalability requires it be maintained statically within the context + data, not dynamically on each tool invocation. *} -section {* Attributes *} +section {* Attributes \label{sec:attributes} *} + +text {* An \emph{attribute} is a function @{text "context \ thm \ + context \ thm"}, which means both a (generic) context and a theorem + can be modified simultaneously. In practice this fully general form + is very rare, instead attributes are presented either as + \emph{declaration attribute:} @{text "thm \ context \ context"} or + \emph{rule attribute:} @{text "context \ thm \ thm"}. + + Attributes can have additional arguments via concrete syntax. There + is a collection of context-sensitive parsers for various logical + entities (types, terms, theorems). These already take care of + applying morphisms to the arguments when attribute expressions are + moved into a different context (see also \secref{sec:morphisms}). + + When implementing declaration attributes, it is important to operate + exactly on the variant of the generic context that is provided by + the system, which is either global theory context or local proof + context. In particular, the background theory of a local context + must not be modified in this situation! *} -text FIXME +text %mlref {* + \begin{mldecls} + @{index_ML_type attribute: "Context.generic * thm -> Context.generic * thm"} \\ + @{index_ML Thm.rule_attribute: "(Context.generic -> thm -> thm) -> attribute"} \\ + @{index_ML Thm.declaration_attribute: " + (thm -> Context.generic -> Context.generic) -> attribute"} \\ + @{index_ML Attrib.setup: "binding -> attribute context_parser -> + string -> theory -> theory"} \\ + \end{mldecls} + + \begin{description} + + \item Type @{ML_type attribute} represents attributes as concrete + type alias. + + \item @{ML Thm.rule_attribute}~@{text "(fn context => rule)"} wraps + a context-dependent rule (mapping on @{ML_type thm}) as attribute. + + \item @{ML Thm.declaration_attribute}~@{text "(fn thm => decl)"} + wraps a theorem-dependent declaration (mapping on @{ML_type + Context.generic}) as attribute. + + \item @{ML Attrib.setup}~@{text "name parser description"} provides + the functionality of the Isar command @{command attribute_setup} as + ML function. + + \end{description} +*} + +text %mlex {* See also @{command attribute_setup} in + \cite{isabelle-isar-ref} which includes some abstract examples. *} end diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/Local_Theory.thy --- a/doc-src/IsarImplementation/Thy/Local_Theory.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/Local_Theory.thy Mon Oct 25 08:08:08 2010 +0200 @@ -31,8 +31,8 @@ Many definitional packages for local theories are available in Isabelle. Although a few old packages only work for global - theories, the local theory interface is already the standard way of - implementing definitional packages in Isabelle. + theories, the standard way of implementing definitional packages in + Isabelle is via the local theory interface. *} @@ -86,13 +86,12 @@ \framebox{\quad auxiliary context \quad\framebox{\quad target context \quad\framebox{\quad background theory\quad}}} \end{center} - \noindent When a definitional package is finished, the auxiliary - context is reset to the target context. The target now holds - definitions for terms and theorems that stem from the hypothetical - @{text "\"} and @{text "\"} elements, transformed by - the particular target policy (see - \cite[\S4--5]{Haftmann-Wenzel:2009} for details). -*} + When a definitional package is finished, the auxiliary context is + reset to the target context. The target now holds definitions for + terms and theorems that stem from the hypothetical @{text + "\"} and @{text "\"} elements, transformed by the + particular target policy (see \cite[\S4--5]{Haftmann-Wenzel:2009} + for details). *} text %mlref {* \begin{mldecls} @@ -106,8 +105,8 @@ \begin{description} - \item @{ML_type local_theory} represents local theories. Although - this is merely an alias for @{ML_type Proof.context}, it is + \item Type @{ML_type local_theory} represents local theories. + Although this is merely an alias for @{ML_type Proof.context}, it is semantically a subtype of the same: a @{ML_type local_theory} holds target information as special context data. Subtyping means that any value @{text "lthy:"}~@{ML_type local_theory} can be also used @@ -155,8 +154,11 @@ *} -section {* Morphisms and declarations *} +section {* Morphisms and declarations \label{sec:morphisms} *} -text FIXME +text {* FIXME + + \medskip See also \cite{Chaieb-Wenzel:2007}. +*} end diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/Logic.thy --- a/doc-src/IsarImplementation/Thy/Logic.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/Logic.thy Mon Oct 25 08:08:08 2010 +0200 @@ -121,8 +121,8 @@ @{index_ML_type sort: "class list"} \\ @{index_ML_type arity: "string * sort list * sort"} \\ @{index_ML_type typ} \\ - @{index_ML map_atyps: "(typ -> typ) -> typ -> typ"} \\ - @{index_ML fold_atyps: "(typ -> 'a -> 'a) -> typ -> 'a -> 'a"} \\ + @{index_ML Term.map_atyps: "(typ -> typ) -> typ -> typ"} \\ + @{index_ML Term.fold_atyps: "(typ -> 'a -> 'a) -> typ -> 'a -> 'a"} \\ \end{mldecls} \begin{mldecls} @{index_ML Sign.subsort: "theory -> sort * sort -> bool"} \\ @@ -136,26 +136,27 @@ \begin{description} - \item @{ML_type class} represents type classes. + \item Type @{ML_type class} represents type classes. - \item @{ML_type sort} represents sorts, i.e.\ finite intersections - of classes. The empty list @{ML "[]: sort"} refers to the empty - class intersection, i.e.\ the ``full sort''. + \item Type @{ML_type sort} represents sorts, i.e.\ finite + intersections of classes. The empty list @{ML "[]: sort"} refers to + the empty class intersection, i.e.\ the ``full sort''. - \item @{ML_type arity} represents type arities. A triple @{text - "(\, \<^vec>s, s) : arity"} represents @{text "\ :: (\<^vec>s)s"} as - described above. + \item Type @{ML_type arity} represents type arities. A triple + @{text "(\, \<^vec>s, s) : arity"} represents @{text "\ :: + (\<^vec>s)s"} as described above. - \item @{ML_type typ} represents types; this is a datatype with + \item Type @{ML_type typ} represents types; this is a datatype with constructors @{ML TFree}, @{ML TVar}, @{ML Type}. - \item @{ML map_atyps}~@{text "f \"} applies the mapping @{text "f"} - to all atomic types (@{ML TFree}, @{ML TVar}) occurring in @{text - "\"}. + \item @{ML Term.map_atyps}~@{text "f \"} applies the mapping @{text + "f"} to all atomic types (@{ML TFree}, @{ML TVar}) occurring in + @{text "\"}. - \item @{ML fold_atyps}~@{text "f \"} iterates the operation @{text - "f"} over all occurrences of atomic types (@{ML TFree}, @{ML TVar}) - in @{text "\"}; the type structure is traversed from left to right. + \item @{ML Term.fold_atyps}~@{text "f \"} iterates the operation + @{text "f"} over all occurrences of atomic types (@{ML TFree}, @{ML + TVar}) in @{text "\"}; the type structure is traversed from left to + right. \item @{ML Sign.subsort}~@{text "thy (s\<^isub>1, s\<^isub>2)"} tests the subsort relation @{text "s\<^isub>1 \ s\<^isub>2"}. @@ -185,6 +186,51 @@ \end{description} *} +text %mlantiq {* + \begin{matharray}{rcl} + @{ML_antiquotation_def "class"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "sort"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "type_name"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "type_abbrev"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "nonterminal"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "typ"} & : & @{text ML_antiquotation} \\ + \end{matharray} + + \begin{rail} + 'class' nameref + ; + 'sort' sort + ; + ('type\_name' | 'type\_abbrev' | 'nonterminal') nameref + ; + 'typ' type + ; + \end{rail} + + \begin{description} + + \item @{text "@{class c}"} inlines the internalized class @{text + "c"} --- as @{ML_type string} literal. + + \item @{text "@{sort s}"} inlines the internalized sort @{text "s"} + --- as @{ML_type "string list"} literal. + + \item @{text "@{type_name c}"} inlines the internalized type + constructor @{text "c"} --- as @{ML_type string} literal. + + \item @{text "@{type_abbrev c}"} inlines the internalized type + abbreviation @{text "c"} --- as @{ML_type string} literal. + + \item @{text "@{nonterminal c}"} inlines the internalized syntactic + type~/ grammar nonterminal @{text "c"} --- as @{ML_type string} + literal. + + \item @{text "@{typ \}"} inlines the internalized type @{text "\"} + --- as constructor term for datatype @{ML_type typ}. + + \end{description} +*} + section {* Terms \label{sec:terms} *} @@ -240,7 +286,7 @@ polymorphic constants that the user-level type-checker would reject due to violation of type class restrictions. - \medskip An \emph{atomic} term is either a variable or constant. + \medskip An \emph{atomic term} is either a variable or constant. The logical category \emph{term} is defined inductively over atomic terms, with abstraction and application as follows: @{text "t = b | x\<^isub>\ | ?x\<^isub>\ | c\<^isub>\ | \\<^isub>\. t | t\<^isub>1 t\<^isub>2"}. Parsing and printing takes care of @@ -309,10 +355,10 @@ \begin{mldecls} @{index_ML_type term} \\ @{index_ML "op aconv": "term * term -> bool"} \\ - @{index_ML map_types: "(typ -> typ) -> term -> term"} \\ - @{index_ML fold_types: "(typ -> 'a -> 'a) -> term -> 'a -> 'a"} \\ - @{index_ML map_aterms: "(term -> term) -> term -> term"} \\ - @{index_ML fold_aterms: "(term -> 'a -> 'a) -> term -> 'a -> 'a"} \\ + @{index_ML Term.map_types: "(typ -> typ) -> term -> term"} \\ + @{index_ML Term.fold_types: "(typ -> 'a -> 'a) -> term -> 'a -> 'a"} \\ + @{index_ML Term.map_aterms: "(term -> term) -> term -> term"} \\ + @{index_ML Term.fold_aterms: "(term -> 'a -> 'a) -> term -> 'a -> 'a"} \\ \end{mldecls} \begin{mldecls} @{index_ML fastype_of: "term -> typ"} \\ @@ -328,8 +374,8 @@ \begin{description} - \item @{ML_type term} represents de-Bruijn terms, with comments in - abstractions, and explicitly named free variables and constants; + \item Type @{ML_type term} represents de-Bruijn terms, with comments + in abstractions, and explicitly named free variables and constants; this is a datatype with constructors @{ML Bound}, @{ML Free}, @{ML Var}, @{ML Const}, @{ML Abs}, @{ML "op $"}. @@ -338,20 +384,20 @@ on type @{ML_type term}; raw datatype equality should only be used for operations related to parsing or printing! - \item @{ML map_types}~@{text "f t"} applies the mapping @{text + \item @{ML Term.map_types}~@{text "f t"} applies the mapping @{text "f"} to all types occurring in @{text "t"}. - \item @{ML fold_types}~@{text "f t"} iterates the operation @{text - "f"} over all occurrences of types in @{text "t"}; the term + \item @{ML Term.fold_types}~@{text "f t"} iterates the operation + @{text "f"} over all occurrences of types in @{text "t"}; the term structure is traversed from left to right. - \item @{ML map_aterms}~@{text "f t"} applies the mapping @{text "f"} - to all atomic terms (@{ML Bound}, @{ML Free}, @{ML Var}, @{ML + \item @{ML Term.map_aterms}~@{text "f t"} applies the mapping @{text + "f"} to all atomic terms (@{ML Bound}, @{ML Free}, @{ML Var}, @{ML Const}) occurring in @{text "t"}. - \item @{ML fold_aterms}~@{text "f t"} iterates the operation @{text - "f"} over all occurrences of atomic terms (@{ML Bound}, @{ML Free}, - @{ML Var}, @{ML Const}) in @{text "t"}; the term structure is + \item @{ML Term.fold_aterms}~@{text "f t"} iterates the operation + @{text "f"} over all occurrences of atomic terms (@{ML Bound}, @{ML + Free}, @{ML Var}, @{ML Const}) in @{text "t"}; the term structure is traversed from left to right. \item @{ML fastype_of}~@{text "t"} determines the type of a @@ -381,6 +427,49 @@ \end{description} *} +text %mlantiq {* + \begin{matharray}{rcl} + @{ML_antiquotation_def "const_name"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "const_abbrev"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "const"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "term"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "prop"} & : & @{text ML_antiquotation} \\ + \end{matharray} + + \begin{rail} + ('const\_name' | 'const\_abbrev') nameref + ; + 'const' ('(' (type + ',') ')')? + ; + 'term' term + ; + 'prop' prop + ; + \end{rail} + + \begin{description} + + \item @{text "@{const_name c}"} inlines the internalized logical + constant name @{text "c"} --- as @{ML_type string} literal. + + \item @{text "@{const_abbrev c}"} inlines the internalized + abbreviated constant name @{text "c"} --- as @{ML_type string} + literal. + + \item @{text "@{const c(\<^vec>\)}"} inlines the internalized + constant @{text "c"} with precise type instantiation in the sense of + @{ML Sign.const_instance} --- as @{ML Const} constructor term for + datatype @{ML_type term}. + + \item @{text "@{term t}"} inlines the internalized term @{text "t"} + --- as constructor term for datatype @{ML_type term}. + + \item @{text "@{prop \}"} inlines the internalized proposition + @{text "\"} --- as constructor term for datatype @{ML_type term}. + + \end{description} +*} + section {* Theorems \label{sec:thms} *} @@ -531,7 +620,7 @@ "d(\\<^isub>i)"} for some @{text "\\<^isub>i"} projected from @{text "\<^vec>\"}. Thus overloaded definitions essentially work by primitive recursion over the syntactic structure of a single type - argument. + argument. See also \cite[\S4.3]{Haftmann-Wenzel:2006:classes}. *} text %mlref {* @@ -552,18 +641,20 @@ @{index_ML Thm.generalize: "string list * string list -> int -> thm -> thm"} \\ @{index_ML Thm.instantiate: "(ctyp * ctyp) list * (cterm * cterm) list -> thm -> thm"} \\ @{index_ML Thm.add_axiom: "binding * term -> theory -> (string * thm) * theory"} \\ - @{index_ML Thm.add_oracle: "binding * ('a -> cterm) -> theory - -> (string * ('a -> thm)) * theory"} \\ - @{index_ML Thm.add_def: "bool -> bool -> binding * term -> theory -> (string * thm) * theory"} \\ + @{index_ML Thm.add_oracle: "binding * ('a -> cterm) -> theory -> + (string * ('a -> thm)) * theory"} \\ + @{index_ML Thm.add_def: "bool -> bool -> binding * term -> theory -> + (string * thm) * theory"} \\ \end{mldecls} \begin{mldecls} - @{index_ML Theory.add_deps: "string -> string * typ -> (string * typ) list -> theory -> theory"} \\ + @{index_ML Theory.add_deps: "string -> string * typ -> (string * typ) list -> + theory -> theory"} \\ \end{mldecls} \begin{description} - \item @{ML_type ctyp} and @{ML_type cterm} represent certified types - and terms, respectively. These are abstract datatypes that + \item Types @{ML_type ctyp} and @{ML_type cterm} represent certified + types and terms, respectively. These are abstract datatypes that guarantee that its values have passed the full well-formedness (and well-typedness) checks, relative to the declarations of type constructors, constants etc. in the theory. @@ -577,8 +668,8 @@ reasoning loops. There are separate operations to decompose certified entities (including actual theorems). - \item @{ML_type thm} represents proven propositions. This is an - abstract datatype that guarantees that its values have been + \item Type @{ML_type thm} represents proven propositions. This is + an abstract datatype that guarantees that its values have been constructed by basic principles of the @{ML_struct Thm} module. Every @{ML_type thm} value contains a sliding back-reference to the enclosing theory, cf.\ \secref{sec:context-theory}. @@ -631,6 +722,68 @@ *} +text %mlantiq {* + \begin{matharray}{rcl} + @{ML_antiquotation_def "ctyp"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "cterm"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "cprop"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "thm"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "thms"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "lemma"} & : & @{text ML_antiquotation} \\ + \end{matharray} + + \begin{rail} + 'ctyp' typ + ; + 'cterm' term + ; + 'cprop' prop + ; + 'thm' thmref + ; + 'thms' thmrefs + ; + 'lemma' ('(open)')? ((prop +) + 'and') \\ 'by' method method? + \end{rail} + + \begin{description} + + \item @{text "@{ctyp \}"} produces a certified type wrt.\ the + current background theory --- as abstract value of type @{ML_type + ctyp}. + + \item @{text "@{cterm t}"} and @{text "@{cprop \}"} produce a + certified term wrt.\ the current background theory --- as abstract + value of type @{ML_type cterm}. + + \item @{text "@{thm a}"} produces a singleton fact --- as abstract + value of type @{ML_type thm}. + + \item @{text "@{thms a}"} produces a general fact --- as abstract + value of type @{ML_type "thm list"}. + + \item @{text "@{lemma \ by meth}"} produces a fact that is proven on + the spot according to the minimal proof, which imitates a terminal + Isar proof. The result is an abstract value of type @{ML_type thm} + or @{ML_type "thm list"}, depending on the number of propositions + given here. + + The internal derivation object lacks a proper theorem name, but it + is formally closed, unless the @{text "(open)"} option is specified + (this may impact performance of applications with proof terms). + + Since ML antiquotations are always evaluated at compile-time, there + is no run-time overhead even for non-trivial proofs. Nonetheless, + the justification is syntactically limited to a single @{command + "by"} step. More complex Isar proofs should be done in regular + theory source, before compiling the corresponding ML text that uses + the result. + + \end{description} + +*} + + subsection {* Auxiliary definitions *} text {* @@ -776,9 +929,9 @@ \end{tabular} \medskip - \noindent Thus we essentially impose nesting levels on propositions - formed from @{text "\"} and @{text "\"}. At each level there is a - prefix of parameters and compound premises, concluding an atomic + Thus we essentially impose nesting levels on propositions formed + from @{text "\"} and @{text "\"}. At each level there is a prefix + of parameters and compound premises, concluding an atomic proposition. Typical examples are @{text "\"}-introduction @{text "(A \ B) \ A \ B"} or mathematical induction @{text "P 0 \ (\n. P n \ P (Suc n)) \ P n"}. Even deeper nesting occurs in well-founded diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/ML.thy --- a/doc-src/IsarImplementation/Thy/ML.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/ML.thy Mon Oct 25 08:08:08 2010 +0200 @@ -2,633 +2,1728 @@ imports Base begin -chapter {* Advanced ML programming *} +chapter {* Isabelle/ML *} + +text {* Isabelle/ML is best understood as a certain culture based on + Standard ML. Thus it is not a new programming language, but a + certain way to use SML at an advanced level within the Isabelle + environment. This covers a variety of aspects that are geared + towards an efficient and robust platform for applications of formal + logic with fully foundational proof construction --- according to + the well-known \emph{LCF principle}. There are specific library + modules and infrastructure to address the needs for such difficult + tasks. For example, the raw parallel programming model of Poly/ML + is presented as considerably more abstract concept of \emph{future + values}, which is then used to augment the inference kernel, proof + interpreter, and theory loader accordingly. + + The main aspects of Isabelle/ML are introduced below. These + first-hand explanations should help to understand how proper + Isabelle/ML is to be read and written, and to get access to the + wealth of experience that is expressed in the source text and its + history of changes.\footnote{See + \url{http://isabelle.in.tum.de/repos/isabelle} for the full + Mercurial history. There are symbolic tags to refer to official + Isabelle releases, as opposed to arbitrary \emph{tip} versions that + merely reflect snapshots that are never really up-to-date.} *} + + +section {* Style and orthography *} + +text {* The sources of Isabelle/Isar are optimized for + \emph{readability} and \emph{maintainability}. The main purpose is + to tell an informed reader what is really going on and how things + really work. This is a non-trivial aim, but it is supported by a + certain style of writing Isabelle/ML that has emerged from long + years of system development.\footnote{See also the interesting style + guide for OCaml + \url{http://caml.inria.fr/resources/doc/guides/guidelines.en.html} + which shares many of our means and ends.} + + The main principle behind any coding style is \emph{consistency}. + For a single author of a small program this merely means ``choose + your style and stick to it''. A complex project like Isabelle, with + long years of development and different contributors, requires more + standardization. A coding style that is changed every few years or + with every new contributor is no style at all, because consistency + is quickly lost. Global consistency is hard to achieve, though. + One should always strife at least for local consistency of modules + and sub-systems, without deviating from some general principles how + to write Isabelle/ML. + + In a sense, a common coding style is like an \emph{orthography} for + the sources: it helps to read quickly over the text and see through + the main points, without getting distracted by accidental + presentation of free-style source. +*} + + +subsection {* Header and sectioning *} + +text {* Isabelle source files have a certain standardized header + format (with precise spacing) that follows ancient traditions + reaching back to the earliest versions of the system by Larry + Paulson. E.g.\ see @{"file" "~~/src/Pure/thm.ML"}. + + The header includes at least @{verbatim Title} and @{verbatim + Author} entries, followed by a prose description of the purpose of + the module. The latter can range from a single line to several + paragraphs of explanations. + + The rest of the file is divided into sections, subsections, + subsubsections, paragraphs etc.\ using some a layout via ML comments + as follows. + +\begin{verbatim} +(*** section ***) + +(** subsection **) + +(* subsubsection *) + +(*short paragraph*) + +(* + long paragraph + more text +*) +\end{verbatim} + + As in regular typography, there is some extra space \emph{before} + section headings that are adjacent to plain text (not other headings + as in the example above). + + \medskip The precise wording of the prose text given in these + headings is chosen carefully to indicate the main theme of the + subsequent formal ML text. +*} + + +subsection {* Naming conventions *} + +text {* Since ML is the primary medium to express the meaning of the + source text, naming of ML entities requires special care. + + \paragraph{Notation.} A name consists of 1--3 \emph{words} (rarely + 4, but not more) that are separated by underscore. There are three + variants concerning upper or lower case letters, which are just for + certain ML categories as follows: + + \medskip + \begin{tabular}{lll} + variant & example & ML categories \\\hline + lower-case & @{verbatim foo_bar} & values, types, record fields \\ + capitalized & @{verbatim Foo_Bar} & datatype constructors, \\ + & & structures, functors \\ + upper-case & @{verbatim FOO_BAR} & special values (combinators), \\ + & & exception constructors, signatures \\ + \end{tabular} + \medskip + + For historical reasons, many capitalized names omit underscores, + e.g.\ old-style @{verbatim FooBar} instead of @{verbatim Foo_Bar}. + Genuine mixed-case names are \emph{not} used --- clear division of + words is essential for readability.\footnote{Camel-case was invented + to workaround the lack of underscore in some early non-ASCII + character sets and keywords. Later it became a habitual in some + language communities that are now strong in numbers.} + + A single (capital) character does not count as ``word'' in this + respect: some Isabelle/ML are suffixed by extra markers like this: + @{verbatim foo_barT}. + + Name variants are produced by adding 1--3 primes, e.g.\ @{verbatim + foo'}, @{verbatim foo''}, or @{verbatim foo'''}, but not @{verbatim + foo''''} or more. Decimal digits scale better to larger numbers, + e.g.\ @{verbatim foo0}, @{verbatim foo1}, @{verbatim foo42}. + + \paragraph{Scopes.} Apart from very basic library modules, ML + structures are not ``opened'', but names are referenced with + explicit qualification, as in @{ML Syntax.string_of_term} for + example. When devising names for structures and their components it + is important aim at eye-catching compositions of both parts, because + this is how they are always seen in the sources and documentation. + For the same reasons, aliases of well-known library functions should + be avoided. -section {* Style *} + Local names of function abstraction or case/lets bindings are + typically shorter, sometimes using only rudiments of ``words'', + while still avoiding cryptic shorthands. An auxiliary function + called @{verbatim helper}, @{verbatim aux}, or @{verbatim f} is + considered bad style. + + Example: + + \begin{verbatim} + (* RIGHT *) + + fun print_foo ctxt foo = + let + fun print t = ... Syntax.string_of_term ctxt t ... + in ... end; + + + (* RIGHT *) + + fun print_foo ctxt foo = + let + val string_of_term = Syntax.string_of_term ctxt; + fun print t = ... string_of_term t ... + in ... end; + + + (* WRONG *) + + val string_of_term = Syntax.string_of_term; + + fun print_foo ctxt foo = + let + fun aux t = ... string_of_term ctxt t ... + in ... end; + + \end{verbatim} + + + \paragraph{Specific conventions.} Here are some important specific + name forms that occur frequently in the sources. + + \begin{itemize} + + \item A function that maps @{verbatim foo} to @{verbatim bar} is + called @{verbatim foo_to_bar} or @{verbatim bar_of_foo} (never + @{verbatim foo2bar}, @{verbatim bar_from_foo}, @{verbatim + bar_for_foo}, or @{verbatim bar4foo}). + + \item The name component @{verbatim legacy} means that the operation + is about to be discontinued soon. + + \item The name component @{verbatim old} means that this is historic + material that might disappear at some later stage. + + \item The name component @{verbatim global} means that this works + with the background theory instead of the regular local context + (\secref{sec:context}), sometimes for historical reasons, sometimes + due a genuine lack of locality of the concept involved, sometimes as + a fall-back for the lack of a proper context in the application + code. Whenever there is a non-global variant available, the + application should be migrated to use it with a proper local + context. + + \item Variables of the main context types of the Isabelle/Isar + framework (\secref{sec:context} and \chref{ch:local-theory}) have + firm naming conventions to allow to visualize the their data flow + via plain regular expressions in the editor. In particular: + + \begin{itemize} + + \item theories are called @{verbatim thy}, rarely @{verbatim theory} + (never @{verbatim thry}) + + \item proof contexts are called @{verbatim ctxt}, rarely @{verbatim + context} (never @{verbatim ctx}) + + \item generic contexts are called @{verbatim context}, rarely + @{verbatim ctxt} + + \item local theories are called @{verbatim lthy}, unless there is an + implicit conversion to a general proof context (semantic super-type) + + \end{itemize} + + Variations with primed or decimal numbers are always possible, as + well as ``sematic prefixes'' like @{verbatim foo_thy} or @{verbatim + bar_ctxt}, but the base conventions above need to be preserved. + + \item The main logical entities (\secref{ch:logic}) also have + established naming convention: + + \begin{itemize} + + \item sorts are called @{verbatim S} + + \item types are called @{verbatim T}, @{verbatim U}, or @{verbatim + ty} (never @{verbatim t}) + + \item terms are called @{verbatim t}, @{verbatim u}, or @{verbatim + tm} (never @{verbatim trm}) + + \item certified types are called @{verbatim cT}, rarely @{verbatim + T}, with variants as for types + + \item certified terms are called @{verbatim ct}, rarely @{verbatim + t}, with variants as for terms + + \item theorems are called @{verbatim th}, or @{verbatim thm} + + \end{itemize} + + Proper semantic names override these conventions completely. For + example, the left-hand side of an equation (as a term) can be called + @{verbatim lhs} (not @{verbatim lhs_tm}). Or a term that is treated + specifically as a variable can be called @{verbatim v} or @{verbatim + x}. + + \end{itemize} +*} + + +subsection {* General source layout *} + +text {* The general Isabelle/ML source layout imitates regular + type-setting to some extent, augmented by the requirements for + deeply nested expressions that are commonplace in functional + programming. + + \paragraph{Line length} is 80 characters according to ancient + standards, but we allow as much as 100 characters, not + more.\footnote{Readability requires to keep the beginning of a line + in view while watching its end. Modern wide-screen displays do not + change the way how the human brain works. As a rule of thumb, + sources need to be printable on plain paper.} The extra 20 + characters acknowledge the space requirements due to qualified + library references in Isabelle/ML. + + \paragraph{White-space} is used to emphasize the structure of + expressions, following mostly standard conventions for mathematical + typesetting, as can be seen in plain {\TeX} or {\LaTeX}. This + defines positioning of spaces for parentheses, punctuation, and + infixes as illustrated here: + + \begin{verbatim} + val x = y + z * (a + b); + val pair = (a, b); + val record = {foo = 1, bar = 2}; + \end{verbatim} + + Lines are normally broken \emph{after} an infix operator or + punctuation character. For example: -text {* - Like any style guide, also this one should not be interpreted dogmatically, but - with care and discernment. It consists of a collection of - recommendations which have been turned out useful over long years of - Isabelle system development and are supposed to support writing of readable - and managable code. Special cases encourage derivations, - as far as you know what you are doing. - \footnote{This style guide is loosely based on - \url{http://caml.inria.fr/resources/doc/guides/guidelines.en.html}} + \begin{verbatim} + val x = + a + + b + + c; + + val tuple = + (a, + b, + c); + \end{verbatim} + + Some special infixes (e.g.\ @{verbatim "|>"}) work better at the + start of the line, but punctuation is always at the end. + + Function application follows the tradition of @{text "\"}-calculus, + not informal mathematics. For example: @{verbatim "f a b"} for a + curried function, or @{verbatim "g (a, b)"} for a tupled function. + Note that the space between @{verbatim g} and the pair @{verbatim + "(a, b)"} follows the important principle of + \emph{compositionality}: the layout of @{verbatim "g p"} does not + change when @{verbatim "p"} is refined to a concrete pair. + + \paragraph{Indentation} uses plain spaces, never hard + tabulators.\footnote{Tabulators were invented to move the carriage + of a type-writer to certain predefined positions. In software they + could be used as a primitive run-length compression of consecutive + spaces, but the precise result would depend on non-standardized + editor configuration.} + + Each level of nesting is indented by 2 spaces, sometimes 1, very + rarely 4, never 8. + + Indentation follows a simple logical format that only depends on the + nesting depth, not the accidental length of the text that initiates + a level of nesting. Example: + + \begin{verbatim} + (* RIGHT *) + + if b then + expr1_part1 + expr1_part2 + else + expr2_part1 + expr2_part2 + + + (* WRONG *) + + if b then expr1_part1 + expr1_part2 + else expr2_part1 + expr2_part2 + \end{verbatim} + + The second form has many problems: it assumes a fixed-width font + when viewing the sources, it uses more space on the line and thus + makes it hard to observe its strict length limit (working against + \emph{readability}), it requires extra editing to adapt the layout + to changes of the initial text (working against + \emph{maintainability}) etc. + + \medskip For similar reasons, any kind of two-dimensional or tabular + layouts, ASCII-art with lines or boxes of stars etc. should be + avoided. + + \paragraph{Complex expressions} consisting of multi-clausal function + definitions, @{verbatim case}, @{verbatim handle}, @{verbatim let}, + or combinations require special attention. The syntax of Standard + ML is a bit too ambitious in admitting too much variance that can + distort the meaning of the text. + + Clauses of @{verbatim fun}, @{verbatim fn}, @{verbatim case}, + @{verbatim handle} get extra indentation to indicate the nesting + clearly. For example: + + \begin{verbatim} + (* RIGHT *) + + fun foo p1 = + expr1 + | foo p2 = + expr2 + + + (* WRONG *) + + fun foo p1 = + expr1 + | foo p2 = + expr2 + \end{verbatim} + + Body expressions consisting of @{verbatim case} or @{verbatim let} + require care to maintain compositionality, to prevent loss of + logical indentation where it is particularly imporant to see the + structure of the text later on. Example: + + \begin{verbatim} + (* RIGHT *) + + fun foo p1 = + (case e of + q1 => ... + | q2 => ...) + | foo p2 = + let + ... + in + ... + end + + + (* WRONG *) + + fun foo p1 = case e of + q1 => ... + | q2 => ... + | foo p2 = + let + ... + in + ... + end + \end{verbatim} + + Extra parentheses around @{verbatim case} expressions are optional, + but help to analyse the nesting easily based on simple string + matching in the editor. + + \medskip There are two main exceptions to the overall principle of + compositionality in the layout of complex expressions. + + \begin{enumerate} + + \item @{verbatim "if"} expressions are iterated as if there would be + a multi-branch conditional in SML, e.g. + + \begin{verbatim} + (* RIGHT *) + + if b1 then e1 + else if b2 then e2 + else e3 + \end{verbatim} + + \item @{verbatim fn} abstractions are often layed-out as if they + would lack any structure by themselves. This traditional form is + motivated by the possibility to shift function arguments back and + forth wrt.\ additional combinators. For example: + + \begin{verbatim} + (* RIGHT *) + + fun foo x y = fold (fn z => + expr) + \end{verbatim} + + Here the visual appearance is that of three arguments @{verbatim x}, + @{verbatim y}, @{verbatim z}. + + \end{enumerate} + + Such weakly structured layout should be use with great care. Here + is a counter-example involving @{verbatim let} expressions: + + \begin{verbatim} + (* WRONG *) + + fun foo x = let + val y = ... + in ... end + + fun foo x = + let + val y = ... + in ... end + \end{verbatim} + + \medskip In general the source layout is meant to emphasize the + structure of complex language expressions, not to pretend that SML + had a completely different syntax (say that of Haskell or Java). +*} + + +section {* SML embedded into Isabelle/Isar *} + +text {* ML and Isar are intertwined via an open-ended bootstrap + process that provides more and more programming facilities and + logical content in an alternating manner. Bootstrapping starts from + the raw environment of existing implementations of Standard ML + (mainly Poly/ML, but also SML/NJ). + + Isabelle/Pure marks the point where the original ML toplevel is + superseded by the Isar toplevel that maintains a uniform environment + for arbitrary ML values (see also \secref{sec:context}). This + formal context holds logical entities as well as ML compiler + bindings, among many other things. Raw Standard ML is never + encountered again after the initial bootstrap of Isabelle/Pure. + + Object-logics such as Isabelle/HOL are built within the + Isabelle/ML/Isar environment of Pure by introducing suitable + theories with associated ML text, either inlined or as separate + files. Thus Isabelle/HOL is defined as a regular user-space + application within the Isabelle framework. Further add-on tools can + be implemented in ML within the Isar context in the same manner: ML + is part of the regular user-space repertoire of Isabelle. +*} + + +subsection {* Isar ML commands *} + +text {* The primary Isar source language provides various facilities + to open a ``window'' to the underlying ML compiler. Especially see + @{command_ref "use"} and @{command_ref "ML"}, which work exactly the + same way, only the source text is provided via a file vs.\ inlined, + respectively. Apart from embedding ML into the main theory + definition like that, there are many more commands that refer to ML + source, such as @{command_ref setup} or @{command_ref declaration}. + An example of even more fine-grained embedding of ML into Isar is + the proof method @{method_ref tactic}, which refines the pending goal state + via a given expression of type @{ML_type tactic}. +*} + +text %mlex {* The following artificial example demonstrates some ML + toplevel declarations within the implicit Isar theory context. This + is regular functional programming without referring to logical + entities yet. +*} + +ML {* + fun factorial 0 = 1 + | factorial n = n * factorial (n - 1) +*} + +text {* Here the ML environment is really managed by Isabelle, i.e.\ + the @{ML factorial} function is not yet accessible in the preceding + paragraph, nor in a different theory that is independent from the + current one in the import hierarchy. + + Removing the above ML declaration from the source text will remove + any trace of this definition as expected. The Isabelle/ML toplevel + environment is managed in a \emph{stateless} way: unlike the raw ML + toplevel or similar command loops of Computer Algebra systems, there + are no global side-effects involved here.\footnote{Such a stateless + compilation environment is also a prerequisite for robust parallel + compilation within independent nodes of the implicit theory + development graph.} + + \medskip The next example shows how to embed ML into Isar proofs. + Instead of @{command_ref "ML"} for theory mode, we use @{command_ref + "ML_prf"} for proof mode. As illustrated below, its effect on the + ML environment is local to the whole proof body, while ignoring its + internal block structure. *} + +example_proof + ML_prf %"ML" {* val a = 1 *} + { -- {* Isar block structure ignored by ML environment *} + ML_prf %"ML" {* val b = a + 1 *} + } -- {* Isar block structure ignored by ML environment *} + ML_prf %"ML" {* val c = b + 1 *} +qed + +text {* By side-stepping the normal scoping rules for Isar proof + blocks, embedded ML code can refer to the different contexts + explicitly, and manipulate corresponding entities, e.g.\ export a + fact from a block context. + + \medskip Two further ML commands are useful in certain situations: + @{command_ref ML_val} and @{command_ref ML_command} are both + \emph{diagnostic} in the sense that there is no effect on the + underlying environment, and can thus used anywhere (even outside a + theory). The examples below produce long strings of digits by + invoking @{ML factorial}: @{command ML_val} already takes care of + printing the ML toplevel result, but @{command ML_command} is silent + so we produce an explicit output message. *} + +ML_val {* factorial 100 *} +ML_command {* writeln (string_of_int (factorial 100)) *} + +example_proof + ML_val {* factorial 100 *} (* FIXME check/fix indentation *) + ML_command {* writeln (string_of_int (factorial 100)) *} +qed + + +subsection {* Compile-time context *} + +text {* Whenever the ML compiler is invoked within Isabelle/Isar, the + formal context is passed as a thread-local reference variable. Thus + ML code may access the theory context during compilation, by reading + or writing the (local) theory under construction. Note that such + direct access to the compile-time context is rare; in practice it is + typically via some derived ML functions. +*} + +text %mlref {* + \begin{mldecls} + @{index_ML ML_Context.the_generic_context: "unit -> Context.generic"} \\ + @{index_ML "Context.>> ": "(Context.generic -> Context.generic) -> unit"} \\ + @{index_ML bind_thms: "string * thm list -> unit"} \\ + @{index_ML bind_thm: "string * thm -> unit"} \\ + \end{mldecls} \begin{description} - \item[fundamental law of programming] - Whenever writing code, keep in mind: A program is - written once, modified ten times, and read - hundred times. So simplify its writing, - always keep future modifications in mind, - and never jeopardize readability. Every second you hesitate - to spend on making your code more clear you will - have to spend ten times understanding what you have - written later on. + \item @{ML "ML_Context.the_generic_context ()"} refers to the theory + context of the ML toplevel --- at compile time. ML code needs to + take care to refer to @{ML "ML_Context.the_generic_context ()"} + correctly. Recall that evaluation of a function body is delayed + until actual run-time. + + \item @{ML "Context.>>"}~@{text f} applies context transformation + @{text f} to the implicit context of the ML toplevel. + + \item @{ML bind_thms}~@{text "(name, thms)"} stores a list of + theorems produced in ML both in the (global) theory context and the + ML toplevel, associating it with the provided name. Theorems are + put into a global ``standard'' format before being stored. + + \item @{ML bind_thm} is similar to @{ML bind_thms} but refers to a + singleton theorem. + + \end{description} + + It is very important to note that the above functions are really + restricted to the compile time, even though the ML compiler is + invoked at run-time. The majority of ML code either uses static + antiquotations (\secref{sec:ML-antiq}) or refers to the theory or + proof context at run-time, by explicit functional abstraction. +*} + + +subsection {* Antiquotations \label{sec:ML-antiq} *} + +text {* A very important consequence of embedding SML into Isar is the + concept of \emph{ML antiquotation}. First, the standard token + language of ML is augmented by special syntactic entities of the + following form: + + \indexouternonterm{antiquote} + \begin{rail} + antiquote: atsign lbrace nameref args rbrace | lbracesym | rbracesym + ; + \end{rail} + + Here @{syntax nameref} and @{syntax args} are regular outer syntax + categories. Note that attributes and proof methods use similar + syntax. + + \medskip A regular antiquotation @{text "@{name args}"} processes + its arguments by the usual means of the Isar source language, and + produces corresponding ML source text, either as literal + \emph{inline} text (e.g. @{text "@{term t}"}) or abstract + \emph{value} (e.g. @{text "@{thm th}"}). This pre-compilation + scheme allows to refer to formal entities in a robust manner, with + proper static scoping and with some degree of logical checking of + small portions of the code. + + Special antiquotations like @{text "@{let \}"} or @{text "@{note + \}"} augment the compilation context without generating code. The + non-ASCII braces @{text "\"} and @{text "\"} allow to delimit the + effect by introducing local blocks within the pre-compilation + environment. + + \medskip See also \cite{Wenzel-Chaieb:2007b} for a slightly broader + perspective on Isabelle/ML antiquotations. *} + +text %mlantiq {* + \begin{matharray}{rcl} + @{ML_antiquotation_def "let"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "note"} & : & @{text ML_antiquotation} \\ + \end{matharray} - \item[white space matters] - Treat white space in your code as if it determines - the meaning of code. + \begin{rail} + 'let' ((term + 'and') '=' term + 'and') + ; + + 'note' (thmdef? thmrefs + 'and') + ; + \end{rail} + + \begin{description} + + \item @{text "@{let p = t}"} binds schematic variables in the + pattern @{text "p"} by higher-order matching against the term @{text + "t"}. This is analogous to the regular @{command_ref let} command + in the Isar proof language. The pre-compilation environment is + augmented by auxiliary term bindings, without emitting ML source. + + \item @{text "@{note a = b\<^sub>1 \ b\<^sub>n}"} recalls existing facts @{text + "b\<^sub>1, \, b\<^sub>n"}, binding the result as @{text a}. This is analogous to + the regular @{command_ref note} command in the Isar proof language. + The pre-compilation environment is augmented by auxiliary fact + bindings, without emitting ML source. + + \end{description} +*} + +text %mlex {* The following artificial example gives a first + impression about using the antiquotation elements introduced so far, + together with the basic @{text "@{thm}"} antiquotation defined + later. +*} - \begin{itemize} +ML {* + \ + @{let ?t = my_term} + @{note my_refl = reflexive [of ?t]} + fun foo th = Thm.transitive th @{thm my_refl} + \ +*} + +text {* + The extra block delimiters do not affect the compiled code itself, i.e.\ + function @{ML foo} is available in the present context. +*} + + +section {* Canonical argument order \label{sec:canonical-argument-order} *} + +text {* Standard ML is a language in the tradition of @{text + "\"}-calculus and \emph{higher-order functional programming}, + similar to OCaml, Haskell, or Isabelle/Pure and HOL as logical + languages. Getting acquainted with the native style of representing + functions in that setting can save a lot of extra boiler-plate of + redundant shuffling of arguments, auxiliary abstractions etc. + + First of all, functions are usually \emph{curried}: the idea of + @{text "f"} taking @{text "n"} arguments of type @{text "\\<^sub>i"} (for + @{text "i \ {1, \ n}"}) with result @{text "\"} is represented by + the iterated function space @{text "\\<^sub>1 \ \ \ \\<^sub>n \ \"}. This is + isomorphic to the encoding via tuples @{text "\\<^sub>1 \ \ \ \\<^sub>n \ \"}, but + the curried version fits more smoothly into the basic + calculus.\footnote{The difference is even more significant in + higher-order logic, because additional the redundant tuple structure + needs to be accommodated by formal reasoning.} + + Currying gives some flexiblity due to \emph{partial application}. A + function @{text "f: \\<^sub>1 \ \\<^bsub>2\<^esub> \ \"} can be applied to @{text "x: \\<^sub>1"} + and the remaining @{text "(f x): \\<^sub>2 \ \"} passed to other functions + etc. How well this works in practice depends on the order of + arguments. In the worst case, arguments are arranged erratically, + and using a function in a certain situation always requires some + glue code. Thus we would get exponentially many oppurtunities to + decorate the code with meaningless permutations of arguments. - \item The space bar is the easiest key to find on the keyboard, - press it as often as necessary. @{verbatim "2 + 2"} is better - than @{verbatim "2+2"}, likewise @{verbatim "f (x, y)"} is - better than @{verbatim "f(x,y)"}. + This can be avoided by \emph{canonical argument order}, which + observes certain standard patterns of function definitions and + minimizes adhoc permutations in their application. In Isabelle/ML, + large portions of non-trivial source code are written without ever + using the infamous function @{text "swap: \ \ \ \ \ \ \"} or the + combinator @{text "C: (\ \ \ \ \) \ (\ \ \ \ \)"}, which is not even + defined in our library. + + \medskip The basic idea is that arguments that vary less are moved + further to the left than those that vary more. Two particularly + important categories of functions are \emph{selectors} and + \emph{updates}. + + The subsequent scheme is based on a hypothetical set-like container + of type @{text "\"} that manages elements of type @{text "\"}. Both + the names and types of the associated operations are canonical for + Isabelle/ML. + + \medskip + \begin{tabular}{ll} + kind & canonical name and type \\\hline + selector & @{text "member: \ \ \ \ bool"} \\ + update & @{text "insert: \ \ \ \ \"} \\ + \end{tabular} + \medskip + + Given a container @{text "B: \"}, the partially applied @{text + "member B"} is a predicate over elements @{text "\ \ bool"}, and + thus represents the intended denotation directly. It is customary + to pass the abstract predicate to further operations, not the + concrete container. The argument order makes it easy to use other + combinators: @{text "forall (member B) list"} will check a list of + elements for membership in @{text "B"} etc. Often the explicit + @{text "list"} is pointless and can be contracted in the application + @{text "forall (member B)"} to get directly a predicate again. + + The update operation naturally ``varies'' the container, so it moves + to the right: @{text "insert a"} is a function @{text "\ \ \"} to + insert a value @{text "a"}. These can be composed naturally as + follows: @{text "insert c \ insert b \ insert a"}. This works well, + apart from some awkwardness due to conventional mathematical + function composition, which can be easily amended as follows. +*} + + +subsection {* Forward application and composition *} + +text {* Regular function application and infix notation works best for + relatively deeply structured expressions, e.g.\ @{text "h (f x y + g + z)"}. The important special case if \emph{linear transformation} + applies a cascade of functions as follows: @{text "f\<^sub>n (\ (f\<^sub>1 x))"}. + This becomes hard to read and maintain if the functions are + themselves complex expressions. The notation can be significantly + improved by introducing \emph{forward} versions of application and + composition as follows: + + \medskip + \begin{tabular}{lll} + @{text "x |> f"} & @{text "\"} & @{text "f x"} \\ + @{text "f #> g"} & @{text "\"} & @{text "x |> f |> g"} \\ + \end{tabular} + \medskip + + This enables to write conveniently @{text "x |> f\<^sub>1 |> \ |> f\<^sub>n"} or + @{text "f\<^sub>1 #> \ #> f\<^sub>n"} for its functional abstraction over @{text + "x"}. + + \medskip There is an additional set of combinators to accommodate + multiple results (via pairs) that are passed on as multiple + arguments (via currying). - \item Restrict your lines to 80 characters. This will allow - you to keep the beginning of a line in view while watching - its end.\footnote{To acknowledge the lax practice of - text editing these days, we tolerate as much as 100 - characters per line, but anything beyond 120 is not - considered proper source text.} + \medskip + \begin{tabular}{lll} + @{text "(x, y) |-> f"} & @{text "\"} & @{text "f x y"} \\ + @{text "f #-> g"} & @{text "\"} & @{text "x |> f |-> g"} \\ + \end{tabular} + \medskip +*} + +text %mlref {* + \begin{mldecls} + @{index_ML "op |> ": "'a * ('a -> 'b) -> 'b"} \\ + @{index_ML "op |-> ": "('c * 'a) * ('c -> 'a -> 'b) -> 'b"} \\ + @{index_ML "op #> ": "('a -> 'b) * ('b -> 'c) -> 'a -> 'c"} \\ + @{index_ML "op #-> ": "('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd"} \\ + \end{mldecls} + + %FIXME description!? +*} + + +subsection {* Canonical iteration *} + +text {* As explained above, a function @{text "f: \ \ \ \ \"} can be + understood as update on a configuration of type @{text "\"} that is + parametrized by arguments of type @{text "\"}. Given @{text "a: \"} + the partial application @{text "(f a): \ \ \"} operates + homogeneously on @{text "\"}. This can be iterated naturally over a + list of parameters @{text "[a\<^sub>1, \, a\<^sub>n]"} as @{text "f a\<^sub>1 #> \ #> f + a\<^bsub>n\<^esub>\<^bsub>\<^esub>"}. The latter expression is again a function @{text "\ \ \"}. + It can be applied to an initial configuration @{text "b: \"} to + start the iteration over the given list of arguments: each @{text + "a"} in @{text "a\<^sub>1, \, a\<^sub>n"} is applied consecutively by updating a + cumulative configuration. + + The @{text fold} combinator in Isabelle/ML lifts a function @{text + "f"} as above to its iterated version over a list of arguments. + Lifting can be repeated, e.g.\ @{text "(fold \ fold) f"} iterates + over a list of lists as expected. + + The variant @{text "fold_rev"} works inside-out over the list of + arguments, such that @{text "fold_rev f \ fold f \ rev"} holds. - \item Ban tabulators; they are a context-sensitive formatting - feature and likely to confuse anyone not using your favorite - editor.\footnote{Some modern programming language even - forbid tabulators altogether according to the formal syntax - definition.} + The @{text "fold_map"} combinator essentially performs @{text + "fold"} and @{text "map"} simultaneously: each application of @{text + "f"} produces an updated configuration together with a side-result; + the iteration collects all such side-results as a separate list. +*} + +text %mlref {* + \begin{mldecls} + @{index_ML fold: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\ + @{index_ML fold_rev: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\ + @{index_ML fold_map: "('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML fold}~@{text f} lifts the parametrized update function + @{text "f"} to a list of parameters. + + \item @{ML fold_rev}~@{text "f"} is similar to @{ML fold}~@{text + "f"}, but works inside-out. - \item Get rid of trailing whitespace. Instead, do not - suppress a trailing newline at the end of your files. + \item @{ML fold_map}~@{text "f"} lifts the parametrized update + function @{text "f"} (with side-result) to a list of parameters and + cumulative side-results. + + \end{description} + + \begin{warn} + The literature on functional programming provides a multitude of + combinators called @{text "foldl"}, @{text "foldr"} etc. SML97 + provides its own variations as @{ML List.foldl} and @{ML + List.foldr}, while the classic Isabelle library still has the + slightly more convenient historic @{ML Library.foldl} and @{ML + Library.foldr}. To avoid further confusion, all of this should be + ignored, and @{ML fold} (or @{ML fold_rev}) used exclusively. + \end{warn} +*} + +text %mlex {* The following example shows how to fill a text buffer + incrementally by adding strings, either individually or from a given + list. +*} - \item Choose a generally accepted style of indentation, - then use it systematically throughout the whole - application. An indentation of two spaces is appropriate. - Avoid dangling indentation. +ML {* + val s = + Buffer.empty + |> Buffer.add "digits: " + |> fold (Buffer.add o string_of_int) (0 upto 9) + |> Buffer.content; + + @{assert} (s = "digits: 0123456789"); +*} + +text {* Note how @{ML "fold (Buffer.add o string_of_int)"} above saves + an extra @{ML "map"} over the given list. This kind of peephole + optimization reduces both the code size and the tree structures in + memory (``deforestation''), but requires some practice to read and + write it fluently. + + \medskip The next example elaborates this idea and demonstrates fast + accumulation of tree content using a text buffer. +*} + +ML {* + datatype tree = Text of string | Elem of string * tree list; + + fun slow_content (Text txt) = txt + | slow_content (Elem (name, ts)) = + "<" ^ name ^ ">" ^ + implode (map slow_content ts) ^ + "" + + fun add_content (Text txt) = Buffer.add txt + | add_content (Elem (name, ts)) = + Buffer.add ("<" ^ name ^ ">") #> + fold add_content ts #> + Buffer.add (""); - \end{itemize} + fun fast_content tree = + Buffer.empty |> add_content tree |> Buffer.content; +*} + +text {* The slow part of @{ML slow_content} is the @{ML implode} of + the recursive results, because it copies previously produced strings + afresh. + + The incremental @{ML add_content} avoids this by operating on a + buffer that is passed-through in a linear fashion. Using @{verbatim + "#>"} and contraction over the actual @{ML_type "Buffer.T"} argument + saves some additional boiler-plate, but requires again some + practice. Of course, the two @{ML "Buffer.add"} invocations with + concatenated strings could have been split into smaller parts, but + this would have obfuscated the source without making a big + difference in allocations. Here we have done some + peephole-optimization for the sake of readability. + + Another benefit of @{ML add_content} is its ``open'' form as a + function on @{ML_type Buffer.T} that can be continued in further + linear transformations, folding etc. Thus it is more compositional + than the naive @{ML slow_content}. As a realistic example, compare + the slightly old-style @{ML "Term.maxidx_of_term: term -> int"} with + the newer @{ML "Term.maxidx_term: term -> int -> int"} in + Isabelle/Pure. + + Note that @{ML fast_content} above is only defined for completeness + of the example. In many practical situations, it is customary to + defined the incremental @{ML add_content} only and leave the + initialization and termination to the concrete application by the + user. +*} + + +section {* Message output channels \label{sec:message-channels} *} + +text {* Isabelle provides output channels for different kinds of + messages: regular output, high-volume tracing information, warnings, + and errors. - \item[cut-and-paste succeeds over copy-and-paste] - \emph{Never} copy-and-paste code when programming. If you - need the same piece of code twice, introduce a - reasonable auxiliary function (if there is no - such function, very likely you got something wrong). - Any copy-and-paste will turn out to be painful - when something has to be changed later on. + Depending on the user interface involved, these messages may appear + in different text styles or colours. The standard output for + terminal sessions prefixes each line of warnings by @{verbatim + "###"} and errors by @{verbatim "***"}, but leaves anything else + unchanged. + + Messages are associated with the transaction context of the running + Isar command. This enables the front-end to manage commands and + resulting messages together. For example, after deleting a command + from a given theory document version, the corresponding message + output can be retracted from the display. +*} + +text %mlref {* + \begin{mldecls} + @{index_ML writeln: "string -> unit"} \\ + @{index_ML tracing: "string -> unit"} \\ + @{index_ML warning: "string -> unit"} \\ + @{index_ML error: "string -> 'a"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML writeln}~@{text "text"} outputs @{text "text"} as regular + message. This is the primary message output operation of Isabelle + and should be used by default. + + \item @{ML tracing}~@{text "text"} outputs @{text "text"} as special + tracing message, indicating potential high-volume output to the + front-end (hundreds or thousands of messages issued by a single + command). The idea is to allow the user-interface to downgrade the + quality of message display to achieve higher throughput. + + Note that the user might have to take special actions to see tracing + output, e.g.\ switch to a different output window. So this channel + should not be used for regular output. - \item[comments] - are a device which requires careful thinking before using - it. The best comment for your code should be the code itself. - Prefer efforts to write clear, understandable code - over efforts to explain nasty code. + \item @{ML warning}~@{text "text"} outputs @{text "text"} as + warning, which typically means some extra emphasis on the front-end + side (color highlighting, icon). + + \item @{ML error}~@{text "text"} raises exception @{ML ERROR}~@{text + "text"} and thus lets the Isar toplevel print @{text "text"} on the + error channel, which typically means some extra emphasis on the + front-end side (color highlighting, icon). + + This assumes that the exception is not handled before the command + terminates. Handling exception @{ML ERROR}~@{text "text"} is a + perfectly legal alternative: it means that the error is absorbed + without any message output. + + \begin{warn} + The actual error channel is accessed via @{ML Output.error_msg}, but + the interaction protocol of Proof~General \emph{crashes} if that + function is used in regular ML code: error output and toplevel + command failure always need to coincide here. + \end{warn} + + \end{description} + + \begin{warn} + Regular Isabelle/ML code should output messages exclusively by the + official channels. Using raw I/O on \emph{stdout} or \emph{stderr} + instead (e.g.\ via @{ML TextIO.output}) is apt to cause problems in + the presence of parallel and asynchronous processing of Isabelle + theories. Such raw output might be displayed by the front-end in + some system console log, with a low chance that the user will ever + see it. Moreover, as a genuine side-effect on global process + channels, there is no proper way to retract output when Isar command + transactions are reset. + \end{warn} + + \begin{warn} + The message channels should be used in a message-oriented manner. + This means that multi-line output that logically belongs together + needs to be issued by a \emph{single} invocation of @{ML writeln} + etc. with the functional concatenation of all message constituents. + \end{warn} +*} - \item[functional programming is based on functions] - Model things as abstract as appropriate. Avoid unnecessarrily - concrete datatype representations. For example, consider a function - taking some table data structure as argument and performing - lookups on it. Instead of taking a table, it could likewise - take just a lookup function. Accustom - your way of coding to the level of expressiveness a functional - programming language is giving onto you. +text %mlex {* The following example demonstrates a multi-line + warning. Note that in some situations the user sees only the first + line, so the most important point should be made first. +*} + +ML_command {* + warning (cat_lines + ["Beware the Jabberwock, my son!", + "The jaws that bite, the claws that catch!", + "Beware the Jubjub Bird, and shun", + "The frumious Bandersnatch!"]); +*} + + +section {* Exceptions \label{sec:exceptions} *} + +text {* The Standard ML semantics of strict functional evaluation + together with exceptions is rather well defined, but some delicate + points need to be observed to avoid that ML programs go wrong + despite static type-checking. Exceptions in Isabelle/ML are + subsequently categorized as follows. + + \paragraph{Regular user errors.} These are meant to provide + informative feedback about malformed input etc. + + The \emph{error} function raises the corresponding \emph{ERROR} + exception, with a plain text message as argument. \emph{ERROR} + exceptions can be handled internally, in order to be ignored, turned + into other exceptions, or cascaded by appending messages. If the + corresponding Isabelle/Isar command terminates with an \emph{ERROR} + exception state, the toplevel will print the result on the error + channel (see \secref{sec:message-channels}). + + It is considered bad style to refer to internal function names or + values in ML source notation in user error messages. + + Grammatical correctness of error messages can be improved by + \emph{omitting} final punctuation: messages are often concatenated + or put into a larger context (e.g.\ augmented with source position). + By not insisting in the final word at the origin of the error, the + system can perform its administrative tasks more easily and + robustly. + + \paragraph{Program failures.} There is a handful of standard + exceptions that indicate general failure situations, or failures of + core operations on logical entities (types, terms, theorems, + theories, see \chref{ch:logic}). + + These exceptions indicate a genuine breakdown of the program, so the + main purpose is to determine quickly what has happened where. + Traditionally, the (short) exception message would include the name + of an ML function, although this no longer necessary, because the ML + runtime system prints a detailed source position of the + corresponding @{verbatim raise} keyword. + + \medskip User modules can always introduce their own custom + exceptions locally, e.g.\ to organize internal failures robustly + without overlapping with existing exceptions. Exceptions that are + exposed in module signatures require extra care, though, and should + \emph{not} be introduced by default. Surprise by end-users or ML + users of a module can be often minimized by using plain user errors. + + \paragraph{Interrupts.} These indicate arbitrary system events: + both the ML runtime system and the Isabelle/ML infrastructure signal + various exceptional situations by raising the special + \emph{Interrupt} exception in user code. + + This is the one and only way that physical events can intrude an + Isabelle/ML program. Such an interrupt can mean out-of-memory, + stack overflow, timeout, internal signaling of threads, or the user + producing a console interrupt manually etc. An Isabelle/ML program + that intercepts interrupts becomes dependent on physical effects of + the environment. Even worse, exception handling patterns that are + too general by accident, e.g.\ by mispelled exception constructors, + will cover interrupts unintentionally, and thus render the program + semantics ill-defined. - \item[tuples] - are often in the way. When there is no striking argument - to tuple function arguments, just write your function curried. + Note that the Interrupt exception dates back to the original SML90 + language definition. It was excluded from the SML97 version to + avoid its malign impact on ML program semantics, but without + providing a viable alternative. Isabelle/ML recovers physical + interruptibility (which an indispensable tool to implement managed + evaluation of Isar command transactions), but requires user code to + be strictly transparent wrt.\ interrupts. + + \begin{warn} + Isabelle/ML user code needs to terminate promptly on interruption, + without guessing at its meaning to the system infrastructure. + Temporary handling of interrupts for cleanup of global resources + etc.\ needs to be followed immediately by re-raising of the original + exception. + \end{warn} +*} + +text %mlref {* + \begin{mldecls} + @{index_ML try: "('a -> 'b) -> 'a -> 'b option"} \\ + @{index_ML can: "('a -> 'b) -> 'a -> bool"} \\ + @{index_ML ERROR: "string -> exn"} \\ + @{index_ML Fail: "string -> exn"} \\ + @{index_ML Exn.is_interrupt: "exn -> bool"} \\ + @{index_ML reraise: "exn -> 'a"} \\ + @{index_ML exception_trace: "(unit -> 'a) -> 'a"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML try}~@{text "f x"} makes the partiality of evaluating + @{text "f x"} explicit via the option datatype. Interrupts are + \emph{not} handled here, i.e.\ this form serves as safe replacement + for the \emph{unsafe} version @{verbatim "(SOME"}~@{text "f + x"}~@{verbatim "handle _ => NONE)"} that is occasionally seen in + books about SML. + + \item @{ML can} is similar to @{ML try} with more abstract result. - \item[telling names] - Any name should tell its purpose as exactly as possible, while - keeping its length to the absolutely necessary minimum. Always - give the same name to function arguments which have the same - meaning. Separate words by underscores (@{verbatim - int_of_string}, not @{verbatim intOfString}).\footnote{Some - recent tools for Emacs include special precautions to cope with - bumpy names in @{text "camelCase"}, e.g.\ for improved on-screen - readability. It is easier to abstain from using such names in the - first place.} + \item @{ML ERROR}~@{text "msg"} represents user errors; this + exception is always raised via the @{ML error} function (see + \secref{sec:message-channels}). + + \item @{ML Fail}~@{text "msg"} represents general program failures. + + \item @{ML Exn.is_interrupt} identifies interrupts robustly, without + mentioning concrete exception constructors in user code. Handled + interrupts need to be re-raised promptly! + + \item @{ML reraise}~@{text "exn"} raises exception @{text "exn"} + while preserving its implicit position information (if possible, + depending on the ML platform). + + \item @{ML exception_trace}~@{verbatim "(fn _ =>"}~@{text + "e"}@{verbatim ")"} evaluates expression @{text "e"} while printing + a full trace of its stack of nested exceptions (if possible, + depending on the ML platform).\footnote{In various versions of + Poly/ML the trace will appear on raw stdout of the Isabelle + process.} + + Inserting @{ML exception_trace} into ML code temporarily is useful + for debugging, but not suitable for production code. + + \end{description} +*} + +text %mlantiq {* + \begin{matharray}{rcl} + @{ML_antiquotation_def "assert"} & : & @{text ML_antiquotation} \\ + \end{matharray} + + \begin{description} + + \item @{text "@{assert}"} inlines a function @{text "bool \ unit"} + that raises @{ML Fail} if the argument is @{ML false}. Due to + inlining the source position of failed assertions is included in the + error output. \end{description} *} -section {* Thread-safe programming *} +section {* Basic data types *} + +text {* The basis library proposal of SML97 need to be treated with + caution. Many of its operations simply do not fit with important + Isabelle/ML conventions (like ``canonical argument order'', see + \secref{sec:canonical-argument-order}), others can cause serious + problems with the parallel evaluation model of Isabelle/ML (such as + @{ML TextIO.print} or @{ML OS.Process.system}). + + Subsequently we give a brief overview of important operations on + basic ML data types. +*} + + +subsection {* Characters *} + +text %mlref {* + \begin{mldecls} + @{index_ML_type char} \\ + \end{mldecls} + + \begin{description} + + \item Type @{ML_type char} is \emph{not} used. The smallest textual + unit in Isabelle is represented a ``symbol'' (see + \secref{sec:symbols}). + + \end{description} +*} + + +subsection {* Integers *} + +text %mlref {* + \begin{mldecls} + @{index_ML_type int} \\ + \end{mldecls} + + \begin{description} + + \item Type @{ML_type int} represents regular mathematical integers, + which are \emph{unbounded}. Overflow never happens in + practice.\footnote{The size limit for integer bit patterns in memory + is 64\,MB for 32-bit Poly/ML, and much higher for 64-bit systems.} + This works uniformly for all supported ML platforms (Poly/ML and + SML/NJ). + + Literal integers in ML text (e.g.\ @{ML + 123456789012345678901234567890}) are forced to be of this one true + integer type --- overloading of SML97 is disabled. + + Structure @{ML_struct IntInf} of SML97 is obsolete, always use + @{ML_struct Int}. Structure @{ML_struct Integer} in @{"file" + "~~/src/Pure/General/integer.ML"} provides some additional + operations. + + \end{description} +*} + + +subsection {* Options *} + +text %mlref {* + \begin{mldecls} + @{index_ML Option.map: "('a -> 'b) -> 'a option -> 'b option"} \\ + @{index_ML is_some: "'a option -> bool"} \\ + @{index_ML is_none: "'a option -> bool"} \\ + @{index_ML the: "'a option -> 'a"} \\ + @{index_ML these: "'a list option -> 'a list"} \\ + @{index_ML the_list: "'a option -> 'a list"} \\ + @{index_ML the_default: "'a -> 'a option -> 'a"} \\ + \end{mldecls} +*} + +text {* Apart from @{ML Option.map} most operations defined in + structure @{ML_struct Option} are alien to Isabelle/ML. The + operations shown above are defined in @{"file" + "~~/src/Pure/General/basics.ML"}, among others. *} + + +subsection {* Lists *} + +text {* Lists are ubiquitous in ML as simple and light-weight + ``collections'' for many everyday programming tasks. Isabelle/ML + provides important additions and improvements over operations that + are predefined in the SML97 library. *} + +text %mlref {* + \begin{mldecls} + @{index_ML cons: "'a -> 'a list -> 'a list"} \\ + @{index_ML member: "('b * 'a -> bool) -> 'a list -> 'b -> bool"} \\ + @{index_ML insert: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\ + @{index_ML remove: "('b * 'a -> bool) -> 'b -> 'a list -> 'a list"} \\ + @{index_ML update: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML cons}~@{text "x xs"} evaluates to @{text "x :: xs"}. + + Tupled infix operators are a historical accident in Standard ML. + The curried @{ML cons} amends this, but it should be only used when + partial application is required. + + \item @{ML member}, @{ML insert}, @{ML remove}, @{ML update} treat + lists as a set-like container that maintains the order of elements. + See @{"file" "~~/src/Pure/library.ML"} for the full specifications + (written in ML). There are some further derived operations like + @{ML union} or @{ML inter}. + + Note that @{ML insert} is conservative about elements that are + already a @{ML member} of the list, while @{ML update} ensures that + the last entry is always put in front. The latter discipline is + often more appropriate in declarations of context data + (\secref{sec:context-data}) that are issued by the user in Isar + source: more recent declarations normally take precedence over + earlier ones. + + \end{description} +*} + +text %mlex {* Using canonical @{ML fold} together with canonical @{ML + cons}, or similar standard operations, alternates the orientation of + data. The is quite natural and should not altered forcible by + inserting extra applications @{ML rev}. The alternative @{ML + fold_rev} can be used in the relatively few situations, where + alternation should be prevented. +*} + +ML {* + val items = [1, 2, 3, 4, 5, 6, 7, 8, 9, 10]; + + val list1 = fold cons items []; + @{assert} (list1 = rev items); + + val list2 = fold_rev cons items []; + @{assert} (list2 = items); +*} + +text {* The subsequent example demonstrates how to \emph{merge} two + lists in a natural way. *} + +ML {* + fun merge_lists eq (xs, ys) = fold_rev (insert eq) ys xs; +*} + +text {* Here the first list is treated conservatively: only the new + elements from the second list are inserted. The inside-out order of + insertion via @{ML fold_rev} attempts to preserve the order of + elements in the result. + + This way of merging lists is typical for context data + (\secref{sec:context-data}). See also @{ML merge} as defined in + @{"file" "~~/src/Pure/library.ML"}. +*} + + +subsection {* Association lists *} + +text {* The operations for association lists interpret a concrete list + of pairs as a finite function from keys to values. Redundant + representations with multiple occurrences of the same key are + implicitly normalized: lookup and update only take the first + occurrence into account. +*} text {* - Recent versions of Poly/ML (5.2.1 or later) support robust - multithreaded execution, based on native operating system threads of - the underlying platform. Thus threads will actually be executed in - parallel on multi-core systems. A speedup-factor of approximately - 1.5--3 can be expected on a regular 4-core machine.\footnote{There - is some inherent limitation of the speedup factor due to garbage - collection, which is still sequential. It helps to provide initial - heap space generously, using the \texttt{-H} option of Poly/ML.} - Threads also help to organize advanced operations of the system, - with explicit communication between sub-components, real-time - conditions, time-outs etc. + \begin{mldecls} + @{index_ML AList.lookup: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> 'c option"} \\ + @{index_ML AList.defined: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> bool"} \\ + @{index_ML AList.update: "('a * 'a -> bool) -> 'a * 'b -> ('a * 'b) list -> ('a * 'b) list"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML AList.lookup}, @{ML AList.defined}, @{ML AList.update} + implement the main ``framework operations'' for mappings in + Isabelle/ML, with standard conventions for names and types. + + Note that a function called @{text lookup} is obliged to express its + partiality via an explicit option element. There is no choice to + raise an exception, without changing the name to something like + @{text "the_element"} or @{text "get"}. + + The @{text "defined"} operation is essentially a contraction of @{ML + is_some} and @{text "lookup"}, but this is sufficiently frequent to + justify its independent existence. This also gives the + implementation some opportunity for peep-hole optimization. + + \end{description} + + Association lists are adequate as simple and light-weight + implementation of finite mappings in many practical situations. A + more heavy-duty table structure is defined in @{"file" + "~~/src/Pure/General/table.ML"}; that version scales easily to + thousands or millions of elements. +*} + + +subsection {* Unsynchronized references *} + +text %mlref {* + \begin{mldecls} + @{index_ML_type "'a Unsynchronized.ref"} \\ + @{index_ML Unsynchronized.ref: "'a -> 'a Unsynchronized.ref"} \\ + @{index_ML "!": "'a Unsynchronized.ref -> 'a"} \\ + @{index_ML "op :=": "'a Unsynchronized.ref * 'a -> unit"} \\ + \end{mldecls} +*} + +text {* Due to ubiquitous parallelism in Isabelle/ML (see also + \secref{sec:multi-threading}), the mutable reference cells of + Standard ML are notorious for causing problems. In a highly + parallel system, both correctness \emph{and} performance are easily + degraded when using mutable data. + + The unwieldy name of @{ML Unsynchronized.ref} for the constructor + for references in Isabelle/ML emphasizes the inconveniences caused by + mutability. Existing operations @{ML "!"} and @{ML "op :="} are + unchanged, but should be used with special precautions, say in a + strictly local situation that is guaranteed to be restricted to + sequential evaluation --- now and in the future. *} + + +section {* Thread-safe programming \label{sec:multi-threading} *} + +text {* Multi-threaded execution has become an everyday reality in + Isabelle since Poly/ML 5.2.1 and Isabelle2008. Isabelle/ML provides + implicit and explicit parallelism by default, and there is no way + for user-space tools to ``opt out''. ML programs that are purely + functional, output messages only via the official channels + (\secref{sec:message-channels}), and do not intercept interrupts + (\secref{sec:exceptions}) can participate in the multi-threaded + environment immediately without further ado. - Threads lack the memory protection of separate processes, and - operate concurrently on shared heap memory. This has the advantage - that results of independent computations are immediately available - to other threads, without requiring untyped character streams, - awkward serialization etc. + More ambitious tools with more fine-grained interaction with the + environment need to observe the principles explained below. +*} + + +subsection {* Multi-threading with shared memory *} + +text {* Multiple threads help to organize advanced operations of the + system, such as real-time conditions on command transactions, + sub-components with explicit communication, general asynchronous + interaction etc. Moreover, parallel evaluation is a prerequisite to + make adequate use of the CPU resources that are available on + multi-core systems.\footnote{Multi-core computing does not mean that + there are ``spare cycles'' to be wasted. It means that the + continued exponential speedup of CPU performance due to ``Moore's + Law'' follows different rules: clock frequency has reached its peak + around 2005, and applications need to be parallelized in order to + avoid a perceived loss of performance. See also + \cite{Sutter:2005}.} + + Isabelle/Isar exploits the inherent structure of theories and proofs + to support \emph{implicit parallelism} to a large extent. LCF-style + theorem provides almost ideal conditions for that; see also + \cite{Wenzel:2009}. This means, significant parts of theory and + proof checking is parallelized by default. A maximum speedup-factor + of 3.0 on 4 cores and 5.0 on 8 cores can be + expected.\footnote{Further scalability is limited due to garbage + collection, which is still sequential in Poly/ML 5.2/5.3/5.4. It + helps to provide initial heap space generously, using the + \texttt{-H} option. Initial heap size needs to be scaled-up + together with the number of CPU cores: approximately 1--2\,GB per + core..} - On the other hand, some programming guidelines need to be observed - in order to make unprotected parallelism work out smoothly. While - the ML system implementation is responsible to maintain basic - integrity of the representation of ML values in memory, the - application programmer needs to ensure that multithreaded execution - does not break the intended semantics. + \medskip ML threads lack the memory protection of separate + processes, and operate concurrently on shared heap memory. This has + the advantage that results of independent computations are + immediately available to other threads: abstract values can be + passed between threads without copying or awkward serialization that + is typically required for explicit message passing between separate + processes. + + To make shared-memory multi-threading work robustly and efficiently, + some programming guidelines need to be observed. While the ML + system is responsible to maintain basic integrity of the + representation of ML values in memory, the application programmer + needs to ensure that multi-threaded execution does not break the + intended semantics. - \medskip \paragraph{Critical shared resources.} Actually only those - parts outside the purely functional world of ML are critical. In - particular, this covers + \begin{warn} + To participate in implicit parallelism, tools need to be + thread-safe. A single ill-behaved tool can affect the stability and + performance of the whole system. + \end{warn} + + Apart from observing the principles of thread-safeness passively, + advanced tools may also exploit parallelism actively, e.g.\ by using + ``future values'' (\secref{sec:futures}) or the more basic library + functions for parallel list operations (\secref{sec:parlist}). + + \begin{warn} + Parallel computing resources are managed centrally by the + Isabelle/ML infrastructure. User programs must not fork their own + ML threads to perform computations. + \end{warn} +*} + + +subsection {* Critical shared resources *} + +text {* Thread-safeness is mainly concerned about concurrent + read/write access to shared resources, which are outside the purely + functional world of ML. This covers the following in particular. \begin{itemize} - \item global references (or arrays), i.e.\ those that persist over - several invocations of associated operations,\footnote{This is - independent of the visibility of such mutable values in the toplevel - scope.} + \item Global references (or arrays), i.e.\ mutable memory cells that + persist over several invocations of associated + operations.\footnote{This is independent of the visibility of such + mutable values in the toplevel scope.} - \item direct I/O on shared channels, notably @{text "stdin"}, @{text - "stdout"}, @{text "stderr"}. + \item Global state of the running Isabelle/ML process, i.e.\ raw I/O + channels, environment variables, current working directory. + + \item Writable resources in the file-system that are shared among + different threads or other processes. \end{itemize} - The majority of tools implemented within the Isabelle/Isar framework - will not require any of these critical elements: nothing special - needs to be observed when staying in the purely functional fragment - of ML. Note that output via the official Isabelle channels does not - count as direct I/O, so the operations @{ML "writeln"}, @{ML - "warning"}, @{ML "tracing"} etc.\ are safe. - - Moreover, ML bindings within the toplevel environment (@{verbatim - "type"}, @{verbatim val}, @{verbatim "structure"} etc.) due to - run-time invocation of the compiler are also safe, because - Isabelle/Isar manages this as part of the theory or proof context. - - \paragraph{Multithreading in Isabelle/Isar.} The theory loader - automatically exploits the overall parallelism of independent nodes - in the development graph, as well as the inherent irrelevance of - proofs for goals being fully specified in advance. This means, - checking of individual Isar proofs is parallelized by default. - Beyond that, very sophisticated proof tools may use local - parallelism internally, via the general programming model of - ``future values'' (see also @{"file" - "~~/src/Pure/Concurrent/future.ML"}). - - Any ML code that works relatively to the present background theory - is already safe. Contextual data may be easily stored within the - theory or proof context, thanks to the generic data concept of - Isabelle/Isar (see \secref{sec:context-data}). This greatly - diminishes the demand for global state information in the first - place. - - \medskip In rare situations where actual mutable content needs to be - manipulated, Isabelle provides a single \emph{critical section} that - may be entered while preventing any other thread from doing the - same. Entering the critical section without contention is very - fast, and several basic system operations do so frequently. This - also means that each thread should leave the critical section - quickly, otherwise parallel execution performance may degrade - significantly. - - Despite this potential bottle-neck, centralized locking is - convenient, because it prevents deadlocks without demanding special - precautions. Explicit communication demands other means, though. - The high-level abstraction of synchronized variables @{"file" - "~~/src/Pure/Concurrent/synchronized.ML"} enables parallel - components to communicate via shared state; see also @{"file" - "~~/src/Pure/Concurrent/mailbox.ML"} as canonical example. - - \paragraph{Good conduct of impure programs.} The following - guidelines enable non-functional programs to participate in - multithreading. + Isabelle/ML provides various mechanisms to avoid critical shared + resources in most practical situations. As last resort there are + some mechanisms for explicit synchronization. The following + guidelines help to make Isabelle/ML programs work smoothly in a + concurrent environment. \begin{itemize} - \item Minimize global state information. Using proper theory and - proof context data will actually return to functional update of - values, without any special precautions for multithreading. Apart - from the fully general functors for theory and proof data (see - \secref{sec:context-data}) there are drop-in replacements that - emulate primitive references for common cases of \emph{configuration - options} for type @{ML_type "bool"}/@{ML_type "int"}/@{ML_type - "string"} (see structure @{ML_struct Config} and @{ML - Attrib.config_bool} etc.), and lists of theorems (see functor - @{ML_functor Named_Thms}). + \item Avoid global references altogether. Isabelle/Isar maintains a + uniform context that incorporates arbitrary data declared by user + programs (\secref{sec:context-data}). This context is passed as + plain value and user tools can get/map their own data in a purely + functional manner. Configuration options within the context + (\secref{sec:config-options}) provide simple drop-in replacements + for formerly writable flags. - \item Keep components with local state information - \emph{re-entrant}. Instead of poking initial values into (private) - global references, create a new state record on each invocation, and - pass that through any auxiliary functions of the component. The - state record may well contain mutable references, without requiring - any special synchronizations, as long as each invocation sees its - own copy. Occasionally, one might even return to plain functional - updates on non-mutable record values here. + \item Keep components with local state information re-entrant. + Instead of poking initial values into (private) global references, a + new state record can be created on each invocation, and passed + through any auxiliary functions of the component. The state record + may well contain mutable references, without requiring any special + synchronizations, as long as each invocation gets its own copy. - \item Isolate process configuration flags. The main legitimate - application of global references is to configure the whole process - in a certain way, essentially affecting all threads. A typical - example is the @{ML show_types} flag, which tells the pretty printer - to output explicit type information for terms. Such flags usually - do not affect the functionality of the core system, but only the - view being presented to the user. + \item Avoid raw output on @{text "stdout"} or @{text "stderr"}. The + Poly/ML library is thread-safe for each individual output operation, + but the ordering of parallel invocations is arbitrary. This means + raw output will appear on some system console with unpredictable + interleaving of atomic chunks. - Occasionally, such global process flags are treated like implicit - arguments to certain operations, by using the @{ML setmp_CRITICAL} combinator - for safe temporary assignment. Its traditional purpose was to - ensure proper recovery of the original value when exceptions are - raised in the body, now the functionality is extended to enter the - \emph{critical section} (with its usual potential of degrading - parallelism). + Note that this does not affect regular message output channels + (\secref{sec:message-channels}). An official message is associated + with the command transaction from where it originates, independently + of other transactions. This means each running Isar command has + effectively its own set of message channels, and interleaving can + only happen when commands use parallelism internally (and only at + message boundaries). - Note that recovery of plain value passing semantics via @{ML - setmp_CRITICAL}~@{text "ref value"} assumes that this @{text "ref"} is - exclusively manipulated within the critical section. In particular, - any persistent global assignment of @{text "ref := value"} needs to - be marked critical as well, to prevent intruding another threads - local view, and a lost-update in the global scope, too. + \item Treat environment variables and the current working directory + of the running process as strictly read-only. + + \item Restrict writing to the file-system to unique temporary files. + Isabelle already provides a temporary directory that is unique for + the running process, and there is a centralized source of unique + serial numbers in Isabelle/ML. Thus temporary files that are passed + to to some external process will be always disjoint, and thus + thread-safe. \end{itemize} +*} - Recall that in an open ``LCF-style'' system like Isabelle/Isar, the - user participates in constructing the overall environment. This - means that state-based facilities offered by one component will - require special caution later on. So minimizing critical elements, - by staying within the plain value-oriented view relative to theory - or proof contexts most of the time, will also reduce the chance of - mishaps occurring to end-users. +text %mlref {* + \begin{mldecls} + @{index_ML File.tmp_path: "Path.T -> Path.T"} \\ + @{index_ML serial_string: "unit -> string"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML File.tmp_path}~@{text "path"} relocates the base + component of @{text "path"} into the unique temporary directory of + the running Isabelle/ML process. + + \item @{ML serial_string}~@{text "()"} creates a new serial number + that is unique over the runtime of the Isabelle/ML process. + + \end{description} +*} + +text %mlex {* The following example shows how to create unique + temporary file names. +*} + +ML {* + val tmp1 = File.tmp_path (Path.basic ("foo" ^ serial_string ())); + val tmp2 = File.tmp_path (Path.basic ("foo" ^ serial_string ())); + @{assert} (tmp1 <> tmp2); +*} + + +subsection {* Explicit synchronization *} + +text {* Isabelle/ML also provides some explicit synchronization + mechanisms, for the rare situations where mutable shared resources + are really required. These are based on the synchronizations + primitives of Poly/ML, which have been adapted to the specific + assumptions of the concurrent Isabelle/ML environment. User code + must not use the Poly/ML primitives directly! + + \medskip The most basic synchronization concept is a single + \emph{critical section} (also called ``monitor'' in the literature). + A thread that enters the critical section prevents all other threads + from doing the same. A thread that is already within the critical + section may re-enter it in an idempotent manner. + + Such centralized locking is convenient, because it prevents + deadlocks by construction. + + \medskip More fine-grained locking works via \emph{synchronized + variables}. An explicit state component is associated with + mechanisms for locking and signaling. There are operations to + await a condition, change the state, and signal the change to all + other waiting threads. + + Here the synchronized access to the state variable is \emph{not} + re-entrant: direct or indirect nesting within the same thread will + cause a deadlock! *} text %mlref {* \begin{mldecls} @{index_ML NAMED_CRITICAL: "string -> (unit -> 'a) -> 'a"} \\ @{index_ML CRITICAL: "(unit -> 'a) -> 'a"} \\ - @{index_ML setmp_CRITICAL: "'a Unsynchronized.ref -> 'a -> ('b -> 'c) -> 'b -> 'c"} \\ + \end{mldecls} + \begin{mldecls} + @{index_ML_type "'a Synchronized.var"} \\ + @{index_ML Synchronized.var: "string -> 'a -> 'a Synchronized.var"} \\ + @{index_ML Synchronized.guarded_access: "'a Synchronized.var -> + ('a -> ('b * 'a) option) -> 'b"} \\ \end{mldecls} \begin{description} - \item @{ML NAMED_CRITICAL}~@{text "name f"} evaluates @{text "f ()"} - while staying within the critical section of Isabelle/Isar. No - other thread may do so at the same time, but non-critical parallel - execution will continue. The @{text "name"} argument serves for - diagnostic purposes and might help to spot sources of congestion. + \item @{ML NAMED_CRITICAL}~@{text "name e"} evaluates @{text "e ()"} + within the central critical section of Isabelle/ML. No other thread + may do so at the same time, but non-critical parallel execution will + continue. The @{text "name"} argument is used for tracing and might + help to spot sources of congestion. + + Entering the critical section without contention is very fast, and + several basic system operations do so frequently. Each thread + should leave the critical section quickly, otherwise parallel + performance may degrade. \item @{ML CRITICAL} is the same as @{ML NAMED_CRITICAL} with empty name argument. - \item @{ML setmp_CRITICAL}~@{text "ref value f x"} evaluates @{text "f x"} - while staying within the critical section and having @{text "ref := - value"} assigned temporarily. This recovers a value-passing - semantics involving global references, regardless of exceptions or - concurrency. + \item Type @{ML_type "'a Synchronized.var"} represents synchronized + variables with state of type @{ML_type 'a}. + + \item @{ML Synchronized.var}~@{text "name x"} creates a synchronized + variable that is initialized with value @{text "x"}. The @{text + "name"} is used for tracing. + + \item @{ML Synchronized.guarded_access}~@{text "var f"} lets the + function @{text "f"} operate within a critical section on the state + @{text "x"} as follows: if @{text "f x"} produces @{ML NONE}, we + continue to wait on the internal condition variable, expecting that + some other thread will eventually change the content in a suitable + manner; if @{text "f x"} produces @{ML SOME}~@{text "(y, x')"} we + are satisfied and assign the new state value @{text "x'"}, broadcast + a signal to all waiting threads on the associated condition + variable, and return the result @{text "y"}. \end{description} -*} - -chapter {* Basic library functions *} - -text {* - Beyond the proposal of the SML/NJ basis library, Isabelle comes - with its own library, from which selected parts are given here. - These should encourage a study of the Isabelle sources, - in particular files \emph{Pure/library.ML} and \emph{Pure/General/*.ML}. -*} - -section {* Linear transformations \label{sec:ML-linear-trans} *} - -text %mlref {* - \begin{mldecls} - @{index_ML "op |> ": "'a * ('a -> 'b) -> 'b"} \\ - \end{mldecls} -*} - -(*<*) -typedecl foo -consts foo :: foo -ML {* -val thy = Theory.copy @{theory} -*} -(*>*) - -text {* - \noindent Many problems in functional programming can be thought of - as linear transformations, i.e.~a caluclation starts with a - particular value @{ML_text "x : foo"} which is then transformed - by application of a function @{ML_text "f : foo -> foo"}, - continued by an application of a function @{ML_text "g : foo -> bar"}, - and so on. As a canoncial example, take functions enriching - a theory by constant declararion and primitive definitions: - - \smallskip\begin{mldecls} - @{ML "Sign.declare_const: (binding * typ) * mixfix - -> theory -> term * theory"} \\ - @{ML "Thm.add_def: bool -> bool -> binding * term -> theory -> (string * thm) * theory"} - \end{mldecls} - - \noindent Written with naive application, an addition of constant - @{term bar} with type @{typ "foo \ foo"} and - a corresponding definition @{term "bar \ \x. x"} would look like: - - \smallskip\begin{mldecls} - @{ML "(fn (t, thy) => Thm.add_def false false - (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})) thy) - (Sign.declare_const - ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn) thy)"} - \end{mldecls} - - \noindent With increasing numbers of applications, this code gets quite - unreadable. Further, it is unintuitive that things are - written down in the opposite order as they actually ``happen''. -*} - -text {* - \noindent At this stage, Isabelle offers some combinators which allow - for more convenient notation, most notably reverse application: - - \smallskip\begin{mldecls} -@{ML "thy -|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn) -|> (fn (t, thy) => thy -|> Thm.add_def false false - (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))"} - \end{mldecls} -*} - -text %mlref {* - \begin{mldecls} - @{index_ML "op |-> ": "('c * 'a) * ('c -> 'a -> 'b) -> 'b"} \\ - @{index_ML "op |>> ": "('a * 'c) * ('a -> 'b) -> 'b * 'c"} \\ - @{index_ML "op ||> ": "('c * 'a) * ('a -> 'b) -> 'c * 'b"} \\ - @{index_ML "op ||>> ": "('c * 'a) * ('a -> 'd * 'b) -> ('c * 'd) * 'b"} \\ - \end{mldecls} -*} - -text {* - \noindent Usually, functions involving theory updates also return - side results; to use these conveniently, yet another - set of combinators is at hand, most notably @{ML "op |->"} - which allows curried access to side results: - - \smallskip\begin{mldecls} -@{ML "thy -|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn) -|-> (fn t => Thm.add_def false false - (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"}))) -"} - \end{mldecls} - - \noindent @{ML "op |>>"} allows for processing side results on their own: - - \smallskip\begin{mldecls} -@{ML "thy -|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn) -|>> (fn t => Logic.mk_equals (t, @{term \"%x. x\"})) -|-> (fn def => Thm.add_def false false (Binding.name \"bar_def\", def)) -"} - \end{mldecls} - - \noindent Orthogonally, @{ML "op ||>"} applies a transformation - in the presence of side results which are left unchanged: - - \smallskip\begin{mldecls} -@{ML "thy -|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn) -||> Sign.add_path \"foobar\" -|-> (fn t => Thm.add_def false false - (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"}))) -||> Sign.restore_naming thy -"} - \end{mldecls} - - \noindent @{ML "op ||>>"} accumulates side results: - - \smallskip\begin{mldecls} -@{ML "thy -|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn) -||>> Sign.declare_const ((Binding.name \"foobar\", @{typ \"foo => foo\"}), NoSyn) -|-> (fn (t1, t2) => Thm.add_def false false - (Binding.name \"bar_def\", Logic.mk_equals (t1, t2))) -"} - \end{mldecls} -*} - -text %mlref {* - \begin{mldecls} - @{index_ML fold: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\ - @{index_ML fold_map: "('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b"} \\ - \end{mldecls} + There are some further variants of the general @{ML + Synchronized.guarded_access} combinator, see @{"file" + "~~/src/Pure/Concurrent/synchronized.ML"} for details. *} -text {* - \noindent This principles naturally lift to \emph{lists} using - the @{ML fold} and @{ML fold_map} combinators. - The first lifts a single function - \begin{quote}\footnotesize - @{ML_text "f : 'a -> 'b -> 'b"} to @{ML_text "'a list -> 'b -> 'b"} - \end{quote} - such that - \begin{quote}\footnotesize - @{ML_text "y |> fold f [x1, x2, ..., x_n]"} \\ - \hspace*{2ex}@{text "\"} @{ML_text "y |> f x1 |> f x2 |> ... |> f x_n"} - \end{quote} - \noindent The second accumulates side results in a list by lifting - a single function - \begin{quote}\footnotesize - @{ML_text "f : 'a -> 'b -> 'c * 'b"} to @{ML_text "'a list -> 'b -> 'c list * 'b"} - \end{quote} - such that - \begin{quote}\footnotesize - @{ML_text "y |> fold_map f [x1, x2, ..., x_n]"} \\ - \hspace*{2ex}@{text "\"} @{ML_text "y |> f x1 ||>> f x2 ||>> ... ||>> f x_n"} \\ - \hspace*{6ex}@{ML_text "||> (fn ((z1, z2), ..., z_n) => [z1, z2, ..., z_n])"} - \end{quote} - - \noindent Example: - - \smallskip\begin{mldecls} -@{ML "let - val consts = [\"foo\", \"bar\"]; -in - thy - |> fold_map (fn const => Sign.declare_const - ((Binding.name const, @{typ \"foo => foo\"}), NoSyn)) consts - |>> map (fn t => Logic.mk_equals (t, @{term \"%x. x\"})) - |-> (fn defs => fold_map (fn def => - Thm.add_def false false (Binding.empty, def)) defs) -end"} - \end{mldecls} -*} - -text %mlref {* - \begin{mldecls} - @{index_ML "op #> ": "('a -> 'b) * ('b -> 'c) -> 'a -> 'c"} \\ - @{index_ML "op #-> ": "('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd"} \\ - @{index_ML "op #>> ": "('a -> 'c * 'b) * ('c -> 'd) -> 'a -> 'd * 'b"} \\ - @{index_ML "op ##> ": "('a -> 'c * 'b) * ('b -> 'd) -> 'a -> 'c * 'd"} \\ - @{index_ML "op ##>> ": "('a -> 'c * 'b) * ('b -> 'e * 'd) -> 'a -> ('c * 'e) * 'd"} \\ - \end{mldecls} -*} +text %mlex {* See @{"file" "~~/src/Pure/Concurrent/mailbox.ML"} how to + implement a mailbox as synchronized variable over a purely + functional queue. -text {* - \noindent All those linear combinators also exist in higher-order - variants which do not expect a value on the left hand side - but a function. -*} - -text %mlref {* - \begin{mldecls} - @{index_ML "op ` ": "('b -> 'a) -> 'b -> 'a * 'b"} \\ - @{index_ML tap: "('b -> 'a) -> 'b -> 'b"} \\ - \end{mldecls} -*} - -text {* - \noindent These operators allow to ``query'' a context - in a series of context transformations: - - \smallskip\begin{mldecls} -@{ML "thy -|> tap (fn _ => writeln \"now adding constant\") -|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn) -||>> `(fn thy => Sign.declared_const thy - (Sign.full_name thy (Binding.name \"foobar\"))) -|-> (fn (t, b) => if b then I - else Sign.declare_const - ((Binding.name \"foobar\", @{typ \"foo => foo\"}), NoSyn) #> snd) -"} - \end{mldecls} -*} - -section {* Options and partiality *} - -text %mlref {* - \begin{mldecls} - @{index_ML is_some: "'a option -> bool"} \\ - @{index_ML is_none: "'a option -> bool"} \\ - @{index_ML the: "'a option -> 'a"} \\ - @{index_ML these: "'a list option -> 'a list"} \\ - @{index_ML the_list: "'a option -> 'a list"} \\ - @{index_ML the_default: "'a -> 'a option -> 'a"} \\ - @{index_ML try: "('a -> 'b) -> 'a -> 'b option"} \\ - @{index_ML can: "('a -> 'b) -> 'a -> bool"} \\ - \end{mldecls} -*} - -text {* - Standard selector functions on @{text option}s are provided. The - @{ML try} and @{ML can} functions provide a convenient interface for - handling exceptions -- both take as arguments a function @{ML_text f} - together with a parameter @{ML_text x} and handle any exception during - the evaluation of the application of @{ML_text f} to @{ML_text x}, either - return a lifted result (@{ML NONE} on failure) or a boolean value - (@{ML false} on failure). + \medskip The following example implements a counter that produces + positive integers that are unique over the runtime of the Isabelle + process: *} - -section {* Common data structures *} - -subsection {* Lists (as set-like data structures) *} - -text {* - \begin{mldecls} - @{index_ML member: "('b * 'a -> bool) -> 'a list -> 'b -> bool"} \\ - @{index_ML insert: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\ - @{index_ML remove: "('b * 'a -> bool) -> 'b -> 'a list -> 'a list"} \\ - @{index_ML merge: "('a * 'a -> bool) -> 'a list * 'a list -> 'a list"} \\ - \end{mldecls} -*} - -text {* - Lists are often used as set-like data structures -- set-like in - the sense that they support a notion of @{ML member}-ship, - @{ML insert}-ing and @{ML remove}-ing, but are order-sensitive. - This is convenient when implementing a history-like mechanism: - @{ML insert} adds an element \emph{to the front} of a list, - if not yet present; @{ML remove} removes \emph{all} occurences - of a particular element. Correspondingly @{ML merge} implements a - a merge on two lists suitable for merges of context data - (\secref{sec:context-theory}). - - Functions are parametrized by an explicit equality function - to accomplish overloaded equality; in most cases of monomorphic - equality, writing @{ML "op ="} should suffice. -*} - -subsection {* Association lists *} - -text {* - \begin{mldecls} - @{index_ML_exn AList.DUP} \\ - @{index_ML AList.lookup: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> 'c option"} \\ - @{index_ML AList.defined: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> bool"} \\ - @{index_ML AList.update: "('a * 'a -> bool) -> ('a * 'b) -> ('a * 'b) list -> ('a * 'b) list"} \\ - @{index_ML AList.default: "('a * 'a -> bool) -> ('a * 'b) -> ('a * 'b) list -> ('a * 'b) list"} \\ - @{index_ML AList.delete: "('a * 'b -> bool) -> 'a -> ('b * 'c) list -> ('b * 'c) list"} \\ - @{index_ML AList.map_entry: "('a * 'b -> bool) -> 'a - -> ('c -> 'c) -> ('b * 'c) list -> ('b * 'c) list"} \\ - @{index_ML AList.map_default: "('a * 'a -> bool) -> 'a * 'b -> ('b -> 'b) - -> ('a * 'b) list -> ('a * 'b) list"} \\ - @{index_ML AList.join: "('a * 'a -> bool) -> ('a -> 'b * 'b -> 'b) (*exception DUP*) - -> ('a * 'b) list * ('a * 'b) list -> ('a * 'b) list (*exception AList.DUP*)"} \\ - @{index_ML AList.merge: "('a * 'a -> bool) -> ('b * 'b -> bool) - -> ('a * 'b) list * ('a * 'b) list -> ('a * 'b) list (*exception AList.DUP*)"} - \end{mldecls} +ML {* + local + val counter = Synchronized.var "counter" 0; + in + fun next () = + Synchronized.guarded_access counter + (fn i => + let val j = i + 1 + in SOME (j, j) end); + end; *} -text {* - Association lists can be seens as an extension of set-like lists: - on the one hand, they may be used to implement finite mappings, - on the other hand, they remain order-sensitive and allow for - multiple key-value-pair with the same key: @{ML AList.lookup} - returns the \emph{first} value corresponding to a particular - key, if present. @{ML AList.update} updates - the \emph{first} occurence of a particular key; if no such - key exists yet, the key-value-pair is added \emph{to the front}. - @{ML AList.delete} only deletes the \emph{first} occurence of a key. - @{ML AList.merge} provides an operation suitable for merges of context data - (\secref{sec:context-theory}), where an equality parameter on - values determines whether a merge should be considered a conflict. - A slightly generalized operation if implementend by the @{ML AList.join} - function which allows for explicit conflict resolution. -*} - -subsection {* Tables *} - -text {* - \begin{mldecls} - @{index_ML_type "'a Symtab.table"} \\ - @{index_ML_exn Symtab.DUP: string} \\ - @{index_ML_exn Symtab.SAME} \\ - @{index_ML_exn Symtab.UNDEF: string} \\ - @{index_ML Symtab.empty: "'a Symtab.table"} \\ - @{index_ML Symtab.lookup: "'a Symtab.table -> string -> 'a option"} \\ - @{index_ML Symtab.defined: "'a Symtab.table -> string -> bool"} \\ - @{index_ML Symtab.update: "(string * 'a) -> 'a Symtab.table -> 'a Symtab.table"} \\ - @{index_ML Symtab.default: "string * 'a -> 'a Symtab.table -> 'a Symtab.table"} \\ - @{index_ML Symtab.delete: "string - -> 'a Symtab.table -> 'a Symtab.table (*exception Symtab.UNDEF*)"} \\ - @{index_ML Symtab.map_entry: "string -> ('a -> 'a) - -> 'a Symtab.table -> 'a Symtab.table"} \\ - @{index_ML Symtab.map_default: "(string * 'a) -> ('a -> 'a) - -> 'a Symtab.table -> 'a Symtab.table"} \\ - @{index_ML Symtab.join: "(string -> 'a * 'a -> 'a) (*exception Symtab.DUP/Symtab.SAME*) - -> 'a Symtab.table * 'a Symtab.table - -> 'a Symtab.table (*exception Symtab.DUP*)"} \\ - @{index_ML Symtab.merge: "('a * 'a -> bool) - -> 'a Symtab.table * 'a Symtab.table - -> 'a Symtab.table (*exception Symtab.DUP*)"} - \end{mldecls} -*} - -text {* - Tables are an efficient representation of finite mappings without - any notion of order; due to their efficiency they should be used - whenever such pure finite mappings are neccessary. - - The key type of tables must be given explicitly by instantiating - the @{ML_functor Table} functor which takes the key type - together with its @{ML_type order}; for convience, we restrict - here to the @{ML_struct Symtab} instance with @{ML_type string} - as key type. - - Most table functions correspond to those of association lists. +ML {* + val a = next (); + val b = next (); + @{assert} (a <> b); *} end \ No newline at end of file diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/Prelim.thy --- a/doc-src/IsarImplementation/Thy/Prelim.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/Prelim.thy Mon Oct 25 08:08:08 2010 +0200 @@ -140,23 +140,26 @@ \medskip There is a separate notion of \emph{theory reference} for maintaining a live link to an evolving theory context: updates on - drafts are propagated automatically. Dynamic updating stops after - an explicit @{text "end"} only. + drafts are propagated automatically. Dynamic updating stops when + the next @{text "checkpoint"} is reached. Derived entities may store a theory reference in order to indicate - the context they belong to. This implicitly assumes monotonic - reasoning, because the referenced context may become larger without - further notice. + the formal context from which they are derived. This implicitly + assumes monotonic reasoning, because the referenced context may + become larger without further notice. *} text %mlref {* \begin{mldecls} @{index_ML_type theory} \\ + @{index_ML Theory.eq_thy: "theory * theory -> bool"} \\ @{index_ML Theory.subthy: "theory * theory -> bool"} \\ @{index_ML Theory.checkpoint: "theory -> theory"} \\ @{index_ML Theory.copy: "theory -> theory"} \\ @{index_ML Theory.merge: "theory * theory -> theory"} \\ @{index_ML Theory.begin_theory: "string -> theory list -> theory"} \\ + @{index_ML Theory.parents_of: "theory -> theory list"} \\ + @{index_ML Theory.ancestors_of: "theory -> theory list"} \\ \end{mldecls} \begin{mldecls} @{index_ML_type theory_ref} \\ @@ -166,10 +169,15 @@ \begin{description} - \item @{ML_type theory} represents theory contexts. This is - essentially a linear type, with explicit runtime checking! Most - internal theory operations destroy the original version, which then - becomes ``stale''. + \item Type @{ML_type theory} represents theory contexts. This is + essentially a linear type, with explicit runtime checking. + Primitive theory operations destroy the original version, which then + becomes ``stale''. This can be prevented by explicit checkpointing, + which the system does at least at the boundary of toplevel command + transactions \secref{sec:isar-toplevel}. + + \item @{ML "Theory.eq_thy"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} check strict + identity of two theories. \item @{ML "Theory.subthy"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} compares theories according to the intrinsic graph structure of the construction. @@ -191,11 +199,17 @@ This version of ad-hoc theory merge fails for unrelated theories! \item @{ML "Theory.begin_theory"}~@{text "name parents"} constructs - a new theory based on the given parents. This {\ML} function is + a new theory based on the given parents. This ML function is normally not invoked directly. - \item @{ML_type theory_ref} represents a sliding reference to an - always valid theory; updates on the original are propagated + \item @{ML "Theory.parents_of"}~@{text "thy"} returns the direct + ancestors of @{text thy}. + + \item @{ML "Theory.ancestors_of"}~@{text "thy"} returns all + ancestors of @{text thy} (not including @{text thy} itself). + + \item Type @{ML_type theory_ref} represents a sliding reference to + an always valid theory; updates on the original are propagated automatically. \item @{ML "Theory.deref"}~@{text "thy_ref"} turns a @{ML_type @@ -209,12 +223,39 @@ \end{description} *} +text %mlantiq {* + \begin{matharray}{rcl} + @{ML_antiquotation_def "theory"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "theory_ref"} & : & @{text ML_antiquotation} \\ + \end{matharray} + + \begin{rail} + ('theory' | 'theory\_ref') nameref? + ; + \end{rail} + + \begin{description} + + \item @{text "@{theory}"} refers to the background theory of the + current context --- as abstract value. + + \item @{text "@{theory A}"} refers to an explicitly named ancestor + theory @{text "A"} of the background theory of the current context + --- as abstract value. + + \item @{text "@{theory_ref}"} is similar to @{text "@{theory}"}, but + produces a @{ML_type theory_ref} via @{ML "Theory.check_thy"} as + explained above. + + \end{description} +*} + subsection {* Proof context \label{sec:context-proof} *} text {* A proof context is a container for pure data with a - back-reference to the theory it belongs to. The @{text "init"} - operation creates a proof context from a given theory. + back-reference to the theory from which it is derived. The @{text + "init"} operation creates a proof context from a given theory. Modifications to draft theories are propagated to the proof context as usual, but there is also an explicit @{text "transfer"} operation to force resynchronization with more substantial updates to the @@ -250,9 +291,9 @@ \begin{description} - \item @{ML_type Proof.context} represents proof contexts. Elements - of this type are essentially pure values, with a sliding reference - to the background theory. + \item Type @{ML_type Proof.context} represents proof contexts. + Elements of this type are essentially pure values, with a sliding + reference to the background theory. \item @{ML ProofContext.init_global}~@{text "thy"} produces a proof context derived from @{text "thy"}, initializing all data. @@ -268,6 +309,23 @@ \end{description} *} +text %mlantiq {* + \begin{matharray}{rcl} + @{ML_antiquotation_def "context"} & : & @{text ML_antiquotation} \\ + \end{matharray} + + \begin{description} + + \item @{text "@{context}"} refers to \emph{the} context at + compile-time --- as abstract value. Independently of (local) theory + or proof mode, this always produces a meaningful result. + + This is probably the most common antiquotation in interactive + experimentation with ML inside Isar. + + \end{description} +*} + subsection {* Generic contexts \label{sec:generic-context} *} @@ -295,7 +353,7 @@ \begin{description} - \item @{ML_type Context.generic} is the direct sum of @{ML_type + \item Type @{ML_type Context.generic} is the direct sum of @{ML_type "theory"} and @{ML_type "Proof.context"}, with the datatype constructors @{ML "Context.Theory"} and @{ML "Context.Proof"}. @@ -331,9 +389,9 @@ \end{tabular} \medskip - \noindent The @{text "empty"} value acts as initial default for - \emph{any} theory that does not declare actual data content; @{text - "extend"} is acts like a unitary version of @{text "merge"}. + The @{text "empty"} value acts as initial default for \emph{any} + theory that does not declare actual data content; @{text "extend"} + is acts like a unitary version of @{text "merge"}. Implementing @{text "merge"} can be tricky. The general idea is that @{text "merge (data\<^sub>1, data\<^sub>2)"} inserts those parts of @{text @@ -355,19 +413,24 @@ \end{tabular} \medskip - \noindent The @{text "init"} operation is supposed to produce a pure - value from the given background theory and should be somehow + The @{text "init"} operation is supposed to produce a pure value + from the given background theory and should be somehow ``immediate''. Whenever a proof context is initialized, which happens frequently, the the system invokes the @{text "init"} - operation of \emph{all} theory data slots ever declared. + operation of \emph{all} theory data slots ever declared. This also + means that one needs to be economic about the total number of proof + data declarations in the system, i.e.\ each ML module should declare + at most one, sometimes two data slots for its internal use. + Repeated data declarations to simulate a record type should be + avoided! \paragraph{Generic data} provides a hybrid interface for both theory and proof data. The @{text "init"} operation for proof contexts is predefined to select the current data value from the background theory. - \bigskip Any of these data declaration over type @{text "T"} result - in an ML structure with the following signature: + \bigskip Any of the above data declarations over type @{text "T"} + result in an ML structure with the following signature: \medskip \begin{tabular}{ll} @@ -377,13 +440,12 @@ \end{tabular} \medskip - \noindent These other operations provide exclusive access for the - particular kind of context (theory, proof, or generic context). - This interface fully observes the ML discipline for types and - scopes: there is no other way to access the corresponding data slot - of a context. By keeping these operations private, an Isabelle/ML - module may maintain abstract values authentically. -*} + These other operations provide exclusive access for the particular + kind of context (theory, proof, or generic context). This interface + observes the ML discipline for types and scopes: there is no other + way to access the corresponding data slot of a context. By keeping + these operations private, an Isabelle/ML module may maintain + abstract values authentically. *} text %mlref {* \begin{mldecls} @@ -422,9 +484,9 @@ end; *} -text {* \noindent The implementation uses private theory data - internally, and only exposes an operation that involves explicit - argument checking wrt.\ the given theory. *} +text {* The implementation uses private theory data internally, and + only exposes an operation that involves explicit argument checking + wrt.\ the given theory. *} ML {* structure Wellformed_Terms: WELLFORMED_TERMS = @@ -437,28 +499,31 @@ val extend = I; fun merge (ts1, ts2) = Ord_List.union Term_Ord.fast_term_ord ts1 ts2; - ) + ); val get = Terms.get; fun add raw_t thy = - let val t = Sign.cert_term thy raw_t - in Terms.map (Ord_List.insert Term_Ord.fast_term_ord t) thy end; + let + val t = Sign.cert_term thy raw_t; + in + Terms.map (Ord_List.insert Term_Ord.fast_term_ord t) thy + end; end; *} -text {* We use @{ML_type "term Ord_List.T"} for reasonably efficient - representation of a set of terms: all operations are linear in the - number of stored elements. Here we assume that our users do not - care about the declaration order, since that data structure forces - its own arrangement of elements. +text {* Type @{ML_type "term Ord_List.T"} is used for reasonably + efficient representation of a set of terms: all operations are + linear in the number of stored elements. Here we assume that users + of this module do not care about the declaration order, since that + data structure forces its own arrangement of elements. Observe how the @{verbatim merge} operation joins the data slots of the two constituents: @{ML Ord_List.union} prevents duplication of common data from different branches, thus avoiding the danger of - exponential blowup. (Plain list append etc.\ must never be used for - theory data merges.) + exponential blowup. Plain list append etc.\ must never be used for + theory data merges! \medskip Our intended invariant is achieved as follows: \begin{enumerate} @@ -479,7 +544,7 @@ type-inference via @{ML Syntax.check_term} (cf.\ \secref{sec:term-check}) is not necessarily monotonic wrt.\ the background theory, since constraints of term constants can be - strengthened by later declarations, for example. + modified by later declarations, for example. In most cases, user-space context data does not have to take such invariants too seriously. The situation is different in the @@ -488,6 +553,111 @@ *} +subsection {* Configuration options \label{sec:config-options} *} + +text {* A \emph{configuration option} is a named optional value of + some basic type (Boolean, integer, string) that is stored in the + context. It is a simple application of general context data + (\secref{sec:context-data}) that is sufficiently common to justify + customized setup, which includes some concrete declarations for + end-users using existing notation for attributes (cf.\ + \secref{sec:attributes}). + + For example, the predefined configuration option @{attribute + show_types} controls output of explicit type constraints for + variables in printed terms (cf.\ \secref{sec:read-print}). Its + value can be modified within Isar text like this: +*} + +declare [[show_types = false]] + -- {* declaration within (local) theory context *} + +example_proof + note [[show_types = true]] + -- {* declaration within proof (forward mode) *} + term x + + have "x = x" + using [[show_types = false]] + -- {* declaration within proof (backward mode) *} + .. +qed + +text {* Configuration options that are not set explicitly hold a + default value that can depend on the application context. This + allows to retrieve the value from another slot within the context, + or fall back on a global preference mechanism, for example. + + The operations to declare configuration options and get/map their + values are modeled as direct replacements for historic global + references, only that the context is made explicit. This allows + easy configuration of tools, without relying on the execution order + as required for old-style mutable references. *} + +text %mlref {* + \begin{mldecls} + @{index_ML Config.get: "Proof.context -> 'a Config.T -> 'a"} \\ + @{index_ML Config.map: "'a Config.T -> ('a -> 'a) -> Proof.context -> Proof.context"} \\ + @{index_ML Attrib.config_bool: "string -> (Context.generic -> bool) -> + bool Config.T * (theory -> theory)"} \\ + @{index_ML Attrib.config_int: "string -> (Context.generic -> int) -> + int Config.T * (theory -> theory)"} \\ + @{index_ML Attrib.config_string: "string -> (Context.generic -> string) -> + string Config.T * (theory -> theory)"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML Config.get}~@{text "ctxt config"} gets the value of + @{text "config"} in the given context. + + \item @{ML Config.map}~@{text "config f ctxt"} updates the context + by updating the value of @{text "config"}. + + \item @{text "(config, setup) ="}~@{ML Attrib.config_bool}~@{text + "name default"} creates a named configuration option of type + @{ML_type bool}, with the given @{text "default"} depending on the + application context. The resulting @{text "config"} can be used to + get/map its value in a given context. The @{text "setup"} function + needs to be applied to the theory initially, in order to make + concrete declaration syntax available to the user. + + \item @{ML Attrib.config_int} and @{ML Attrib.config_string} work + like @{ML Attrib.config_bool}, but for types @{ML_type int} and + @{ML_type string}, respectively. + + \end{description} +*} + +text %mlex {* The following example shows how to declare and use a + Boolean configuration option called @{text "my_flag"} with constant + default value @{ML false}. *} + +ML {* + val (my_flag, my_flag_setup) = + Attrib.config_bool "my_flag" (K false) +*} +setup my_flag_setup + +text {* Now the user can refer to @{attribute my_flag} in + declarations, while we can retrieve the current value from the + context via @{ML Config.get}. *} + +ML_val {* @{assert} (Config.get @{context} my_flag = false) *} + +declare [[my_flag = true]] + +ML_val {* @{assert} (Config.get @{context} my_flag = true) *} + +example_proof + { + note [[my_flag = false]] + ML_val {* @{assert} (Config.get @{context} my_flag = false) *} + } + ML_val {* @{assert} (Config.get @{context} my_flag = true) *} +qed + + section {* Names \label{sec:names} *} text {* In principle, a name is just a string, but there are various @@ -517,7 +687,7 @@ *} -subsection {* Strings of symbols *} +subsection {* Strings of symbols \label{sec:symbols} *} text {* A \emph{symbol} constitutes the smallest textual unit in Isabelle --- raw ML characters are normally not encountered at all! @@ -549,13 +719,13 @@ \end{enumerate} - \noindent The @{text "ident"} syntax for symbol names is @{text - "letter (letter | digit)\<^sup>*"}, where @{text "letter = - A..Za..z"} and @{text "digit = 0..9"}. There are infinitely many - regular symbols and control symbols, but a fixed collection of - standard symbols is treated specifically. For example, - ``\verb,\,\verb,,'' is classified as a letter, which means it - may occur within regular Isabelle identifiers. + The @{text "ident"} syntax for symbol names is @{text "letter + (letter | digit)\<^sup>*"}, where @{text "letter = A..Za..z"} and @{text + "digit = 0..9"}. There are infinitely many regular symbols and + control symbols, but a fixed collection of standard symbols is + treated specifically. For example, ``\verb,\,\verb,,'' is + classified as a letter, which means it may occur within regular + Isabelle identifiers. The character set underlying Isabelle symbols is 7-bit ASCII, but 8-bit character sequences are passed-through unchanged. Unicode/UCS @@ -594,26 +764,25 @@ \begin{description} - \item @{ML_type "Symbol.symbol"} represents individual Isabelle + \item Type @{ML_type "Symbol.symbol"} represents individual Isabelle symbols. \item @{ML "Symbol.explode"}~@{text "str"} produces a symbol list - from the packed form. This function supercedes @{ML + from the packed form. This function supersedes @{ML "String.explode"} for virtually all purposes of manipulating text in Isabelle!\footnote{The runtime overhead for exploded strings is mainly that of the list structure: individual symbols that happen to - be a singleton string --- which is the most common case --- do not - require extra memory in Poly/ML.} + be a singleton string do not require extra memory in Poly/ML.} \item @{ML "Symbol.is_letter"}, @{ML "Symbol.is_digit"}, @{ML "Symbol.is_quasi"}, @{ML "Symbol.is_blank"} classify standard symbols according to fixed syntactic conventions of Isabelle, cf.\ \cite{isabelle-isar-ref}. - \item @{ML_type "Symbol.sym"} is a concrete datatype that represents - the different kinds of symbols explicitly, with constructors @{ML - "Symbol.Char"}, @{ML "Symbol.Sym"}, @{ML "Symbol.UTF8"}, @{ML - "Symbol.Ctrl"}, @{ML "Symbol.Raw"}. + \item Type @{ML_type "Symbol.sym"} is a concrete datatype that + represents the different kinds of symbols explicitly, with + constructors @{ML "Symbol.Char"}, @{ML "Symbol.Sym"}, @{ML + "Symbol.UTF8"}, @{ML "Symbol.Ctrl"}, @{ML "Symbol.Raw"}. \item @{ML "Symbol.decode"} converts the string representation of a symbol into the datatype version. @@ -692,8 +861,8 @@ \item @{ML Name.skolem}~@{text "name"} produces a Skolem name by adding two underscores. - \item @{ML_type Name.context} represents the context of already used - names; the initial value is @{ML "Name.context"}. + \item Type @{ML_type Name.context} represents the context of already + used names; the initial value is @{ML "Name.context"}. \item @{ML Name.declare}~@{text "name"} enters a used name into the context. @@ -716,6 +885,37 @@ \end{description} *} +text %mlex {* The following simple examples demonstrate how to produce + fresh names from the initial @{ML Name.context}. *} + +ML {* + val list1 = Name.invents Name.context "a" 5; + @{assert} (list1 = ["a", "b", "c", "d", "e"]); + + val list2 = + #1 (Name.variants ["x", "x", "a", "a", "'a", "'a"] Name.context); + @{assert} (list2 = ["x", "xa", "a", "aa", "'a", "'aa"]); +*} + +text {* \medskip The same works reletively to the formal context as + follows. *} + +locale ex = fixes a b c :: 'a +begin + +ML {* + val names = Variable.names_of @{context}; + + val list1 = Name.invents names "a" 5; + @{assert} (list1 = ["d", "e", "f", "g", "h"]); + + val list2 = + #1 (Name.variants ["x", "x", "a", "a", "'a", "'a"] names); + @{assert} (list2 = ["x", "xa", "aa", "ab", "'aa", "'ab"]); +*} + +end + subsection {* Indexed names \label{sec:indexname} *} @@ -756,15 +956,15 @@ text %mlref {* \begin{mldecls} - @{index_ML_type indexname} \\ + @{index_ML_type indexname: "string * int"} \\ \end{mldecls} \begin{description} - \item @{ML_type indexname} represents indexed names. This is an - abbreviation for @{ML_type "string * int"}. The second component is - usually non-negative, except for situations where @{text "(x, -1)"} - is used to inject basic names into this type. Other negative + \item Type @{ML_type indexname} represents indexed names. This is + an abbreviation for @{ML_type "string * int"}. The second component + is usually non-negative, except for situations where @{text "(x, + -1)"} is used to inject basic names into this type. Other negative indexes should not be used. \end{description} @@ -847,13 +1047,12 @@ main equation of this ``chemical reaction'' when binding new entities in a context is as follows: - \smallskip + \medskip \begin{tabular}{l} @{text "binding + naming \ long name + name space accesses"} \end{tabular} - \smallskip - \medskip As a general principle, there is a separate name space for + \bigskip As a general principle, there is a separate name space for each kind of formal entity, e.g.\ fact, logical constant, type constructor, type class. It is usually clear from the occurrence in concrete syntax (or from the scope) which kind of entity a name @@ -866,10 +1065,17 @@ constant @{text "c"}. This technique of mapping names from one space into another requires some care in order to avoid conflicts. In particular, theorem names derived from a type constructor or type - class are better suffixed in addition to the usual qualification, - e.g.\ @{text "c_type.intro"} and @{text "c_class.intro"} for - theorems related to type @{text "c"} and class @{text "c"}, - respectively. + class should get an additional suffix in addition to the usual + qualification. This leads to the following conventions for derived + names: + + \medskip + \begin{tabular}{ll} + logical entity & fact name \\\hline + constant @{text "c"} & @{text "c.intro"} \\ + type @{text "c"} & @{text "c_type.intro"} \\ + class @{text "c"} & @{text "c_class.intro"} \\ + \end{tabular} *} text %mlref {* @@ -901,13 +1107,15 @@ \begin{description} - \item @{ML_type binding} represents the abstract concept of name - bindings. + \item Type @{ML_type binding} represents the abstract concept of + name bindings. \item @{ML Binding.empty} is the empty binding. \item @{ML Binding.name}~@{text "name"} produces a binding with base - name @{text "name"}. + name @{text "name"}. Note that this lacks proper source position + information; see also the ML antiquotation @{ML_antiquotation + binding}. \item @{ML Binding.qualify}~@{text "mandatory name binding"} prefixes qualifier @{text "name"} to @{text "binding"}. The @{text @@ -932,8 +1140,8 @@ representation for human-readable output, together with some formal markup that might get used in GUI front-ends, for example. - \item @{ML_type Name_Space.naming} represents the abstract concept of - a naming policy. + \item Type @{ML_type Name_Space.naming} represents the abstract + concept of a naming policy. \item @{ML Name_Space.default_naming} is the default naming policy. In a theory context, this is usually augmented by a path prefix @@ -946,7 +1154,7 @@ name binding (usually a basic name) into the fully qualified internal name, according to the given naming policy. - \item @{ML_type Name_Space.T} represents name spaces. + \item Type @{ML_type Name_Space.T} represents name spaces. \item @{ML Name_Space.empty}~@{text "kind"} and @{ML Name_Space.merge}~@{text "(space\<^isub>1, space\<^isub>2)"} are the canonical operations for @@ -981,4 +1189,42 @@ \end{description} *} +text %mlantiq {* + \begin{matharray}{rcl} + @{ML_antiquotation_def "binding"} & : & @{text ML_antiquotation} \\ + \end{matharray} + + \begin{rail} + 'binding' name + ; + \end{rail} + + \begin{description} + + \item @{text "@{binding name}"} produces a binding with base name + @{text "name"} and the source position taken from the concrete + syntax of this antiquotation. In many situations this is more + appropriate than the more basic @{ML Binding.name} function. + + \end{description} +*} + +text %mlex {* The following example yields the source position of some + concrete binding inlined into the text. +*} + +ML {* Binding.pos_of @{binding here} *} + +text {* \medskip That position can be also printed in a message as + follows. *} + +ML_command {* + writeln + ("Look here" ^ Position.str_of (Binding.pos_of @{binding here})) +*} + +text {* This illustrates a key virtue of formalized bindings as + opposed to raw specifications of base names: the system can use this + additional information for advanced feedback given to the user. *} + end diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/Proof.thy --- a/doc-src/IsarImplementation/Thy/Proof.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/Proof.thy Mon Oct 25 08:08:08 2010 +0200 @@ -52,13 +52,26 @@ operation of exporting results from a context: a type variable @{text "\"} is considered fixed as long as it occurs in some fixed term variable of the context. For example, exporting @{text "x: - term, \: type \ x\<^isub>\ = x\<^isub>\"} produces in the first step - @{text "x: term \ x\<^isub>\ = x\<^isub>\"} for fixed @{text "\"}, - and only in the second step @{text "\ ?x\<^isub>?\<^isub>\ = - ?x\<^isub>?\<^isub>\"} for schematic @{text "?x"} and @{text "?\"}. + term, \: type \ x\<^isub>\ \ x\<^isub>\"} produces in the first step @{text "x: term + \ x\<^isub>\ \ x\<^isub>\"} for fixed @{text "\"}, and only in the second step + @{text "\ ?x\<^isub>?\<^isub>\ \ ?x\<^isub>?\<^isub>\"} for schematic @{text "?x"} and @{text "?\"}. + The following Isar source text illustrates this scenario. +*} - \medskip The Isabelle/Isar proof context manages the gory details of - term vs.\ type variables, with high-level principles for moving the +example_proof + { + fix x -- {* all potential occurrences of some @{text "x::\"} are fixed *} + { + have "x::'a \ x" -- {* implicit type assigment by concrete occurrence *} + by (rule reflexive) + } + thm this -- {* result still with fixed type @{text "'a"} *} + } + thm this -- {* fully general result for arbitrary @{text "?x::?'a"} *} +qed + +text {* The Isabelle/Isar proof context manages the details of term + vs.\ type variables, with high-level principles for moving the frontier between fixed and schematic variables. The @{text "add_fixes"} operation explictly declares fixed @@ -145,12 +158,8 @@ \end{description} *} -text %mlex {* The following example (in theory @{theory Pure}) shows - how to work with fixed term and type parameters work with - type-inference. -*} - -typedecl foo -- {* some basic type for testing purposes *} +text %mlex {* The following example shows how to work with fixed term + and type parameters and with type-inference. *} ML {* (*static compile-time context -- for testing only*) @@ -164,11 +173,11 @@ val t1' = singleton (Variable.polymorphic ctxt1) t1; (*term u enforces specific type assignment*) - val u = Syntax.read_term ctxt1 "(x::foo) \ y"; + val u = Syntax.read_term ctxt1 "(x::nat) \ y"; (*official declaration of u -- propagates constraints etc.*) val ctxt2 = ctxt1 |> Variable.declare_term u; - val t2 = Syntax.read_term ctxt2 "x"; (*x::foo is enforced*) + val t2 = Syntax.read_term ctxt2 "x"; (*x::nat is enforced*) *} text {* In the above example, the starting context had been derived @@ -185,15 +194,13 @@ ctxt0 |> Variable.variant_fixes ["x", "x", "x"]; *} -text {* \noindent Subsequent ML code can now work with the invented - names of @{verbatim x1}, @{verbatim x2}, @{verbatim x3}, without - depending on the details on the system policy for introducing these - variants. Recall that within a proof body the system always invents - fresh ``skolem constants'', e.g.\ as follows: -*} +text {* The following ML code can now work with the invented names of + @{verbatim x1}, @{verbatim x2}, @{verbatim x3}, without depending on + the details on the system policy for introducing these variants. + Recall that within a proof body the system always invents fresh + ``skolem constants'', e.g.\ as follows: *} -lemma "PROP XXX" -proof - +example_proof ML_prf %"ML" {* val ctxt0 = @{context}; @@ -206,10 +213,10 @@ *} oops -text {* \noindent In this situation @{ML Variable.add_fixes} and @{ML +text {* In this situation @{ML Variable.add_fixes} and @{ML Variable.variant_fixes} are very similar, but identical name proposals given in a row are only accepted by the second version. -*} + *} section {* Assumptions \label{sec:assumptions} *} @@ -282,11 +289,11 @@ \begin{description} - \item @{ML_type Assumption.export} represents arbitrary export - rules, which is any function of type @{ML_type "bool -> cterm list -> thm -> thm"}, - where the @{ML_type "bool"} indicates goal mode, and the @{ML_type - "cterm list"} the collection of assumptions to be discharged - simultaneously. + \item Type @{ML_type Assumption.export} represents arbitrary export + rules, which is any function of type @{ML_type "bool -> cterm list + -> thm -> thm"}, where the @{ML_type "bool"} indicates goal mode, + and the @{ML_type "cterm list"} the collection of assumptions to be + discharged simultaneously. \item @{ML Assumption.assume}~@{text "A"} turns proposition @{text "A"} into a primitive assumption @{text "A \ A'"}, where the @@ -330,11 +337,10 @@ val r = Assumption.export false ctxt1 ctxt0 eq'; *} -text {* \noindent Note that the variables of the resulting rule are - not generalized. This would have required to fix them properly in - the context beforehand, and export wrt.\ variables afterwards (cf.\ - @{ML Variable.export} or the combined @{ML "ProofContext.export"}). -*} +text {* Note that the variables of the resulting rule are not + generalized. This would have required to fix them properly in the + context beforehand, and export wrt.\ variables afterwards (cf.\ @{ML + Variable.export} or the combined @{ML "ProofContext.export"}). *} section {* Structured goals and results \label{sec:struct-goals} *} @@ -373,18 +379,24 @@ conclusion: the proof demonstrates that the context may be augmented by parameters and assumptions, without affecting any conclusions that do not mention these parameters. See also - \cite{isabelle-isar-ref} for the user-level @{text "\"} and - @{text "\"} elements. Final results, which may not refer to + \cite{isabelle-isar-ref} for the user-level @{command obtain} and + @{command guess} elements. Final results, which may not refer to the parameters in the conclusion, need to exported explicitly into - the original context. -*} + the original context. *} text %mlref {* \begin{mldecls} - @{index_ML SUBPROOF: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\ - @{index_ML Subgoal.FOCUS: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\ - @{index_ML Subgoal.FOCUS_PREMS: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\ - @{index_ML Subgoal.FOCUS_PARAMS: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\ + @{index_ML SUBPROOF: "(Subgoal.focus -> tactic) -> + Proof.context -> int -> tactic"} \\ + @{index_ML Subgoal.FOCUS: "(Subgoal.focus -> tactic) -> + Proof.context -> int -> tactic"} \\ + @{index_ML Subgoal.FOCUS_PREMS: "(Subgoal.focus -> tactic) -> + Proof.context -> int -> tactic"} \\ + @{index_ML Subgoal.FOCUS_PARAMS: "(Subgoal.focus -> tactic) -> + Proof.context -> int -> tactic"} \\ + @{index_ML Subgoal.focus: "Proof.context -> int -> thm -> Subgoal.focus * thm"} \\ + @{index_ML Subgoal.focus_prems: "Proof.context -> int -> thm -> Subgoal.focus * thm"} \\ + @{index_ML Subgoal.focus_params: "Proof.context -> int -> thm -> Subgoal.focus * thm"} \\ \end{mldecls} \begin{mldecls} @@ -394,8 +406,8 @@ ({prems: thm list, context: Proof.context} -> tactic) -> thm list"} \\ \end{mldecls} \begin{mldecls} - @{index_ML Obtain.result: "(Proof.context -> tactic) -> - thm list -> Proof.context -> ((string * cterm) list * thm list) * Proof.context"} \\ + @{index_ML Obtain.result: "(Proof.context -> tactic) -> thm list -> + Proof.context -> ((string * cterm) list * thm list) * Proof.context"} \\ \end{mldecls} \begin{description} @@ -412,6 +424,12 @@ imported into the context, and the body tactic may introduce new subgoals and schematic variables. + \item @{ML Subgoal.focus}, @{ML Subgoal.focus_prems}, @{ML + Subgoal.focus_params} extract the focus information from a goal + state in the same way as the corresponding tacticals above. This is + occasionally useful to experiment without writing actual tactics + yet. + \item @{ML Goal.prove}~@{text "ctxt xs As C tac"} states goal @{text "C"} in the context augmented by fixed variables @{text "xs"} and assumptions @{text "As"}, and applies tactic @{text "tac"} to solve @@ -431,4 +449,41 @@ \end{description} *} +text %mlex {* The following minimal example illustrates how to access + the focus information of a structured goal state. *} + +example_proof + fix A B C :: "'a \ bool" + + have "\x. A x \ B x \ C x" + ML_val + {* + val {goal, context = goal_ctxt, ...} = @{Isar.goal}; + val (focus as {params, asms, concl, ...}, goal') = + Subgoal.focus goal_ctxt 1 goal; + val [A, B] = #prems focus; + val [(_, x)] = #params focus; + *} + oops + +text {* \medskip The next example demonstrates forward-elimination in + a local context, using @{ML Obtain.result}. *} + +example_proof + assume ex: "\x. B x" + + ML_prf %"ML" {* + val ctxt0 = @{context}; + val (([(_, x)], [B]), ctxt1) = ctxt0 + |> Obtain.result (fn _ => etac @{thm exE} 1) [@{thm ex}]; + *} + ML_prf %"ML" {* + singleton (ProofContext.export ctxt1 ctxt0) @{thm refl}; + *} + ML_prf %"ML" {* + ProofContext.export ctxt1 ctxt0 [Thm.reflexive x] + handle ERROR msg => (warning msg; []); + *} +qed + end diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/ROOT.ML diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/Syntax.thy --- a/doc-src/IsarImplementation/Thy/Syntax.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/Syntax.thy Mon Oct 25 08:08:08 2010 +0200 @@ -6,10 +6,91 @@ text FIXME -section {* Parsing and printing *} +section {* Reading and pretty printing \label{sec:read-print} *} + +text FIXME + +text %mlref {* + \begin{mldecls} + @{index_ML Syntax.read_typ: "Proof.context -> string -> typ"} \\ + @{index_ML Syntax.read_term: "Proof.context -> string -> term"} \\ + @{index_ML Syntax.read_prop: "Proof.context -> string -> term"} \\ + @{index_ML Syntax.pretty_typ: "Proof.context -> typ -> Pretty.T"} \\ + @{index_ML Syntax.pretty_term: "Proof.context -> term -> Pretty.T"} \\ + \end{mldecls} + + \begin{description} + + \item FIXME + + \end{description} +*} + + +section {* Parsing and unparsing \label{sec:parse-unparse} *} text FIXME +text %mlref {* + \begin{mldecls} + @{index_ML Syntax.parse_typ: "Proof.context -> string -> typ"} \\ + @{index_ML Syntax.parse_term: "Proof.context -> string -> term"} \\ + @{index_ML Syntax.parse_prop: "Proof.context -> string -> term"} \\ + @{index_ML Syntax.unparse_typ: "Proof.context -> typ -> Pretty.T"} \\ + @{index_ML Syntax.unparse_term: "Proof.context -> term -> Pretty.T"} \\ + \end{mldecls} + + \begin{description} + + \item FIXME + + \end{description} +*} + + section {* Checking and unchecking \label{sec:term-check} *} +text FIXME + +text %mlref {* + \begin{mldecls} + @{index_ML Syntax.check_typs: "Proof.context -> typ list -> typ list"} \\ + @{index_ML Syntax.check_terms: "Proof.context -> term list -> term list"} \\ + @{index_ML Syntax.check_props: "Proof.context -> term list -> term list"} \\ + @{index_ML Syntax.uncheck_typs: "Proof.context -> typ list -> typ list"} \\ + @{index_ML Syntax.uncheck_terms: "Proof.context -> term list -> term list"} \\ + \end{mldecls} + + \begin{description} + + \item FIXME + + \end{description} +*} + + +section {* Syntax translations *} + +text FIXME + +text %mlantiq {* + \begin{matharray}{rcl} + @{ML_antiquotation_def "class_syntax"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "type_syntax"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "const_syntax"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "syntax_const"} & : & @{text ML_antiquotation} \\ + \end{matharray} + + \begin{rail} + ('class_syntax' | 'type_syntax' | 'const_syntax' | 'syntax_const') name + ; + \end{rail} + + \begin{description} + + \item FIXME + + \end{description} +*} + end diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/Tactic.thy --- a/doc-src/IsarImplementation/Thy/Tactic.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/Tactic.thy Mon Oct 25 08:08:08 2010 +0200 @@ -86,7 +86,7 @@ *} -section {* Tactics *} +section {* Tactics\label{sec:tactics} *} text {* A @{text "tactic"} is a function @{text "goal \ goal\<^sup>*\<^sup>*"} that maps a given goal state (represented as a theorem, cf.\ @@ -180,14 +180,14 @@ \begin{description} - \item @{ML_type tactic} represents tactics. The well-formedness - conditions described above need to be observed. See also @{"file" - "~~/src/Pure/General/seq.ML"} for the underlying implementation of - lazy sequences. + \item Type @{ML_type tactic} represents tactics. The + well-formedness conditions described above need to be observed. See + also @{"file" "~~/src/Pure/General/seq.ML"} for the underlying + implementation of lazy sequences. - \item @{ML_type "int -> tactic"} represents tactics with explicit - subgoal addressing, with well-formedness conditions as described - above. + \item Type @{ML_type "int -> tactic"} represents tactics with + explicit subgoal addressing, with well-formedness conditions as + described above. \item @{ML no_tac} is a tactic that always fails, returning the empty sequence. @@ -406,6 +406,8 @@ Various search strategies may be expressed via tacticals. \medskip FIXME -*} + + \medskip The chapter on tacticals in \cite{isabelle-ref} is still + applicable, despite a few outdated details. *} end diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/document/Base.tex --- a/doc-src/IsarImplementation/Thy/document/Base.tex Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/document/Base.tex Mon Oct 25 08:08:08 2010 +0200 @@ -9,10 +9,43 @@ \isatagtheory \isacommand{theory}\isamarkupfalse% \ Base\isanewline -\isakeyword{imports}\ Pure\isanewline +\isakeyword{imports}\ Main\isanewline \isakeyword{uses}\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isacharslash}{\isachardot}{\isachardot}{\isacharslash}antiquote{\isacharunderscore}setup{\isachardot}ML{\isachardoublequoteclose}\isanewline -\isakeyword{begin}\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +\isanewline +% +\endisadelimtheory +\isanewline +% +\isadelimML \isanewline +% +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ ML{\isacharunderscore}Antiquote{\isachardot}inline\ {\isachardoublequote}assert{\isachardoublequote}\isanewline +\ \ \ \ {\isacharparenleft}Scan{\isachardot}succeed\ {\isachardoublequote}{\isacharparenleft}fn\ b\ {\isacharequal}{\isachargreater}\ if\ b\ then\ {\isacharparenleft}{\isacharparenright}\ else\ raise\ General{\isachardot}Fail\ {\isacharbackslash}{\isachardoublequote}Assertion\ failed{\isacharbackslash}{\isachardoublequote}{\isacharparenright}{\isachardoublequote}{\isacharparenright}\isanewline +{\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +\isanewline +% +\endisadelimML +% +\isadelimtheory +\isanewline +% +\endisadelimtheory +% +\isatagtheory \isacommand{end}\isamarkupfalse% % \endisatagtheory diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/document/Integration.tex --- a/doc-src/IsarImplementation/Thy/document/Integration.tex Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/document/Integration.tex Mon Oct 25 08:08:08 2010 +0200 @@ -29,13 +29,12 @@ \begin{isamarkuptext}% The Isar toplevel may be considered the centeral hub of the Isabelle/Isar system, where all key components and sub-systems are - integrated into a single read-eval-print loop of Isar commands. We - shall even incorporate the existing {\ML} toplevel of the compiler - and run-time system (cf.\ \secref{sec:ML-toplevel}). + integrated into a single read-eval-print loop of Isar commands, + which also incorporates the underlying ML compiler. Isabelle/Isar departs from the original ``LCF system architecture'' - where {\ML} was really The Meta Language for defining theories and - conducting proofs. Instead, {\ML} now only serves as the + where ML was really The Meta Language for defining theories and + conducting proofs. Instead, ML now only serves as the implementation language for the system (and user extensions), while the specific Isar toplevel supports the concepts of theory and proof development natively. This includes the graph structure of theories @@ -93,11 +92,11 @@ \begin{description} - \item \verb|Toplevel.state| represents Isar toplevel states, - which are normally manipulated through the concept of toplevel - transitions only (\secref{sec:toplevel-transition}). Also note that - a raw toplevel state is subject to the same linearity restrictions - as a theory context (cf.~\secref{sec:context-theory}). + \item Type \verb|Toplevel.state| represents Isar toplevel + states, which are normally manipulated through the concept of + toplevel transitions only (\secref{sec:toplevel-transition}). Also + note that a raw toplevel state is subject to the same linearity + restrictions as a theory context (cf.~\secref{sec:context-theory}). \item \verb|Toplevel.UNDEF| is raised for undefined toplevel operations. Many operations work only partially for certain cases, @@ -121,7 +120,7 @@ information for each Isar command being executed. \item \verb|Toplevel.profiling|~\verb|:=|~\isa{n} controls - low-level profiling of the underlying {\ML} runtime system. For + low-level profiling of the underlying ML runtime system. For Poly/ML, \isa{n\ {\isacharequal}\ {\isadigit{1}}} means time and \isa{n\ {\isacharequal}\ {\isadigit{2}}} space profiling. @@ -136,6 +135,35 @@ % \endisadelimmlref % +\isadelimmlantiq +% +\endisadelimmlantiq +% +\isatagmlantiq +% +\begin{isamarkuptext}% +\begin{matharray}{rcl} + \indexdef{}{ML antiquotation}{Isar.state}\hypertarget{ML antiquotation.Isar.state}{\hyperlink{ML antiquotation.Isar.state}{\mbox{\isa{Isar{\isachardot}state}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \end{matharray} + + \begin{description} + + \item \isa{{\isacharat}{\isacharbraceleft}Isar{\isachardot}state{\isacharbraceright}} refers to Isar toplevel state at that + point --- as abstract value. + + This only works for diagnostic ML commands, such as \hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isacharunderscore}val}}}} or \hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isacharunderscore}command}}}}. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlantiq +{\isafoldmlantiq}% +% +\isadelimmlantiq +% +\endisadelimmlantiq +% \isamarkupsubsection{Toplevel transitions \label{sec:toplevel-transition}% } \isamarkuptrue% @@ -232,141 +260,6 @@ % \endisadelimmlref % -\isamarkupsubsection{Toplevel control% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -There are a few special control commands that modify the behavior - the toplevel itself, and only make sense in interactive mode. Under - normal circumstances, the user encounters these only implicitly as - part of the protocol between the Isabelle/Isar system and a - user-interface such as Proof~General. - - \begin{description} - - \item \isacommand{undo} follows the three-level hierarchy of empty - toplevel vs.\ theory vs.\ proof: undo within a proof reverts to the - previous proof context, undo after a proof reverts to the theory - before the initial goal statement, undo of a theory command reverts - to the previous theory value, undo of a theory header discontinues - the current theory development and removes it from the theory - database (\secref{sec:theory-database}). - - \item \isacommand{kill} aborts the current level of development: - kill in a proof context reverts to the theory before the initial - goal statement, kill in a theory context aborts the current theory - development, removing it from the database. - - \item \isacommand{exit} drops out of the Isar toplevel into the - underlying {\ML} toplevel (\secref{sec:ML-toplevel}). The Isar - toplevel state is preserved and may be continued later. - - \item \isacommand{quit} terminates the Isabelle/Isar process without - saving. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{ML toplevel \label{sec:ML-toplevel}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The {\ML} toplevel provides a read-compile-eval-print loop for {\ML} - values, types, structures, and functors. {\ML} declarations operate - on the global system state, which consists of the compiler - environment plus the values of {\ML} reference variables. There is - no clean way to undo {\ML} declarations, except for reverting to a - previously saved state of the whole Isabelle process. {\ML} input - is either read interactively from a TTY, or from a string (usually - within a theory text), or from a source file (usually loaded from a - theory). - - Whenever the {\ML} toplevel is active, the current Isabelle theory - context is passed as an internal reference variable. Thus {\ML} - code may access the theory context during compilation, it may even - change the value of a theory being under construction --- while - observing the usual linearity restrictions - (cf.~\secref{sec:context-theory}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexdef{}{ML}{ML\_Context.the\_generic\_context}\verb|ML_Context.the_generic_context: unit -> Context.generic| \\ - \indexdef{}{ML}{Context.$>$$>$ }\verb|Context.>> : (Context.generic -> Context.generic) -> unit| \\ - \end{mldecls} - - \begin{description} - - \item \verb|ML_Context.the_generic_context ()| refers to the theory - context of the {\ML} toplevel --- at compile time! {\ML} code needs - to take care to refer to \verb|ML_Context.the_generic_context ()| - correctly. Recall that evaluation of a function body is delayed - until actual runtime. Moreover, persistent {\ML} toplevel bindings - to an unfinished theory should be avoided: code should either - project out the desired information immediately, or produce an - explicit \verb|theory_ref| (cf.\ \secref{sec:context-theory}). - - \item \verb|Context.>>|~\isa{f} applies context transformation - \isa{f} to the implicit context of the {\ML} toplevel. - - \end{description} - - It is very important to note that the above functions are really - restricted to the compile time, even though the {\ML} compiler is - invoked at runtime! The majority of {\ML} code uses explicit - functional arguments of a theory or proof context instead. Thus it - may be invoked for an arbitrary context later on, without having to - worry about any operational details. - - \bigskip - - \begin{mldecls} - \indexdef{}{ML}{Isar.main}\verb|Isar.main: unit -> unit| \\ - \indexdef{}{ML}{Isar.loop}\verb|Isar.loop: unit -> unit| \\ - \indexdef{}{ML}{Isar.state}\verb|Isar.state: unit -> Toplevel.state| \\ - \indexdef{}{ML}{Isar.exn}\verb|Isar.exn: unit -> (exn * string) option| \\ - \indexdef{}{ML}{Isar.goal}\verb|Isar.goal: unit ->|\isasep\isanewline% -\verb| {context: Proof.context, facts: thm list, goal: thm}| \\ - \end{mldecls} - - \begin{description} - - \item \verb|Isar.main ()| invokes the Isar toplevel from {\ML}, - initializing an empty toplevel state. - - \item \verb|Isar.loop ()| continues the Isar toplevel with the - current state, after having dropped out of the Isar toplevel loop. - - \item \verb|Isar.state ()| and \verb|Isar.exn ()| get current - toplevel state and error condition, respectively. This only works - after having dropped out of the Isar toplevel loop. - - \item \verb|Isar.goal ()| produces the full Isar goal state, - consisting of proof context, facts that have been indicated for - immediate use, and the tactical goal according to - \secref{sec:tactical-goals}. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% \isamarkupsection{Theory database \label{sec:theory-database}% } \isamarkuptrue% @@ -384,10 +277,10 @@ name space is flat! Theory \isa{A} is associated with the main theory file \isa{A}\verb,.thy,, which needs to be accessible through the theory - loader path. Any number of additional {\ML} source files may be + loader path. Any number of additional ML source files may be associated with each theory, by declaring these dependencies in the theory header as \isa{{\isasymUSES}}, and loading them consecutively - within the theory context. The system keeps track of incoming {\ML} + within the theory context. The system keeps track of incoming ML sources and associates them with the current theory. The basic internal actions of the theory database are \isa{update} and \isa{remove}: diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/document/Isar.tex --- a/doc-src/IsarImplementation/Thy/document/Isar.tex Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/document/Isar.tex Mon Oct 25 08:08:08 2010 +0200 @@ -23,16 +23,33 @@ \isamarkuptrue% % \begin{isamarkuptext}% -The primary Isar language consists of three main categories of - language elements: +The Isar proof language (see also \cite[\S2]{isabelle-isar-ref}) + consists of three main categories of language elements: \begin{enumerate} - \item Proof commands + \item Proof \emph{commands} define the primary language of + transactions of the underlying Isar/VM interpreter. Typical + examples are \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}, \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}, \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}, \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}, and \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}. + + Composing proof commands according to the rules of the Isar/VM leads + to expressions of structured proof text, such that both the machine + and the human reader can give it a meaning as formal reasoning. - \item Proof methods + \item Proof \emph{methods} define a secondary language of mixed + forward-backward refinement steps involving facts and goals. + Typical examples are \hyperlink{method.rule}{\mbox{\isa{rule}}}, \hyperlink{method.unfold}{\mbox{\isa{unfold}}}, and \hyperlink{method.simp}{\mbox{\isa{simp}}}. - \item Attributes + Methods can occur in certain well-defined parts of the Isar proof + language, say as arguments to \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}, \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}, + or \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}. + + \item \emph{Attributes} define a tertiary language of small + annotations to theorems being defined or referenced. Attributes can + modify both the context and the theorem. + + Typical examples are \hyperlink{attribute.intro}{\mbox{\isa{intro}}} (which affects the context), + and \hyperlink{attribute.symmetric}{\mbox{\isa{symmetric}}} (which affects the theorem). \end{enumerate}% \end{isamarkuptext}% @@ -43,7 +60,227 @@ \isamarkuptrue% % \begin{isamarkuptext}% -FIXME% +A \emph{proof command} is state transition of the Isar/VM + proof interpreter. + + In principle, Isar proof commands could be defined in + user-space as well. The system is built like that in the first + place: part of the commands are primitive, the other part is defined + as derived elements. Adding to the genuine structured proof + language requires profound understanding of the Isar/VM machinery, + though, so this is beyond the scope of this manual. + + What can be done realistically is to define some diagnostic commands + that inspect the general state of the Isar/VM, and report some + feedback to the user. Typically this involves checking of the + linguistic \emph{mode} of a proof state, or peeking at the pending + goals (if available). + + Another common application is to define a toplevel command that + poses a problem to the user as Isar proof state and stores the final + result in a suitable context data slot. Thus a proof can be + incorporated into the context of some user-space tool, without + modifying the Isar proof language itself.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML type}{Proof.state}\verb|type Proof.state| \\ + \indexdef{}{ML}{Proof.assert\_forward}\verb|Proof.assert_forward: Proof.state -> Proof.state| \\ + \indexdef{}{ML}{Proof.assert\_chain}\verb|Proof.assert_chain: Proof.state -> Proof.state| \\ + \indexdef{}{ML}{Proof.assert\_backward}\verb|Proof.assert_backward: Proof.state -> Proof.state| \\ + \indexdef{}{ML}{Proof.simple\_goal}\verb|Proof.simple_goal: Proof.state -> {context: Proof.context, goal: thm}| \\ + \indexdef{}{ML}{Proof.goal}\verb|Proof.goal: Proof.state ->|\isasep\isanewline% +\verb| {context: Proof.context, facts: thm list, goal: thm}| \\ + \indexdef{}{ML}{Proof.raw\_goal}\verb|Proof.raw_goal: Proof.state ->|\isasep\isanewline% +\verb| {context: Proof.context, facts: thm list, goal: thm}| \\ + \indexdef{}{ML}{Proof.theorem}\verb|Proof.theorem: Method.text option ->|\isasep\isanewline% +\verb| (thm list list -> Proof.context -> Proof.context) ->|\isasep\isanewline% +\verb| (term * term list) list list -> Proof.context -> Proof.state| \\ + \end{mldecls} + + \begin{description} + + \item Type \verb|Proof.state| represents Isar proof states. + This is a block-structured configuration with proof context, + linguistic mode, and optional goal. The latter consists of goal + context, goal facts (``\isa{using}''), and tactical goal state + (see \secref{sec:tactical-goals}). + + The general idea is that the facts shall contribute to the + refinement of some parts of the tactical goal --- how exactly is + defined by the proof method that is applied in that situation. + + \item \verb|Proof.assert_forward|, \verb|Proof.assert_chain|, \verb|Proof.assert_backward| are partial identity functions that fail + unless a certain linguistic mode is active, namely ``\isa{proof{\isacharparenleft}state{\isacharparenright}}'', ``\isa{proof{\isacharparenleft}chain{\isacharparenright}}'', ``\isa{proof{\isacharparenleft}prove{\isacharparenright}}'', respectively (using the terminology of + \cite{isabelle-isar-ref}). + + It is advisable study the implementations of existing proof commands + for suitable modes to be asserted. + + \item \verb|Proof.simple_goal|~\isa{state} returns the structured + Isar goal (if available) in the form seen by ``simple'' methods + (like \hyperlink{method.simp}{\mbox{\isa{simp}}} or \hyperlink{method.blast}{\mbox{\isa{blast}}}). The Isar goal facts are + already inserted as premises into the subgoals, which are presented + individually as in \verb|Proof.goal|. + + \item \verb|Proof.goal|~\isa{state} returns the structured Isar + goal (if available) in the form seen by regular methods (like + \hyperlink{method.rule}{\mbox{\isa{rule}}}). The auxiliary internal encoding of Pure + conjunctions is split into individual subgoals as usual. + + \item \verb|Proof.raw_goal|~\isa{state} returns the structured + Isar goal (if available) in the raw internal form seen by ``raw'' + methods (like \hyperlink{method.induct}{\mbox{\isa{induct}}}). This form is rarely appropriate + for dignostic tools; \verb|Proof.simple_goal| or \verb|Proof.goal| + should be used in most situations. + + \item \verb|Proof.theorem|~\isa{before{\isacharunderscore}qed\ after{\isacharunderscore}qed\ statement\ ctxt} + initializes a toplevel Isar proof state within a given context. + + The optional \isa{before{\isacharunderscore}qed} method is applied at the end of + the proof, just before extracting the result (this feature is rarely + used). + + The \isa{after{\isacharunderscore}qed} continuation receives the extracted result + in order to apply it to the final context in a suitable way (e.g.\ + storing named facts). Note that at this generic level the target + context is specified as \verb|Proof.context|, but the usual + wrapping of toplevel proofs into command transactions will provide a + \verb|local_theory| here (see also \chref{ch:local-theory}). + This usually affects the way how results are stored. + + The \isa{statement} is given as a nested list of terms, each + associated with optional \hyperlink{keyword.is}{\mbox{\isa{\isakeyword{is}}}} patterns as usual in the + Isar source language. The original list structure over terms is + turned into one over theorems when \isa{after{\isacharunderscore}qed} is invoked. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isadelimmlantiq +% +\endisadelimmlantiq +% +\isatagmlantiq +% +\begin{isamarkuptext}% +\begin{matharray}{rcl} + \indexdef{}{ML antiquotation}{Isar.goal}\hypertarget{ML antiquotation.Isar.goal}{\hyperlink{ML antiquotation.Isar.goal}{\mbox{\isa{Isar{\isachardot}goal}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \end{matharray} + + \begin{description} + + \item \isa{{\isacharat}{\isacharbraceleft}Isar{\isachardot}goal{\isacharbraceright}} refers to the regular goal state (if + available) of the current proof state managed by the Isar toplevel + --- as abstract value. + + This only works for diagnostic ML commands, such as \hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isacharunderscore}val}}}} or \hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isacharunderscore}command}}}}. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlantiq +{\isafoldmlantiq}% +% +\isadelimmlantiq +% +\endisadelimmlantiq +% +\isadelimmlex +% +\endisadelimmlex +% +\isatagmlex +% +\begin{isamarkuptext}% +The following example peeks at a certain goal configuration.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlex +{\isafoldmlex}% +% +\isadelimmlex +% +\endisadelimmlex +\isacommand{example{\isacharunderscore}proof}\isamarkupfalse% +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{have}\isamarkupfalse% +\ A\ \isakeyword{and}\ B\ \isakeyword{and}\ C% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +% +\isadelimML +\ \ \ \ % +\endisadelimML +% +\isatagML +\isacommand{ML{\isacharunderscore}val}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ \ \ \ \ val\ n\ {\isacharequal}\ Thm{\isachardot}nprems{\isacharunderscore}of\ {\isacharparenleft}{\isacharhash}goal\ % +\isaantiq +Isar{\isachardot}goal% +\endisaantiq +{\isacharparenright}{\isacharsemicolon}\isanewline +\ \ \ \ \ \ % +\isaantiq +assert% +\endisaantiq +\ {\isacharparenleft}n\ {\isacharequal}\ {\isadigit{3}}{\isacharparenright}{\isacharsemicolon}\isanewline +\ \ \ \ {\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +\isanewline +% +\endisadelimML +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{oops}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Here we see 3 individual subgoals in the same way as regular + proof methods would do.% \end{isamarkuptext}% \isamarkuptrue% % @@ -52,19 +289,627 @@ \isamarkuptrue% % \begin{isamarkuptext}% -FIXME% +A \isa{method} is a function \isa{context\ {\isasymrightarrow}\ thm\isactrlsup {\isacharasterisk}\ {\isasymrightarrow}\ goal\ {\isasymrightarrow}\ {\isacharparenleft}cases\ {\isasymtimes}\ goal{\isacharparenright}\isactrlsup {\isacharasterisk}\isactrlsup {\isacharasterisk}} that operates on the full Isar goal + configuration with context, goal facts, and tactical goal state and + enumerates possible follow-up goal states, with the potential + addition of named extensions of the proof context (\emph{cases}). + The latter feature is rarely used. + + This means a proof method is like a structurally enhanced tactic + (cf.\ \secref{sec:tactics}). The well-formedness conditions for + tactics need to hold for methods accordingly, with the following + additions. + + \begin{itemize} + + \item Goal addressing is further limited either to operate either + uniformly on \emph{all} subgoals, or specifically on the + \emph{first} subgoal. + + Exception: old-style tactic emulations that are embedded into the + method space, e.g.\ \hyperlink{method.rule-tac}{\mbox{\isa{rule{\isacharunderscore}tac}}}. + + \item A non-trivial method always needs to make progress: an + identical follow-up goal state has to be avoided.\footnote{This + enables the user to write method expressions like \isa{meth\isactrlsup {\isacharplus}} + without looping, while the trivial do-nothing case can be recovered + via \isa{meth\isactrlsup {\isacharquery}}.} + + Exception: trivial stuttering steps, such as ``\hyperlink{method.-}{\mbox{\isa{{\isacharminus}}}}'' or + \hyperlink{method.succeed}{\mbox{\isa{succeed}}}. + + \item Goal facts passed to the method must not be ignored. If there + is no sensible use of facts outside the goal state, facts should be + inserted into the subgoals that are addressed by the method. + + \end{itemize} + + \medskip Syntactically, the language of proof methods is embedded + into Isar as arguments to certain commands, e.g.\ \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}} or + \hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}. User-space additions are reasonably easy by + plugging suitable method-valued parser functions into the framework, + using the \hyperlink{command.method-setup}{\mbox{\isa{\isacommand{method{\isacharunderscore}setup}}}} command, for example. + + To get a better idea about the range of possibilities, consider the + following Isar proof schemes. First some general form of structured + proof text: + + \medskip + \begin{tabular}{l} + \hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{facts\isactrlsub {\isadigit{1}}}~\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{props}~\hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{facts\isactrlsub {\isadigit{2}}} \\ + \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isacharparenleft}initial{\isacharunderscore}method{\isacharparenright}} \\ + \quad\isa{body} \\ + \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}~\isa{{\isacharparenleft}terminal{\isacharunderscore}method{\isacharparenright}} \\ + \end{tabular} + \medskip + + The goal configuration consists of \isa{facts\isactrlsub {\isadigit{1}}} and \isa{facts\isactrlsub {\isadigit{2}}} appended in that order, and various \isa{props} being + claimed here. The \isa{initial{\isacharunderscore}method} is invoked with facts + and goals together and refines the problem to something that is + handled recursively in the proof \isa{body}. The \isa{terminal{\isacharunderscore}method} has another chance to finish-off any remaining + subgoals, but it does not see the facts of the initial step. + + \medskip The second pattern illustrates unstructured proof scripts: + + \medskip + \begin{tabular}{l} + \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{props} \\ + \quad\hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{facts\isactrlsub {\isadigit{1}}}~\hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}~\isa{method\isactrlsub {\isadigit{1}}} \\ + \quad\hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}~\isa{method\isactrlsub {\isadigit{2}}} \\ + \quad\hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{facts\isactrlsub {\isadigit{3}}}~\hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}~\isa{method\isactrlsub {\isadigit{3}}} \\ + \quad\hyperlink{command.done}{\mbox{\isa{\isacommand{done}}}} \\ + \end{tabular} + \medskip + + The \isa{method\isactrlsub {\isadigit{1}}} operates on the original claim together while + using \isa{facts\isactrlsub {\isadigit{1}}}. Since the \hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}} command + structurally resets the facts, the \isa{method\isactrlsub {\isadigit{2}}} will operate on + the remaining goal state without facts. The \isa{method\isactrlsub {\isadigit{3}}} will + see again a collection of \isa{facts\isactrlsub {\isadigit{3}}} that has been inserted + into the script explicitly. + + \medskip Empirically, Isar proof methods can be categorized as + follows. + + \begin{enumerate} + + \item \emph{Special method with cases} with named context additions + associated with the follow-up goal state. + + Example: \hyperlink{method.induct}{\mbox{\isa{induct}}}, which is also a ``raw'' method since it + operates on the internal representation of simultaneous claims as + Pure conjunction, instead of separate subgoals. + + \item \emph{Structured method} with strong emphasis on facts outside + the goal state. + + Example: \hyperlink{method.rule}{\mbox{\isa{rule}}}, which captures the key ideas behind + structured reasoning in Isar in purest form. + + \item \emph{Simple method} with weaker emphasis on facts, which are + inserted into subgoals to emulate old-style tactical as + ``premises''. + + Examples: \hyperlink{method.simp}{\mbox{\isa{simp}}}, \hyperlink{method.blast}{\mbox{\isa{blast}}}, \hyperlink{method.auto}{\mbox{\isa{auto}}}. + + \item \emph{Old-style tactic emulation} with detailed numeric goal + addressing and explicit references to entities of the internal goal + state (which are otherwise invisible from proper Isar proof text). + To make the special non-standard status clear, the naming convention + \isa{foo{\isacharunderscore}tac} needs to be observed. + + Example: \hyperlink{method.rule-tac}{\mbox{\isa{rule{\isacharunderscore}tac}}}. + + \end{enumerate} + + When implementing proof methods, it is advisable to study existing + implementations carefully and imitate the typical ``boiler plate'' + for context-sensitive parsing and further combinators to wrap-up + tactic expressions as methods.\footnote{Aliases or abbreviations of + the standard method combinators should be avoided. Note that from + Isabelle99 until Isabelle2009 the system did provide various odd + combinations of method wrappers that made user applications more + complicated than necessary.}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML type}{Proof.method}\verb|type Proof.method| \\ + \indexdef{}{ML}{METHOD\_CASES}\verb|METHOD_CASES: (thm list -> cases_tactic) -> Proof.method| \\ + \indexdef{}{ML}{METHOD}\verb|METHOD: (thm list -> tactic) -> Proof.method| \\ + \indexdef{}{ML}{SIMPLE\_METHOD}\verb|SIMPLE_METHOD: tactic -> Proof.method| \\ + \indexdef{}{ML}{SIMPLE\_METHOD'}\verb|SIMPLE_METHOD': (int -> tactic) -> Proof.method| \\ + \indexdef{}{ML}{HEADGOAL}\verb|HEADGOAL: (int -> tactic) -> tactic| \\ + \indexdef{}{ML}{Method.insert\_tac}\verb|Method.insert_tac: thm list -> int -> tactic| \\ + \indexdef{}{ML}{Method.setup}\verb|Method.setup: binding -> (Proof.context -> Proof.method) context_parser ->|\isasep\isanewline% +\verb| string -> theory -> theory| \\ + \end{mldecls} + + \begin{description} + + \item Type \verb|Proof.method| represents proof methods as + abstract type. + + \item \verb|METHOD_CASES|~\isa{{\isacharparenleft}fn\ facts\ {\isacharequal}{\isachargreater}\ cases{\isacharunderscore}tactic{\isacharparenright}} wraps + \isa{cases{\isacharunderscore}tactic} depending on goal facts as proof method with + cases; the goal context is passed via method syntax. + + \item \verb|METHOD|~\isa{{\isacharparenleft}fn\ facts\ {\isacharequal}{\isachargreater}\ tactic{\isacharparenright}} wraps \isa{tactic} depending on goal facts as regular proof method; the goal + context is passed via method syntax. + + \item \verb|SIMPLE_METHOD|~\isa{tactic} wraps a tactic that + addresses all subgoals uniformly as simple proof method. Goal facts + are already inserted into all subgoals before \isa{tactic} is + applied. + + \item \verb|SIMPLE_METHOD'|~\isa{tactic} wraps a tactic that + addresses a specific subgoal as simple proof method. Goal facts are + already inserted into the first subgoal before \isa{tactic} is + applied to the same. + + \item \verb|HEADGOAL|~\isa{tactic} applies \isa{tactic} to + the first subgoal. This is convenient to reproduce part the \verb|SIMPLE_METHOD'| wrapping within regular \verb|METHOD|, for example. + + \item \verb|Method.insert_tac|~\isa{facts\ i} inserts \isa{facts} into subgoal \isa{i}. This is convenient to reproduce + part of the \verb|SIMPLE_METHOD| or \verb|SIMPLE_METHOD'| wrapping + within regular \verb|METHOD|, for example. + + \item \verb|Method.setup|~\isa{name\ parser\ description} provides + the functionality of the Isar command \hyperlink{command.method-setup}{\mbox{\isa{\isacommand{method{\isacharunderscore}setup}}}} as ML + function. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isadelimmlex +% +\endisadelimmlex +% +\isatagmlex +% +\begin{isamarkuptext}% +See also \hyperlink{command.method-setup}{\mbox{\isa{\isacommand{method{\isacharunderscore}setup}}}} in + \cite{isabelle-isar-ref} which includes some abstract examples. + + \medskip The following toy examples illustrate how the goal facts + and state are passed to proof methods. The pre-defined proof method + called ``\hyperlink{method.tactic}{\mbox{\isa{tactic}}}'' wraps ML source of type \verb|tactic| (abstracted over \verb|facts|). This allows immediate + experimentation without parsing of concrete syntax.% \end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsection{Attributes% +\endisatagmlex +{\isafoldmlex}% +% +\isadelimmlex +% +\endisadelimmlex +\isacommand{example{\isacharunderscore}proof}\isamarkupfalse% +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{assume}\isamarkupfalse% +\ a{\isacharcolon}\ A\ \isakeyword{and}\ b{\isacharcolon}\ B\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}tactic\ {\isacharverbatimopen}\ rtac\ % +\isaantiq +thm\ conjI% +\endisaantiq +\ {\isadigit{1}}\ {\isacharverbatimclose}{\isacharparenright}\isanewline +\ \ \ \ \isacommand{using}\isamarkupfalse% +\ a\ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}tactic\ {\isacharverbatimopen}\ resolve{\isacharunderscore}tac\ facts\ {\isadigit{1}}\ {\isacharverbatimclose}{\isacharparenright}\isanewline +\ \ \ \ \isacommand{using}\isamarkupfalse% +\ b\ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}tactic\ {\isacharverbatimopen}\ resolve{\isacharunderscore}tac\ facts\ {\isadigit{1}}\ {\isacharverbatimclose}{\isacharparenright}\isanewline +\ \ \ \ \isacommand{done}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{using}\isamarkupfalse% +\ a\ \isakeyword{and}\ b% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +% +\isadelimML +\ \ \ \ % +\endisadelimML +% +\isatagML +\isacommand{ML{\isacharunderscore}val}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharat}{\isacharbraceleft}Isar{\isachardot}goal{\isacharbraceright}{\isachardoublequoteclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +\isanewline +% +\endisadelimML +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}tactic\ {\isacharverbatimopen}\ Method{\isachardot}insert{\isacharunderscore}tac\ facts\ {\isadigit{1}}\ {\isacharverbatimclose}{\isacharparenright}\isanewline +\ \ \ \ \isacommand{apply}\isamarkupfalse% +\ {\isacharparenleft}tactic\ {\isacharverbatimopen}\ {\isacharparenleft}rtac\ % +\isaantiq +thm\ conjI% +\endisaantiq +\ THEN{\isacharunderscore}ALL{\isacharunderscore}NEW\ atac{\isacharparenright}\ {\isadigit{1}}\ {\isacharverbatimclose}{\isacharparenright}\isanewline +\ \ \ \ \isacommand{done}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\medskip The next example implements a method that simplifies + the first subgoal by rewrite rules given as arguments.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{method{\isacharunderscore}setup}\isamarkupfalse% +\ my{\isacharunderscore}simp\ {\isacharequal}\ {\isacharverbatimopen}\isanewline +\ \ Attrib{\isachardot}thms\ {\isachargreater}{\isachargreater}\ {\isacharparenleft}fn\ thms\ {\isacharequal}{\isachargreater}\ fn\ ctxt\ {\isacharequal}{\isachargreater}\isanewline +\ \ \ \ SIMPLE{\isacharunderscore}METHOD{\isacharprime}\ {\isacharparenleft}fn\ i\ {\isacharequal}{\isachargreater}\isanewline +\ \ \ \ \ \ CHANGED\ {\isacharparenleft}asm{\isacharunderscore}full{\isacharunderscore}simp{\isacharunderscore}tac\isanewline +\ \ \ \ \ \ \ \ {\isacharparenleft}HOL{\isacharunderscore}basic{\isacharunderscore}ss\ addsimps\ thms{\isacharparenright}\ i{\isacharparenright}{\isacharparenright}{\isacharparenright}\isanewline +{\isacharverbatimclose}\ {\isachardoublequoteopen}rewrite\ subgoal\ by\ given\ rules{\isachardoublequoteclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +The concrete syntax wrapping of \hyperlink{command.method-setup}{\mbox{\isa{\isacommand{method{\isacharunderscore}setup}}}} always + passes-through the proof context at the end of parsing, but it is + not used in this example. + + The \verb|Attrib.thms| parser produces a list of theorems from the + usual Isar syntax involving attribute expressions etc.\ (syntax + category \hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}). The resulting \verb|thms| are + added to \verb|HOL_basic_ss| which already contains the basic + Simplifier setup for HOL. + + The tactic \verb|asm_full_simp_tac| is the one that is also used in + method \hyperlink{method.simp}{\mbox{\isa{simp}}} by default. The extra wrapping by the \verb|CHANGED| tactical ensures progress of simplification: identical goal + states are filtered out explicitly to make the raw tactic conform to + standard Isar method behaviour. + + \medskip Method \hyperlink{method.my-simp}{\mbox{\isa{my{\isacharunderscore}simp}}} can be used in Isar proofs like + this:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{example{\isacharunderscore}proof}\isamarkupfalse% +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{fix}\isamarkupfalse% +\ a\ b\ c\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ a{\isacharcolon}\ {\isachardoublequoteopen}a\ {\isacharequal}\ b{\isachardoublequoteclose}\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ b{\isacharcolon}\ {\isachardoublequoteopen}b\ {\isacharequal}\ c{\isachardoublequoteclose}\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}a\ {\isacharequal}\ c{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}my{\isacharunderscore}simp\ a\ b{\isacharparenright}\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Here is a similar method that operates on all subgoals, + instead of just the first one.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{method{\isacharunderscore}setup}\isamarkupfalse% +\ my{\isacharunderscore}simp{\isacharunderscore}all\ {\isacharequal}\ {\isacharverbatimopen}\isanewline +\ \ Attrib{\isachardot}thms\ {\isachargreater}{\isachargreater}\ {\isacharparenleft}fn\ thms\ {\isacharequal}{\isachargreater}\ fn\ ctxt\ {\isacharequal}{\isachargreater}\isanewline +\ \ \ \ SIMPLE{\isacharunderscore}METHOD\isanewline +\ \ \ \ \ \ {\isacharparenleft}CHANGED\isanewline +\ \ \ \ \ \ \ \ {\isacharparenleft}ALLGOALS\ {\isacharparenleft}asm{\isacharunderscore}full{\isacharunderscore}simp{\isacharunderscore}tac\isanewline +\ \ \ \ \ \ \ \ \ \ {\isacharparenleft}HOL{\isacharunderscore}basic{\isacharunderscore}ss\ addsimps\ thms{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}\isanewline +{\isacharverbatimclose}\ {\isachardoublequoteopen}rewrite\ all\ subgoals\ by\ given\ rules{\isachardoublequoteclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +\isanewline +\isanewline +\isacommand{example{\isacharunderscore}proof}\isamarkupfalse% +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{fix}\isamarkupfalse% +\ a\ b\ c\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ a{\isacharcolon}\ {\isachardoublequoteopen}a\ {\isacharequal}\ b{\isachardoublequoteclose}\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ b{\isacharcolon}\ {\isachardoublequoteopen}b\ {\isacharequal}\ c{\isachardoublequoteclose}\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}a\ {\isacharequal}\ c{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}c\ {\isacharequal}\ b{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}my{\isacharunderscore}simp{\isacharunderscore}all\ a\ b{\isacharparenright}\isanewline +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\medskip Apart from explicit arguments, common proof methods + typically work with a default configuration provided by the context. + As a shortcut to rule management we use a cheap solution via functor + \verb|Named_Thms| (see also \hyperlink{file.~~/src/Pure/Tools/named-thms.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Tools{\isacharslash}named{\isacharunderscore}thms{\isachardot}ML}}}}).% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ structure\ My{\isacharunderscore}Simps\ {\isacharequal}\isanewline +\ \ \ \ Named{\isacharunderscore}Thms\isanewline +\ \ \ \ \ \ {\isacharparenleft}val\ name\ {\isacharequal}\ {\isachardoublequote}my{\isacharunderscore}simp{\isachardoublequote}\ val\ description\ {\isacharequal}\ {\isachardoublequote}my{\isacharunderscore}simp\ rule{\isachardoublequote}{\isacharparenright}\isanewline +{\isacharverbatimclose}\isanewline +\isacommand{setup}\isamarkupfalse% +\ My{\isacharunderscore}Simps{\isachardot}setup% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +This provides ML access to a list of theorems in canonical + declaration order via \verb|My_Simps.get|. The user can add or + delete rules via the attribute \hyperlink{attribute.my-simp}{\mbox{\isa{my{\isacharunderscore}simp}}}. The actual + proof method is now defined as before, but we append the explicit + arguments and the rules from the context.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{method{\isacharunderscore}setup}\isamarkupfalse% +\ my{\isacharunderscore}simp{\isacharprime}\ {\isacharequal}\ {\isacharverbatimopen}\isanewline +\ \ Attrib{\isachardot}thms\ {\isachargreater}{\isachargreater}\ {\isacharparenleft}fn\ thms\ {\isacharequal}{\isachargreater}\ fn\ ctxt\ {\isacharequal}{\isachargreater}\isanewline +\ \ \ \ SIMPLE{\isacharunderscore}METHOD{\isacharprime}\ {\isacharparenleft}fn\ i\ {\isacharequal}{\isachargreater}\isanewline +\ \ \ \ \ \ CHANGED\ {\isacharparenleft}asm{\isacharunderscore}full{\isacharunderscore}simp{\isacharunderscore}tac\isanewline +\ \ \ \ \ \ \ \ {\isacharparenleft}HOL{\isacharunderscore}basic{\isacharunderscore}ss\ addsimps\ {\isacharparenleft}thms\ {\isacharat}\ My{\isacharunderscore}Simps{\isachardot}get\ ctxt{\isacharparenright}{\isacharparenright}\ i{\isacharparenright}{\isacharparenright}{\isacharparenright}\isanewline +{\isacharverbatimclose}\ {\isachardoublequoteopen}rewrite\ subgoal\ by\ given\ rules\ and\ my{\isacharunderscore}simp\ rules\ from\ the\ context{\isachardoublequoteclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +\medskip Method \hyperlink{method.my-simp'}{\mbox{\isa{my{\isacharunderscore}simp{\isacharprime}}}} can be used in Isar proofs + like this:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{example{\isacharunderscore}proof}\isamarkupfalse% +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{fix}\isamarkupfalse% +\ a\ b\ c\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isacharbrackleft}my{\isacharunderscore}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}a\ {\isasymequiv}\ b{\isachardoublequoteclose}\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isacharbrackleft}my{\isacharunderscore}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}b\ {\isasymequiv}\ c{\isachardoublequoteclose}\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}a\ {\isasymequiv}\ c{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ my{\isacharunderscore}simp{\isacharprime}\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\medskip The \hyperlink{method.my-simp}{\mbox{\isa{my{\isacharunderscore}simp}}} variants defined above are + ``simple'' methods, i.e.\ the goal facts are merely inserted as goal + premises by the \verb|SIMPLE_METHOD'| or \verb|SIMPLE_METHOD| wrapper. + For proof methods that are similar to the standard collection of + \hyperlink{method.simp}{\mbox{\isa{simp}}}, \hyperlink{method.blast}{\mbox{\isa{blast}}}, \hyperlink{method.auto}{\mbox{\isa{auto}}} little more can be + done here. + + Note that using the primary goal facts in the same manner as the + method arguments obtained via concrete syntax or the context does + not meet the requirement of ``strong emphasis on facts'' of regular + proof methods, because rewrite rules as used above can be easily + ignored. A proof text ``\hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{foo}~\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{my{\isacharunderscore}simp}'' where \isa{foo} is not used would + deceive the reader. + + \medskip The technical treatment of rules from the context requires + further attention. Above we rebuild a fresh \verb|simpset| from + the arguments and \emph{all} rules retrieved from the context on + every invocation of the method. This does not scale to really large + collections of rules, which easily emerges in the context of a big + theory library, for example. + + This is an inherent limitation of the simplistic rule management via + functor \verb|Named_Thms|, because it lacks tool-specific + storage and retrieval. More realistic applications require + efficient index-structures that organize theorems in a customized + manner, such as a discrimination net that is indexed by the + left-hand sides of rewrite rules. For variations on the Simplifier, + re-use of the existing type \verb|simpset| is adequate, but + scalability requires it be maintained statically within the context + data, not dynamically on each tool invocation.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Attributes \label{sec:attributes}% } \isamarkuptrue% % \begin{isamarkuptext}% -FIXME% +An \emph{attribute} is a function \isa{context\ {\isasymtimes}\ thm\ {\isasymrightarrow}\ context\ {\isasymtimes}\ thm}, which means both a (generic) context and a theorem + can be modified simultaneously. In practice this fully general form + is very rare, instead attributes are presented either as + \emph{declaration attribute:} \isa{thm\ {\isasymrightarrow}\ context\ {\isasymrightarrow}\ context} or + \emph{rule attribute:} \isa{context\ {\isasymrightarrow}\ thm\ {\isasymrightarrow}\ thm}. + + Attributes can have additional arguments via concrete syntax. There + is a collection of context-sensitive parsers for various logical + entities (types, terms, theorems). These already take care of + applying morphisms to the arguments when attribute expressions are + moved into a different context (see also \secref{sec:morphisms}). + + When implementing declaration attributes, it is important to operate + exactly on the variant of the generic context that is provided by + the system, which is either global theory context or local proof + context. In particular, the background theory of a local context + must not be modified in this situation!% \end{isamarkuptext}% \isamarkuptrue% % +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML type}{attribute}\verb|type attribute = Context.generic * thm -> Context.generic * thm| \\ + \indexdef{}{ML}{Thm.rule\_attribute}\verb|Thm.rule_attribute: (Context.generic -> thm -> thm) -> attribute| \\ + \indexdef{}{ML}{Thm.declaration\_attribute}\verb|Thm.declaration_attribute: |\isasep\isanewline% +\verb| (thm -> Context.generic -> Context.generic) -> attribute| \\ + \indexdef{}{ML}{Attrib.setup}\verb|Attrib.setup: binding -> attribute context_parser ->|\isasep\isanewline% +\verb| string -> theory -> theory| \\ + \end{mldecls} + + \begin{description} + + \item Type \verb|attribute| represents attributes as concrete + type alias. + + \item \verb|Thm.rule_attribute|~\isa{{\isacharparenleft}fn\ context\ {\isacharequal}{\isachargreater}\ rule{\isacharparenright}} wraps + a context-dependent rule (mapping on \verb|thm|) as attribute. + + \item \verb|Thm.declaration_attribute|~\isa{{\isacharparenleft}fn\ thm\ {\isacharequal}{\isachargreater}\ decl{\isacharparenright}} + wraps a theorem-dependent declaration (mapping on \verb|Context.generic|) as attribute. + + \item \verb|Attrib.setup|~\isa{name\ parser\ description} provides + the functionality of the Isar command \hyperlink{command.attribute-setup}{\mbox{\isa{\isacommand{attribute{\isacharunderscore}setup}}}} as + ML function. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isadelimmlex +% +\endisadelimmlex +% +\isatagmlex +% +\begin{isamarkuptext}% +See also \hyperlink{command.attribute-setup}{\mbox{\isa{\isacommand{attribute{\isacharunderscore}setup}}}} in + \cite{isabelle-isar-ref} which includes some abstract examples.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlex +{\isafoldmlex}% +% +\isadelimmlex +% +\endisadelimmlex +% \isadelimtheory % \endisadelimtheory diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/document/Local_Theory.tex --- a/doc-src/IsarImplementation/Thy/document/Local_Theory.tex Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/document/Local_Theory.tex Mon Oct 25 08:08:08 2010 +0200 @@ -49,8 +49,8 @@ Many definitional packages for local theories are available in Isabelle. Although a few old packages only work for global - theories, the local theory interface is already the standard way of - implementing definitional packages in Isabelle.% + theories, the standard way of implementing definitional packages in + Isabelle is via the local theory interface.% \end{isamarkuptext}% \isamarkuptrue% % @@ -105,12 +105,11 @@ \framebox{\quad auxiliary context \quad\framebox{\quad target context \quad\framebox{\quad background theory\quad}}} \end{center} - \noindent When a definitional package is finished, the auxiliary - context is reset to the target context. The target now holds - definitions for terms and theorems that stem from the hypothetical - \isa{{\isasymDEFINE}} and \isa{{\isasymNOTE}} elements, transformed by - the particular target policy (see - \cite[\S4--5]{Haftmann-Wenzel:2009} for details).% + When a definitional package is finished, the auxiliary context is + reset to the target context. The target now holds definitions for + terms and theorems that stem from the hypothetical \isa{{\isasymDEFINE}} and \isa{{\isasymNOTE}} elements, transformed by the + particular target policy (see \cite[\S4--5]{Haftmann-Wenzel:2009} + for details).% \end{isamarkuptext}% \isamarkuptrue% % @@ -132,8 +131,8 @@ \begin{description} - \item \verb|local_theory| represents local theories. Although - this is merely an alias for \verb|Proof.context|, it is + \item Type \verb|local_theory| represents local theories. + Although this is merely an alias for \verb|Proof.context|, it is semantically a subtype of the same: a \verb|local_theory| holds target information as special context data. Subtyping means that any value \isa{lthy{\isacharcolon}}~\verb|local_theory| can be also used @@ -185,12 +184,14 @@ % \endisadelimmlref % -\isamarkupsection{Morphisms and declarations% +\isamarkupsection{Morphisms and declarations \label{sec:morphisms}% } \isamarkuptrue% % \begin{isamarkuptext}% -FIXME% +FIXME + + \medskip See also \cite{Chaieb-Wenzel:2007}.% \end{isamarkuptext}% \isamarkuptrue% % diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/document/Logic.tex --- a/doc-src/IsarImplementation/Thy/document/Logic.tex Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/document/Logic.tex Mon Oct 25 08:08:08 2010 +0200 @@ -132,8 +132,8 @@ \indexdef{}{ML type}{sort}\verb|type sort = class list| \\ \indexdef{}{ML type}{arity}\verb|type arity = string * sort list * sort| \\ \indexdef{}{ML type}{typ}\verb|type typ| \\ - \indexdef{}{ML}{map\_atyps}\verb|map_atyps: (typ -> typ) -> typ -> typ| \\ - \indexdef{}{ML}{fold\_atyps}\verb|fold_atyps: (typ -> 'a -> 'a) -> typ -> 'a -> 'a| \\ + \indexdef{}{ML}{Term.map\_atyps}\verb|Term.map_atyps: (typ -> typ) -> typ -> typ| \\ + \indexdef{}{ML}{Term.fold\_atyps}\verb|Term.fold_atyps: (typ -> 'a -> 'a) -> typ -> 'a -> 'a| \\ \end{mldecls} \begin{mldecls} \indexdef{}{ML}{Sign.subsort}\verb|Sign.subsort: theory -> sort * sort -> bool| \\ @@ -147,23 +147,24 @@ \begin{description} - \item \verb|class| represents type classes. + \item Type \verb|class| represents type classes. - \item \verb|sort| represents sorts, i.e.\ finite intersections - of classes. The empty list \verb|[]: sort| refers to the empty - class intersection, i.e.\ the ``full sort''. + \item Type \verb|sort| represents sorts, i.e.\ finite + intersections of classes. The empty list \verb|[]: sort| refers to + the empty class intersection, i.e.\ the ``full sort''. - \item \verb|arity| represents type arities. A triple \isa{{\isacharparenleft}{\isasymkappa}{\isacharcomma}\ \isactrlvec s{\isacharcomma}\ s{\isacharparenright}\ {\isacharcolon}\ arity} represents \isa{{\isasymkappa}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}\isactrlvec s{\isacharparenright}s} as - described above. + \item Type \verb|arity| represents type arities. A triple + \isa{{\isacharparenleft}{\isasymkappa}{\isacharcomma}\ \isactrlvec s{\isacharcomma}\ s{\isacharparenright}\ {\isacharcolon}\ arity} represents \isa{{\isasymkappa}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}\isactrlvec s{\isacharparenright}s} as described above. - \item \verb|typ| represents types; this is a datatype with + \item Type \verb|typ| represents types; this is a datatype with constructors \verb|TFree|, \verb|TVar|, \verb|Type|. - \item \verb|map_atyps|~\isa{f\ {\isasymtau}} applies the mapping \isa{f} - to all atomic types (\verb|TFree|, \verb|TVar|) occurring in \isa{{\isasymtau}}. + \item \verb|Term.map_atyps|~\isa{f\ {\isasymtau}} applies the mapping \isa{f} to all atomic types (\verb|TFree|, \verb|TVar|) occurring in + \isa{{\isasymtau}}. - \item \verb|fold_atyps|~\isa{f\ {\isasymtau}} iterates the operation \isa{f} over all occurrences of atomic types (\verb|TFree|, \verb|TVar|) - in \isa{{\isasymtau}}; the type structure is traversed from left to right. + \item \verb|Term.fold_atyps|~\isa{f\ {\isasymtau}} iterates the operation + \isa{f} over all occurrences of atomic types (\verb|TFree|, \verb|TVar|) in \isa{{\isasymtau}}; the type structure is traversed from left to + right. \item \verb|Sign.subsort|~\isa{thy\ {\isacharparenleft}s\isactrlisub {\isadigit{1}}{\isacharcomma}\ s\isactrlisub {\isadigit{2}}{\isacharparenright}} tests the subsort relation \isa{s\isactrlisub {\isadigit{1}}\ {\isasymsubseteq}\ s\isactrlisub {\isadigit{2}}}. @@ -196,6 +197,64 @@ % \endisadelimmlref % +\isadelimmlantiq +% +\endisadelimmlantiq +% +\isatagmlantiq +% +\begin{isamarkuptext}% +\begin{matharray}{rcl} + \indexdef{}{ML antiquotation}{class}\hypertarget{ML antiquotation.class}{\hyperlink{ML antiquotation.class}{\mbox{\isa{class}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{sort}\hypertarget{ML antiquotation.sort}{\hyperlink{ML antiquotation.sort}{\mbox{\isa{sort}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{type\_name}\hypertarget{ML antiquotation.type-name}{\hyperlink{ML antiquotation.type-name}{\mbox{\isa{type{\isacharunderscore}name}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{type\_abbrev}\hypertarget{ML antiquotation.type-abbrev}{\hyperlink{ML antiquotation.type-abbrev}{\mbox{\isa{type{\isacharunderscore}abbrev}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{nonterminal}\hypertarget{ML antiquotation.nonterminal}{\hyperlink{ML antiquotation.nonterminal}{\mbox{\isa{nonterminal}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{typ}\hypertarget{ML antiquotation.typ}{\hyperlink{ML antiquotation.typ}{\mbox{\isa{typ}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \end{matharray} + + \begin{rail} + 'class' nameref + ; + 'sort' sort + ; + ('type\_name' | 'type\_abbrev' | 'nonterminal') nameref + ; + 'typ' type + ; + \end{rail} + + \begin{description} + + \item \isa{{\isacharat}{\isacharbraceleft}class\ c{\isacharbraceright}} inlines the internalized class \isa{c} --- as \verb|string| literal. + + \item \isa{{\isacharat}{\isacharbraceleft}sort\ s{\isacharbraceright}} inlines the internalized sort \isa{s} + --- as \verb|string list| literal. + + \item \isa{{\isacharat}{\isacharbraceleft}type{\isacharunderscore}name\ c{\isacharbraceright}} inlines the internalized type + constructor \isa{c} --- as \verb|string| literal. + + \item \isa{{\isacharat}{\isacharbraceleft}type{\isacharunderscore}abbrev\ c{\isacharbraceright}} inlines the internalized type + abbreviation \isa{c} --- as \verb|string| literal. + + \item \isa{{\isacharat}{\isacharbraceleft}nonterminal\ c{\isacharbraceright}} inlines the internalized syntactic + type~/ grammar nonterminal \isa{c} --- as \verb|string| + literal. + + \item \isa{{\isacharat}{\isacharbraceleft}typ\ {\isasymtau}{\isacharbraceright}} inlines the internalized type \isa{{\isasymtau}} + --- as constructor term for datatype \verb|typ|. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlantiq +{\isafoldmlantiq}% +% +\isadelimmlantiq +% +\endisadelimmlantiq +% \isamarkupsection{Terms \label{sec:terms}% } \isamarkuptrue% @@ -248,7 +307,7 @@ polymorphic constants that the user-level type-checker would reject due to violation of type class restrictions. - \medskip An \emph{atomic} term is either a variable or constant. + \medskip An \emph{atomic term} is either a variable or constant. The logical category \emph{term} is defined inductively over atomic terms, with abstraction and application as follows: \isa{t\ {\isacharequal}\ b\ {\isacharbar}\ x\isactrlisub {\isasymtau}\ {\isacharbar}\ {\isacharquery}x\isactrlisub {\isasymtau}\ {\isacharbar}\ c\isactrlisub {\isasymtau}\ {\isacharbar}\ {\isasymlambda}\isactrlisub {\isasymtau}{\isachardot}\ t\ {\isacharbar}\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}}. Parsing and printing takes care of converting between an external representation with named bound @@ -317,11 +376,11 @@ \begin{isamarkuptext}% \begin{mldecls} \indexdef{}{ML type}{term}\verb|type term| \\ - \indexdef{}{ML}{op aconv}\verb|op aconv: term * term -> bool| \\ - \indexdef{}{ML}{map\_types}\verb|map_types: (typ -> typ) -> term -> term| \\ - \indexdef{}{ML}{fold\_types}\verb|fold_types: (typ -> 'a -> 'a) -> term -> 'a -> 'a| \\ - \indexdef{}{ML}{map\_aterms}\verb|map_aterms: (term -> term) -> term -> term| \\ - \indexdef{}{ML}{fold\_aterms}\verb|fold_aterms: (term -> 'a -> 'a) -> term -> 'a -> 'a| \\ + \indexdef{}{ML}{aconv}\verb|op aconv: term * term -> bool| \\ + \indexdef{}{ML}{Term.map\_types}\verb|Term.map_types: (typ -> typ) -> term -> term| \\ + \indexdef{}{ML}{Term.fold\_types}\verb|Term.fold_types: (typ -> 'a -> 'a) -> term -> 'a -> 'a| \\ + \indexdef{}{ML}{Term.map\_aterms}\verb|Term.map_aterms: (term -> term) -> term -> term| \\ + \indexdef{}{ML}{Term.fold\_aterms}\verb|Term.fold_aterms: (term -> 'a -> 'a) -> term -> 'a -> 'a| \\ \end{mldecls} \begin{mldecls} \indexdef{}{ML}{fastype\_of}\verb|fastype_of: term -> typ| \\ @@ -337,24 +396,24 @@ \begin{description} - \item \verb|term| represents de-Bruijn terms, with comments in - abstractions, and explicitly named free variables and constants; + \item Type \verb|term| represents de-Bruijn terms, with comments + in abstractions, and explicitly named free variables and constants; this is a datatype with constructors \verb|Bound|, \verb|Free|, \verb|Var|, \verb|Const|, \verb|Abs|, \verb|op $|. \item \isa{t}~\verb|aconv|~\isa{u} checks \isa{{\isasymalpha}}-equivalence of two terms. This is the basic equality relation on type \verb|term|; raw datatype equality should only be used for operations related to parsing or printing! - \item \verb|map_types|~\isa{f\ t} applies the mapping \isa{f} to all types occurring in \isa{t}. + \item \verb|Term.map_types|~\isa{f\ t} applies the mapping \isa{f} to all types occurring in \isa{t}. - \item \verb|fold_types|~\isa{f\ t} iterates the operation \isa{f} over all occurrences of types in \isa{t}; the term + \item \verb|Term.fold_types|~\isa{f\ t} iterates the operation + \isa{f} over all occurrences of types in \isa{t}; the term structure is traversed from left to right. - \item \verb|map_aterms|~\isa{f\ t} applies the mapping \isa{f} - to all atomic terms (\verb|Bound|, \verb|Free|, \verb|Var|, \verb|Const|) occurring in \isa{t}. + \item \verb|Term.map_aterms|~\isa{f\ t} applies the mapping \isa{f} to all atomic terms (\verb|Bound|, \verb|Free|, \verb|Var|, \verb|Const|) occurring in \isa{t}. - \item \verb|fold_aterms|~\isa{f\ t} iterates the operation \isa{f} over all occurrences of atomic terms (\verb|Bound|, \verb|Free|, - \verb|Var|, \verb|Const|) in \isa{t}; the term structure is + \item \verb|Term.fold_aterms|~\isa{f\ t} iterates the operation + \isa{f} over all occurrences of atomic terms (\verb|Bound|, \verb|Free|, \verb|Var|, \verb|Const|) in \isa{t}; the term structure is traversed from left to right. \item \verb|fastype_of|~\isa{t} determines the type of a @@ -389,6 +448,63 @@ % \endisadelimmlref % +\isadelimmlantiq +% +\endisadelimmlantiq +% +\isatagmlantiq +% +\begin{isamarkuptext}% +\begin{matharray}{rcl} + \indexdef{}{ML antiquotation}{const\_name}\hypertarget{ML antiquotation.const-name}{\hyperlink{ML antiquotation.const-name}{\mbox{\isa{const{\isacharunderscore}name}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{const\_abbrev}\hypertarget{ML antiquotation.const-abbrev}{\hyperlink{ML antiquotation.const-abbrev}{\mbox{\isa{const{\isacharunderscore}abbrev}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{const}\hypertarget{ML antiquotation.const}{\hyperlink{ML antiquotation.const}{\mbox{\isa{const}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{term}\hypertarget{ML antiquotation.term}{\hyperlink{ML antiquotation.term}{\mbox{\isa{term}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{prop}\hypertarget{ML antiquotation.prop}{\hyperlink{ML antiquotation.prop}{\mbox{\isa{prop}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \end{matharray} + + \begin{rail} + ('const\_name' | 'const\_abbrev') nameref + ; + 'const' ('(' (type + ',') ')')? + ; + 'term' term + ; + 'prop' prop + ; + \end{rail} + + \begin{description} + + \item \isa{{\isacharat}{\isacharbraceleft}const{\isacharunderscore}name\ c{\isacharbraceright}} inlines the internalized logical + constant name \isa{c} --- as \verb|string| literal. + + \item \isa{{\isacharat}{\isacharbraceleft}const{\isacharunderscore}abbrev\ c{\isacharbraceright}} inlines the internalized + abbreviated constant name \isa{c} --- as \verb|string| + literal. + + \item \isa{{\isacharat}{\isacharbraceleft}const\ c{\isacharparenleft}\isactrlvec {\isasymtau}{\isacharparenright}{\isacharbraceright}} inlines the internalized + constant \isa{c} with precise type instantiation in the sense of + \verb|Sign.const_instance| --- as \verb|Const| constructor term for + datatype \verb|term|. + + \item \isa{{\isacharat}{\isacharbraceleft}term\ t{\isacharbraceright}} inlines the internalized term \isa{t} + --- as constructor term for datatype \verb|term|. + + \item \isa{{\isacharat}{\isacharbraceleft}prop\ {\isasymphi}{\isacharbraceright}} inlines the internalized proposition + \isa{{\isasymphi}} --- as constructor term for datatype \verb|term|. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlantiq +{\isafoldmlantiq}% +% +\isadelimmlantiq +% +\endisadelimmlantiq +% \isamarkupsection{Theorems \label{sec:thms}% } \isamarkuptrue% @@ -530,7 +646,7 @@ distinct variables \isa{\isactrlvec {\isasymalpha}}). The RHS may mention previously defined constants as above, or arbitrary constants \isa{d{\isacharparenleft}{\isasymalpha}\isactrlisub i{\isacharparenright}} for some \isa{{\isasymalpha}\isactrlisub i} projected from \isa{\isactrlvec {\isasymalpha}}. Thus overloaded definitions essentially work by primitive recursion over the syntactic structure of a single type - argument.% + argument. See also \cite[\S4.3]{Haftmann-Wenzel:2006:classes}.% \end{isamarkuptext}% \isamarkuptrue% % @@ -558,18 +674,20 @@ \indexdef{}{ML}{Thm.generalize}\verb|Thm.generalize: string list * string list -> int -> thm -> thm| \\ \indexdef{}{ML}{Thm.instantiate}\verb|Thm.instantiate: (ctyp * ctyp) list * (cterm * cterm) list -> thm -> thm| \\ \indexdef{}{ML}{Thm.add\_axiom}\verb|Thm.add_axiom: binding * term -> theory -> (string * thm) * theory| \\ - \indexdef{}{ML}{Thm.add\_oracle}\verb|Thm.add_oracle: binding * ('a -> cterm) -> theory|\isasep\isanewline% -\verb| -> (string * ('a -> thm)) * theory| \\ - \indexdef{}{ML}{Thm.add\_def}\verb|Thm.add_def: bool -> bool -> binding * term -> theory -> (string * thm) * theory| \\ + \indexdef{}{ML}{Thm.add\_oracle}\verb|Thm.add_oracle: binding * ('a -> cterm) -> theory ->|\isasep\isanewline% +\verb| (string * ('a -> thm)) * theory| \\ + \indexdef{}{ML}{Thm.add\_def}\verb|Thm.add_def: bool -> bool -> binding * term -> theory ->|\isasep\isanewline% +\verb| (string * thm) * theory| \\ \end{mldecls} \begin{mldecls} - \indexdef{}{ML}{Theory.add\_deps}\verb|Theory.add_deps: string -> string * typ -> (string * typ) list -> theory -> theory| \\ + \indexdef{}{ML}{Theory.add\_deps}\verb|Theory.add_deps: string -> string * typ -> (string * typ) list ->|\isasep\isanewline% +\verb| theory -> theory| \\ \end{mldecls} \begin{description} - \item \verb|ctyp| and \verb|cterm| represent certified types - and terms, respectively. These are abstract datatypes that + \item Types \verb|ctyp| and \verb|cterm| represent certified + types and terms, respectively. These are abstract datatypes that guarantee that its values have passed the full well-formedness (and well-typedness) checks, relative to the declarations of type constructors, constants etc. in the theory. @@ -582,8 +700,8 @@ reasoning loops. There are separate operations to decompose certified entities (including actual theorems). - \item \verb|thm| represents proven propositions. This is an - abstract datatype that guarantees that its values have been + \item Type \verb|thm| represents proven propositions. This is + an abstract datatype that guarantees that its values have been constructed by basic principles of the \verb|Thm| module. Every \verb|thm| value contains a sliding back-reference to the enclosing theory, cf.\ \secref{sec:context-theory}. @@ -639,6 +757,78 @@ % \endisadelimmlref % +\isadelimmlantiq +% +\endisadelimmlantiq +% +\isatagmlantiq +% +\begin{isamarkuptext}% +\begin{matharray}{rcl} + \indexdef{}{ML antiquotation}{ctyp}\hypertarget{ML antiquotation.ctyp}{\hyperlink{ML antiquotation.ctyp}{\mbox{\isa{ctyp}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{cterm}\hypertarget{ML antiquotation.cterm}{\hyperlink{ML antiquotation.cterm}{\mbox{\isa{cterm}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{cprop}\hypertarget{ML antiquotation.cprop}{\hyperlink{ML antiquotation.cprop}{\mbox{\isa{cprop}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{thm}\hypertarget{ML antiquotation.thm}{\hyperlink{ML antiquotation.thm}{\mbox{\isa{thm}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{thms}\hypertarget{ML antiquotation.thms}{\hyperlink{ML antiquotation.thms}{\mbox{\isa{thms}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{lemma}\hypertarget{ML antiquotation.lemma}{\hyperlink{ML antiquotation.lemma}{\mbox{\isa{lemma}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \end{matharray} + + \begin{rail} + 'ctyp' typ + ; + 'cterm' term + ; + 'cprop' prop + ; + 'thm' thmref + ; + 'thms' thmrefs + ; + 'lemma' ('(open)')? ((prop +) + 'and') \\ 'by' method method? + \end{rail} + + \begin{description} + + \item \isa{{\isacharat}{\isacharbraceleft}ctyp\ {\isasymtau}{\isacharbraceright}} produces a certified type wrt.\ the + current background theory --- as abstract value of type \verb|ctyp|. + + \item \isa{{\isacharat}{\isacharbraceleft}cterm\ t{\isacharbraceright}} and \isa{{\isacharat}{\isacharbraceleft}cprop\ {\isasymphi}{\isacharbraceright}} produce a + certified term wrt.\ the current background theory --- as abstract + value of type \verb|cterm|. + + \item \isa{{\isacharat}{\isacharbraceleft}thm\ a{\isacharbraceright}} produces a singleton fact --- as abstract + value of type \verb|thm|. + + \item \isa{{\isacharat}{\isacharbraceleft}thms\ a{\isacharbraceright}} produces a general fact --- as abstract + value of type \verb|thm list|. + + \item \isa{{\isacharat}{\isacharbraceleft}lemma\ {\isasymphi}\ by\ meth{\isacharbraceright}} produces a fact that is proven on + the spot according to the minimal proof, which imitates a terminal + Isar proof. The result is an abstract value of type \verb|thm| + or \verb|thm list|, depending on the number of propositions + given here. + + The internal derivation object lacks a proper theorem name, but it + is formally closed, unless the \isa{{\isacharparenleft}open{\isacharparenright}} option is specified + (this may impact performance of applications with proof terms). + + Since ML antiquotations are always evaluated at compile-time, there + is no run-time overhead even for non-trivial proofs. Nonetheless, + the justification is syntactically limited to a single \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}} step. More complex Isar proofs should be done in regular + theory source, before compiling the corresponding ML text that uses + the result. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlantiq +{\isafoldmlantiq}% +% +\isadelimmlantiq +% +\endisadelimmlantiq +% \isamarkupsubsection{Auxiliary definitions% } \isamarkuptrue% @@ -795,9 +985,9 @@ \end{tabular} \medskip - \noindent Thus we essentially impose nesting levels on propositions - formed from \isa{{\isasymAnd}} and \isa{{\isasymLongrightarrow}}. At each level there is a - prefix of parameters and compound premises, concluding an atomic + Thus we essentially impose nesting levels on propositions formed + from \isa{{\isasymAnd}} and \isa{{\isasymLongrightarrow}}. At each level there is a prefix + of parameters and compound premises, concluding an atomic proposition. Typical examples are \isa{{\isasymlongrightarrow}}-introduction \isa{{\isacharparenleft}A\ {\isasymLongrightarrow}\ B{\isacharparenright}\ {\isasymLongrightarrow}\ A\ {\isasymlongrightarrow}\ B} or mathematical induction \isa{P\ {\isadigit{0}}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymAnd}n{\isachardot}\ P\ n\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ P\ n}. Even deeper nesting occurs in well-founded induction \isa{{\isacharparenleft}{\isasymAnd}x{\isachardot}\ {\isacharparenleft}{\isasymAnd}y{\isachardot}\ y\ {\isasymprec}\ x\ {\isasymLongrightarrow}\ P\ y{\isacharparenright}\ {\isasymLongrightarrow}\ P\ x{\isacharparenright}\ {\isasymLongrightarrow}\ P\ x}, but this already marks the limit of rule complexity that is usually seen in @@ -920,8 +1110,8 @@ % \begin{isamarkuptext}% \begin{mldecls} - \indexdef{}{ML}{op RS}\verb|op RS: thm * thm -> thm| \\ - \indexdef{}{ML}{op OF}\verb|op OF: thm * thm list -> thm| \\ + \indexdef{}{ML}{RS}\verb|op RS: thm * thm -> thm| \\ + \indexdef{}{ML}{OF}\verb|op OF: thm * thm list -> thm| \\ \end{mldecls} \begin{description} diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/document/ML.tex --- a/doc-src/IsarImplementation/Thy/document/ML.tex Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/document/ML.tex Mon Oct 25 08:08:08 2010 +0200 @@ -18,252 +18,2165 @@ % \endisadelimtheory % -\isamarkupchapter{Advanced ML programming% +\isamarkupchapter{Isabelle/ML% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isabelle/ML is best understood as a certain culture based on + Standard ML. Thus it is not a new programming language, but a + certain way to use SML at an advanced level within the Isabelle + environment. This covers a variety of aspects that are geared + towards an efficient and robust platform for applications of formal + logic with fully foundational proof construction --- according to + the well-known \emph{LCF principle}. There are specific library + modules and infrastructure to address the needs for such difficult + tasks. For example, the raw parallel programming model of Poly/ML + is presented as considerably more abstract concept of \emph{future + values}, which is then used to augment the inference kernel, proof + interpreter, and theory loader accordingly. + + The main aspects of Isabelle/ML are introduced below. These + first-hand explanations should help to understand how proper + Isabelle/ML is to be read and written, and to get access to the + wealth of experience that is expressed in the source text and its + history of changes.\footnote{See + \url{http://isabelle.in.tum.de/repos/isabelle} for the full + Mercurial history. There are symbolic tags to refer to official + Isabelle releases, as opposed to arbitrary \emph{tip} versions that + merely reflect snapshots that are never really up-to-date.}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Style and orthography% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The sources of Isabelle/Isar are optimized for + \emph{readability} and \emph{maintainability}. The main purpose is + to tell an informed reader what is really going on and how things + really work. This is a non-trivial aim, but it is supported by a + certain style of writing Isabelle/ML that has emerged from long + years of system development.\footnote{See also the interesting style + guide for OCaml + \url{http://caml.inria.fr/resources/doc/guides/guidelines.en.html} + which shares many of our means and ends.} + + The main principle behind any coding style is \emph{consistency}. + For a single author of a small program this merely means ``choose + your style and stick to it''. A complex project like Isabelle, with + long years of development and different contributors, requires more + standardization. A coding style that is changed every few years or + with every new contributor is no style at all, because consistency + is quickly lost. Global consistency is hard to achieve, though. + One should always strife at least for local consistency of modules + and sub-systems, without deviating from some general principles how + to write Isabelle/ML. + + In a sense, a common coding style is like an \emph{orthography} for + the sources: it helps to read quickly over the text and see through + the main points, without getting distracted by accidental + presentation of free-style source.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Header and sectioning% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isabelle source files have a certain standardized header + format (with precise spacing) that follows ancient traditions + reaching back to the earliest versions of the system by Larry + Paulson. E.g.\ see \hyperlink{file.~~/src/Pure/thm.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}thm{\isachardot}ML}}}}. + + The header includes at least \verb|Title| and \verb|Author| entries, followed by a prose description of the purpose of + the module. The latter can range from a single line to several + paragraphs of explanations. + + The rest of the file is divided into sections, subsections, + subsubsections, paragraphs etc.\ using some a layout via ML comments + as follows. + +\begin{verbatim} +(*** section ***) + +(** subsection **) + +(* subsubsection *) + +(*short paragraph*) + +(* + long paragraph + more text +*) +\end{verbatim} + + As in regular typography, there is some extra space \emph{before} + section headings that are adjacent to plain text (not other headings + as in the example above). + + \medskip The precise wording of the prose text given in these + headings is chosen carefully to indicate the main theme of the + subsequent formal ML text.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Naming conventions% } \isamarkuptrue% % -\isamarkupsection{Style% +\begin{isamarkuptext}% +Since ML is the primary medium to express the meaning of the + source text, naming of ML entities requires special care. + + \paragraph{Notation.} A name consists of 1--3 \emph{words} (rarely + 4, but not more) that are separated by underscore. There are three + variants concerning upper or lower case letters, which are just for + certain ML categories as follows: + + \medskip + \begin{tabular}{lll} + variant & example & ML categories \\\hline + lower-case & \verb|foo_bar| & values, types, record fields \\ + capitalized & \verb|Foo_Bar| & datatype constructors, \\ + & & structures, functors \\ + upper-case & \verb|FOO_BAR| & special values (combinators), \\ + & & exception constructors, signatures \\ + \end{tabular} + \medskip + + For historical reasons, many capitalized names omit underscores, + e.g.\ old-style \verb|FooBar| instead of \verb|Foo_Bar|. + Genuine mixed-case names are \emph{not} used --- clear division of + words is essential for readability.\footnote{Camel-case was invented + to workaround the lack of underscore in some early non-ASCII + character sets and keywords. Later it became a habitual in some + language communities that are now strong in numbers.} + + A single (capital) character does not count as ``word'' in this + respect: some Isabelle/ML are suffixed by extra markers like this: + \verb|foo_barT|. + + Name variants are produced by adding 1--3 primes, e.g.\ \verb|foo'|, \verb|foo''|, or \verb|foo'''|, but not \verb|foo''''| or more. Decimal digits scale better to larger numbers, + e.g.\ \verb|foo0|, \verb|foo1|, \verb|foo42|. + + \paragraph{Scopes.} Apart from very basic library modules, ML + structures are not ``opened'', but names are referenced with + explicit qualification, as in \verb|Syntax.string_of_term| for + example. When devising names for structures and their components it + is important aim at eye-catching compositions of both parts, because + this is how they are always seen in the sources and documentation. + For the same reasons, aliases of well-known library functions should + be avoided. + + Local names of function abstraction or case/lets bindings are + typically shorter, sometimes using only rudiments of ``words'', + while still avoiding cryptic shorthands. An auxiliary function + called \verb|helper|, \verb|aux|, or \verb|f| is + considered bad style. + + Example: + + \begin{verbatim} + (* RIGHT *) + + fun print_foo ctxt foo = + let + fun print t = ... Syntax.string_of_term ctxt t ... + in ... end; + + + (* RIGHT *) + + fun print_foo ctxt foo = + let + val string_of_term = Syntax.string_of_term ctxt; + fun print t = ... string_of_term t ... + in ... end; + + + (* WRONG *) + + val string_of_term = Syntax.string_of_term; + + fun print_foo ctxt foo = + let + fun aux t = ... string_of_term ctxt t ... + in ... end; + + \end{verbatim} + + + \paragraph{Specific conventions.} Here are some important specific + name forms that occur frequently in the sources. + + \begin{itemize} + + \item A function that maps \verb|foo| to \verb|bar| is + called \verb|foo_to_bar| or \verb|bar_of_foo| (never + \verb|foo2bar|, \verb|bar_from_foo|, \verb|bar_for_foo|, or \verb|bar4foo|). + + \item The name component \verb|legacy| means that the operation + is about to be discontinued soon. + + \item The name component \verb|old| means that this is historic + material that might disappear at some later stage. + + \item The name component \verb|global| means that this works + with the background theory instead of the regular local context + (\secref{sec:context}), sometimes for historical reasons, sometimes + due a genuine lack of locality of the concept involved, sometimes as + a fall-back for the lack of a proper context in the application + code. Whenever there is a non-global variant available, the + application should be migrated to use it with a proper local + context. + + \item Variables of the main context types of the Isabelle/Isar + framework (\secref{sec:context} and \chref{ch:local-theory}) have + firm naming conventions to allow to visualize the their data flow + via plain regular expressions in the editor. In particular: + + \begin{itemize} + + \item theories are called \verb|thy|, rarely \verb|theory| + (never \verb|thry|) + + \item proof contexts are called \verb|ctxt|, rarely \verb|context| (never \verb|ctx|) + + \item generic contexts are called \verb|context|, rarely + \verb|ctxt| + + \item local theories are called \verb|lthy|, unless there is an + implicit conversion to a general proof context (semantic super-type) + + \end{itemize} + + Variations with primed or decimal numbers are always possible, as + well as ``sematic prefixes'' like \verb|foo_thy| or \verb|bar_ctxt|, but the base conventions above need to be preserved. + + \item The main logical entities (\secref{ch:logic}) also have + established naming convention: + + \begin{itemize} + + \item sorts are called \verb|S| + + \item types are called \verb|T|, \verb|U|, or \verb|ty| (never \verb|t|) + + \item terms are called \verb|t|, \verb|u|, or \verb|tm| (never \verb|trm|) + + \item certified types are called \verb|cT|, rarely \verb|T|, with variants as for types + + \item certified terms are called \verb|ct|, rarely \verb|t|, with variants as for terms + + \item theorems are called \verb|th|, or \verb|thm| + + \end{itemize} + + Proper semantic names override these conventions completely. For + example, the left-hand side of an equation (as a term) can be called + \verb|lhs| (not \verb|lhs_tm|). Or a term that is treated + specifically as a variable can be called \verb|v| or \verb|x|. + + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{General source layout% } \isamarkuptrue% % \begin{isamarkuptext}% -Like any style guide, also this one should not be interpreted dogmatically, but - with care and discernment. It consists of a collection of - recommendations which have been turned out useful over long years of - Isabelle system development and are supposed to support writing of readable - and managable code. Special cases encourage derivations, - as far as you know what you are doing. - \footnote{This style guide is loosely based on - \url{http://caml.inria.fr/resources/doc/guides/guidelines.en.html}} +The general Isabelle/ML source layout imitates regular + type-setting to some extent, augmented by the requirements for + deeply nested expressions that are commonplace in functional + programming. + + \paragraph{Line length} is 80 characters according to ancient + standards, but we allow as much as 100 characters, not + more.\footnote{Readability requires to keep the beginning of a line + in view while watching its end. Modern wide-screen displays do not + change the way how the human brain works. As a rule of thumb, + sources need to be printable on plain paper.} The extra 20 + characters acknowledge the space requirements due to qualified + library references in Isabelle/ML. + + \paragraph{White-space} is used to emphasize the structure of + expressions, following mostly standard conventions for mathematical + typesetting, as can be seen in plain {\TeX} or {\LaTeX}. This + defines positioning of spaces for parentheses, punctuation, and + infixes as illustrated here: + + \begin{verbatim} + val x = y + z * (a + b); + val pair = (a, b); + val record = {foo = 1, bar = 2}; + \end{verbatim} + + Lines are normally broken \emph{after} an infix operator or + punctuation character. For example: + + \begin{verbatim} + val x = + a + + b + + c; + + val tuple = + (a, + b, + c); + \end{verbatim} + + Some special infixes (e.g.\ \verb||\verb,|,\verb|>|) work better at the + start of the line, but punctuation is always at the end. + + Function application follows the tradition of \isa{{\isasymlambda}}-calculus, + not informal mathematics. For example: \verb|f a b| for a + curried function, or \verb|g (a, b)| for a tupled function. + Note that the space between \verb|g| and the pair \verb|(a, b)| follows the important principle of + \emph{compositionality}: the layout of \verb|g p| does not + change when \verb|p| is refined to a concrete pair. + + \paragraph{Indentation} uses plain spaces, never hard + tabulators.\footnote{Tabulators were invented to move the carriage + of a type-writer to certain predefined positions. In software they + could be used as a primitive run-length compression of consecutive + spaces, but the precise result would depend on non-standardized + editor configuration.} + + Each level of nesting is indented by 2 spaces, sometimes 1, very + rarely 4, never 8. + + Indentation follows a simple logical format that only depends on the + nesting depth, not the accidental length of the text that initiates + a level of nesting. Example: + + \begin{verbatim} + (* RIGHT *) + + if b then + expr1_part1 + expr1_part2 + else + expr2_part1 + expr2_part2 + + + (* WRONG *) + + if b then expr1_part1 + expr1_part2 + else expr2_part1 + expr2_part2 + \end{verbatim} + + The second form has many problems: it assumes a fixed-width font + when viewing the sources, it uses more space on the line and thus + makes it hard to observe its strict length limit (working against + \emph{readability}), it requires extra editing to adapt the layout + to changes of the initial text (working against + \emph{maintainability}) etc. + + \medskip For similar reasons, any kind of two-dimensional or tabular + layouts, ASCII-art with lines or boxes of stars etc. should be + avoided. + + \paragraph{Complex expressions} consisting of multi-clausal function + definitions, \verb|case|, \verb|handle|, \verb|let|, + or combinations require special attention. The syntax of Standard + ML is a bit too ambitious in admitting too much variance that can + distort the meaning of the text. + + Clauses of \verb|fun|, \verb|fn|, \verb|case|, + \verb|handle| get extra indentation to indicate the nesting + clearly. For example: + + \begin{verbatim} + (* RIGHT *) + + fun foo p1 = + expr1 + | foo p2 = + expr2 + + + (* WRONG *) + + fun foo p1 = + expr1 + | foo p2 = + expr2 + \end{verbatim} + + Body expressions consisting of \verb|case| or \verb|let| + require care to maintain compositionality, to prevent loss of + logical indentation where it is particularly imporant to see the + structure of the text later on. Example: + + \begin{verbatim} + (* RIGHT *) + + fun foo p1 = + (case e of + q1 => ... + | q2 => ...) + | foo p2 = + let + ... + in + ... + end + + + (* WRONG *) + + fun foo p1 = case e of + q1 => ... + | q2 => ... + | foo p2 = + let + ... + in + ... + end + \end{verbatim} + + Extra parentheses around \verb|case| expressions are optional, + but help to analyse the nesting easily based on simple string + matching in the editor. + + \medskip There are two main exceptions to the overall principle of + compositionality in the layout of complex expressions. + + \begin{enumerate} + + \item \verb|if| expressions are iterated as if there would be + a multi-branch conditional in SML, e.g. + + \begin{verbatim} + (* RIGHT *) + + if b1 then e1 + else if b2 then e2 + else e3 + \end{verbatim} + + \item \verb|fn| abstractions are often layed-out as if they + would lack any structure by themselves. This traditional form is + motivated by the possibility to shift function arguments back and + forth wrt.\ additional combinators. For example: + + \begin{verbatim} + (* RIGHT *) + + fun foo x y = fold (fn z => + expr) + \end{verbatim} + + Here the visual appearance is that of three arguments \verb|x|, + \verb|y|, \verb|z|. + + \end{enumerate} + + Such weakly structured layout should be use with great care. Here + is a counter-example involving \verb|let| expressions: + + \begin{verbatim} + (* WRONG *) + + fun foo x = let + val y = ... + in ... end + + fun foo x = + let + val y = ... + in ... end + \end{verbatim} + + \medskip In general the source layout is meant to emphasize the + structure of complex language expressions, not to pretend that SML + had a completely different syntax (say that of Haskell or Java).% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{SML embedded into Isabelle/Isar% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +ML and Isar are intertwined via an open-ended bootstrap + process that provides more and more programming facilities and + logical content in an alternating manner. Bootstrapping starts from + the raw environment of existing implementations of Standard ML + (mainly Poly/ML, but also SML/NJ). + + Isabelle/Pure marks the point where the original ML toplevel is + superseded by the Isar toplevel that maintains a uniform environment + for arbitrary ML values (see also \secref{sec:context}). This + formal context holds logical entities as well as ML compiler + bindings, among many other things. Raw Standard ML is never + encountered again after the initial bootstrap of Isabelle/Pure. + + Object-logics such as Isabelle/HOL are built within the + Isabelle/ML/Isar environment of Pure by introducing suitable + theories with associated ML text, either inlined or as separate + files. Thus Isabelle/HOL is defined as a regular user-space + application within the Isabelle framework. Further add-on tools can + be implemented in ML within the Isar context in the same manner: ML + is part of the regular user-space repertoire of Isabelle.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Isar ML commands% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The primary Isar source language provides various facilities + to open a ``window'' to the underlying ML compiler. Especially see + \indexref{}{command}{use}\hyperlink{command.use}{\mbox{\isa{\isacommand{use}}}} and \indexref{}{command}{ML}\hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}}, which work exactly the + same way, only the source text is provided via a file vs.\ inlined, + respectively. Apart from embedding ML into the main theory + definition like that, there are many more commands that refer to ML + source, such as \indexref{}{command}{setup}\hyperlink{command.setup}{\mbox{\isa{\isacommand{setup}}}} or \indexref{}{command}{declaration}\hyperlink{command.declaration}{\mbox{\isa{\isacommand{declaration}}}}. + An example of even more fine-grained embedding of ML into Isar is + the proof method \indexref{}{method}{tactic}\hyperlink{method.tactic}{\mbox{\isa{tactic}}}, which refines the pending goal state + via a given expression of type \verb|tactic|.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlex +% +\endisadelimmlex +% +\isatagmlex +% +\begin{isamarkuptext}% +The following artificial example demonstrates some ML + toplevel declarations within the implicit Isar theory context. This + is regular functional programming without referring to logical + entities yet.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlex +{\isafoldmlex}% +% +\isadelimmlex +% +\endisadelimmlex +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ fun\ factorial\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}\isanewline +\ \ \ \ {\isacharbar}\ factorial\ n\ {\isacharequal}\ n\ {\isacharasterisk}\ factorial\ {\isacharparenleft}n\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}\isanewline +{\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +Here the ML environment is really managed by Isabelle, i.e.\ + the \verb|factorial| function is not yet accessible in the preceding + paragraph, nor in a different theory that is independent from the + current one in the import hierarchy. + + Removing the above ML declaration from the source text will remove + any trace of this definition as expected. The Isabelle/ML toplevel + environment is managed in a \emph{stateless} way: unlike the raw ML + toplevel or similar command loops of Computer Algebra systems, there + are no global side-effects involved here.\footnote{Such a stateless + compilation environment is also a prerequisite for robust parallel + compilation within independent nodes of the implicit theory + development graph.} + + \medskip The next example shows how to embed ML into Isar proofs. + Instead of \indexref{}{command}{ML}\hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}} for theory mode, we use \indexref{}{command}{ML\_prf}\hyperlink{command.ML-prf}{\mbox{\isa{\isacommand{ML{\isacharunderscore}prf}}}} for proof mode. As illustrated below, its effect on the + ML environment is local to the whole proof body, while ignoring its + internal block structure.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{example{\isacharunderscore}proof}\isamarkupfalse% +\isanewline +% +\isadelimML +\ \ % +\endisadelimML +% +\isatagML +\isacommand{ML{\isacharunderscore}prf}\isamarkupfalse% +\ {\isacharverbatimopen}\ val\ a\ {\isacharequal}\ {\isadigit{1}}\ {\isacharverbatimclose}\isanewline +\ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse% +\ % +\isamarkupcmt{Isar block structure ignored by ML environment% +} +\isanewline +\ \ \ \ \isacommand{ML{\isacharunderscore}prf}\isamarkupfalse% +\ {\isacharverbatimopen}\ val\ b\ {\isacharequal}\ a\ {\isacharplus}\ {\isadigit{1}}\ {\isacharverbatimclose}\isanewline +\ \ \isacommand{{\isacharbraceright}}\isamarkupfalse% +\ % +\isamarkupcmt{Isar block structure ignored by ML environment% +} +\isanewline +\ \ \isacommand{ML{\isacharunderscore}prf}\isamarkupfalse% +\ {\isacharverbatimopen}\ val\ c\ {\isacharequal}\ b\ {\isacharplus}\ {\isadigit{1}}\ {\isacharverbatimclose}\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +By side-stepping the normal scoping rules for Isar proof + blocks, embedded ML code can refer to the different contexts + explicitly, and manipulate corresponding entities, e.g.\ export a + fact from a block context. + + \medskip Two further ML commands are useful in certain situations: + \indexref{}{command}{ML\_val}\hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isacharunderscore}val}}}} and \indexref{}{command}{ML\_command}\hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isacharunderscore}command}}}} are both + \emph{diagnostic} in the sense that there is no effect on the + underlying environment, and can thus used anywhere (even outside a + theory). The examples below produce long strings of digits by + invoking \verb|factorial|: \hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isacharunderscore}val}}}} already takes care of + printing the ML toplevel result, but \hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isacharunderscore}command}}}} is silent + so we produce an explicit output message.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML{\isacharunderscore}val}\isamarkupfalse% +\ {\isacharverbatimopen}\ factorial\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharverbatimclose}\isanewline +\isacommand{ML{\isacharunderscore}command}\isamarkupfalse% +\ {\isacharverbatimopen}\ writeln\ {\isacharparenleft}string{\isacharunderscore}of{\isacharunderscore}int\ {\isacharparenleft}factorial\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharparenright}{\isacharparenright}\ {\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +\isanewline +\isanewline +\isacommand{example{\isacharunderscore}proof}\isamarkupfalse% +\isanewline +% +\isadelimML +\ \ % +\endisadelimML +% +\isatagML +\isacommand{ML{\isacharunderscore}val}\isamarkupfalse% +\ {\isacharverbatimopen}\ factorial\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharverbatimclose}\ \ \isanewline +\ \ \isacommand{ML{\isacharunderscore}command}\isamarkupfalse% +\ {\isacharverbatimopen}\ writeln\ {\isacharparenleft}string{\isacharunderscore}of{\isacharunderscore}int\ {\isacharparenleft}factorial\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharparenright}{\isacharparenright}\ {\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +\isanewline +% +\endisadelimML +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Compile-time context% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Whenever the ML compiler is invoked within Isabelle/Isar, the + formal context is passed as a thread-local reference variable. Thus + ML code may access the theory context during compilation, by reading + or writing the (local) theory under construction. Note that such + direct access to the compile-time context is rare; in practice it is + typically via some derived ML functions.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{ML\_Context.the\_generic\_context}\verb|ML_Context.the_generic_context: unit -> Context.generic| \\ + \indexdef{}{ML}{Context.$>$$>$}\verb|Context.>> : (Context.generic -> Context.generic) -> unit| \\ + \indexdef{}{ML}{bind\_thms}\verb|bind_thms: string * thm list -> unit| \\ + \indexdef{}{ML}{bind\_thm}\verb|bind_thm: string * thm -> unit| \\ + \end{mldecls} \begin{description} - \item[fundamental law of programming] - Whenever writing code, keep in mind: A program is - written once, modified ten times, and read - hundred times. So simplify its writing, - always keep future modifications in mind, - and never jeopardize readability. Every second you hesitate - to spend on making your code more clear you will - have to spend ten times understanding what you have - written later on. + \item \verb|ML_Context.the_generic_context ()| refers to the theory + context of the ML toplevel --- at compile time. ML code needs to + take care to refer to \verb|ML_Context.the_generic_context ()| + correctly. Recall that evaluation of a function body is delayed + until actual run-time. + + \item \verb|Context.>>|~\isa{f} applies context transformation + \isa{f} to the implicit context of the ML toplevel. - \item[white space matters] - Treat white space in your code as if it determines - the meaning of code. + \item \verb|bind_thms|~\isa{{\isacharparenleft}name{\isacharcomma}\ thms{\isacharparenright}} stores a list of + theorems produced in ML both in the (global) theory context and the + ML toplevel, associating it with the provided name. Theorems are + put into a global ``standard'' format before being stored. - \begin{itemize} + \item \verb|bind_thm| is similar to \verb|bind_thms| but refers to a + singleton theorem. + + \end{description} - \item The space bar is the easiest key to find on the keyboard, - press it as often as necessary. \verb|2 + 2| is better - than \verb|2+2|, likewise \verb|f (x, y)| is - better than \verb|f(x,y)|. + It is very important to note that the above functions are really + restricted to the compile time, even though the ML compiler is + invoked at run-time. The majority of ML code either uses static + antiquotations (\secref{sec:ML-antiq}) or refers to the theory or + proof context at run-time, by explicit functional abstraction.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsubsection{Antiquotations \label{sec:ML-antiq}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +A very important consequence of embedding SML into Isar is the + concept of \emph{ML antiquotation}. First, the standard token + language of ML is augmented by special syntactic entities of the + following form: - \item Restrict your lines to 80 characters. This will allow - you to keep the beginning of a line in view while watching - its end.\footnote{To acknowledge the lax practice of - text editing these days, we tolerate as much as 100 - characters per line, but anything beyond 120 is not - considered proper source text.} - - \item Ban tabulators; they are a context-sensitive formatting - feature and likely to confuse anyone not using your favorite - editor.\footnote{Some modern programming language even - forbid tabulators altogether according to the formal syntax - definition.} - - \item Get rid of trailing whitespace. Instead, do not - suppress a trailing newline at the end of your files. + \indexouternonterm{antiquote} + \begin{rail} + antiquote: atsign lbrace nameref args rbrace | lbracesym | rbracesym + ; + \end{rail} - \item Choose a generally accepted style of indentation, - then use it systematically throughout the whole - application. An indentation of two spaces is appropriate. - Avoid dangling indentation. - - \end{itemize} + Here \hyperlink{syntax.nameref}{\mbox{\isa{nameref}}} and \hyperlink{syntax.args}{\mbox{\isa{args}}} are regular outer syntax + categories. Note that attributes and proof methods use similar + syntax. - \item[cut-and-paste succeeds over copy-and-paste] - \emph{Never} copy-and-paste code when programming. If you - need the same piece of code twice, introduce a - reasonable auxiliary function (if there is no - such function, very likely you got something wrong). - Any copy-and-paste will turn out to be painful - when something has to be changed later on. + \medskip A regular antiquotation \isa{{\isacharat}{\isacharbraceleft}name\ args{\isacharbraceright}} processes + its arguments by the usual means of the Isar source language, and + produces corresponding ML source text, either as literal + \emph{inline} text (e.g. \isa{{\isacharat}{\isacharbraceleft}term\ t{\isacharbraceright}}) or abstract + \emph{value} (e.g. \isa{{\isacharat}{\isacharbraceleft}thm\ th{\isacharbraceright}}). This pre-compilation + scheme allows to refer to formal entities in a robust manner, with + proper static scoping and with some degree of logical checking of + small portions of the code. - \item[comments] - are a device which requires careful thinking before using - it. The best comment for your code should be the code itself. - Prefer efforts to write clear, understandable code - over efforts to explain nasty code. + Special antiquotations like \isa{{\isacharat}{\isacharbraceleft}let\ {\isasymdots}{\isacharbraceright}} or \isa{{\isacharat}{\isacharbraceleft}note\ {\isasymdots}{\isacharbraceright}} augment the compilation context without generating code. The + non-ASCII braces \isa{{\isasymlbrace}} and \isa{{\isasymrbrace}} allow to delimit the + effect by introducing local blocks within the pre-compilation + environment. - \item[functional programming is based on functions] - Model things as abstract as appropriate. Avoid unnecessarrily - concrete datatype representations. For example, consider a function - taking some table data structure as argument and performing - lookups on it. Instead of taking a table, it could likewise - take just a lookup function. Accustom - your way of coding to the level of expressiveness a functional - programming language is giving onto you. + \medskip See also \cite{Wenzel-Chaieb:2007b} for a slightly broader + perspective on Isabelle/ML antiquotations.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlantiq +% +\endisadelimmlantiq +% +\isatagmlantiq +% +\begin{isamarkuptext}% +\begin{matharray}{rcl} + \indexdef{}{ML antiquotation}{let}\hypertarget{ML antiquotation.let}{\hyperlink{ML antiquotation.let}{\mbox{\isa{let}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{note}\hypertarget{ML antiquotation.note}{\hyperlink{ML antiquotation.note}{\mbox{\isa{note}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \end{matharray} - \item[tuples] - are often in the way. When there is no striking argument - to tuple function arguments, just write your function curried. + \begin{rail} + 'let' ((term + 'and') '=' term + 'and') + ; + + 'note' (thmdef? thmrefs + 'and') + ; + \end{rail} + + \begin{description} - \item[telling names] - Any name should tell its purpose as exactly as possible, while - keeping its length to the absolutely necessary minimum. Always - give the same name to function arguments which have the same - meaning. Separate words by underscores (\verb|int_of_string|, not \verb|intOfString|).\footnote{Some - recent tools for Emacs include special precautions to cope with - bumpy names in \isa{camelCase}, e.g.\ for improved on-screen - readability. It is easier to abstain from using such names in the - first place.} + \item \isa{{\isacharat}{\isacharbraceleft}let\ p\ {\isacharequal}\ t{\isacharbraceright}} binds schematic variables in the + pattern \isa{p} by higher-order matching against the term \isa{t}. This is analogous to the regular \indexref{}{command}{let}\hyperlink{command.let}{\mbox{\isa{\isacommand{let}}}} command + in the Isar proof language. The pre-compilation environment is + augmented by auxiliary term bindings, without emitting ML source. + + \item \isa{{\isacharat}{\isacharbraceleft}note\ a\ {\isacharequal}\ b\isactrlsub {\isadigit{1}}\ {\isasymdots}\ b\isactrlsub n{\isacharbraceright}} recalls existing facts \isa{b\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ b\isactrlsub n}, binding the result as \isa{a}. This is analogous to + the regular \indexref{}{command}{note}\hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}} command in the Isar proof language. + The pre-compilation environment is augmented by auxiliary fact + bindings, without emitting ML source. \end{description}% \end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsection{Thread-safe programming% +\endisatagmlantiq +{\isafoldmlantiq}% +% +\isadelimmlantiq +% +\endisadelimmlantiq +% +\isadelimmlex +% +\endisadelimmlex +% +\isatagmlex +% +\begin{isamarkuptext}% +The following artificial example gives a first + impression about using the antiquotation elements introduced so far, + together with the basic \isa{{\isacharat}{\isacharbraceleft}thm{\isacharbraceright}} antiquotation defined + later.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlex +{\isafoldmlex}% +% +\isadelimmlex +% +\endisadelimmlex +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ {\isaantiqopen}\isanewline +\ \ \ \ % +\isaantiq +let\ {\isacharquery}t\ {\isacharequal}\ my{\isacharunderscore}term% +\endisaantiq +\isanewline +\ \ \ \ % +\isaantiq +note\ my{\isacharunderscore}refl\ {\isacharequal}\ reflexive\ {\isacharbrackleft}of\ {\isacharquery}t{\isacharbrackright}% +\endisaantiq +\isanewline +\ \ \ \ fun\ foo\ th\ {\isacharequal}\ Thm{\isachardot}transitive\ th\ % +\isaantiq +thm\ my{\isacharunderscore}refl% +\endisaantiq +\isanewline +\ \ {\isaantiqclose}\isanewline +{\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +The extra block delimiters do not affect the compiled code itself, i.e.\ + function \verb|foo| is available in the present context.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Canonical argument order \label{sec:canonical-argument-order}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Standard ML is a language in the tradition of \isa{{\isasymlambda}}-calculus and \emph{higher-order functional programming}, + similar to OCaml, Haskell, or Isabelle/Pure and HOL as logical + languages. Getting acquainted with the native style of representing + functions in that setting can save a lot of extra boiler-plate of + redundant shuffling of arguments, auxiliary abstractions etc. + + First of all, functions are usually \emph{curried}: the idea of + \isa{f} taking \isa{n} arguments of type \isa{{\isasymtau}\isactrlsub i} (for + \isa{i\ {\isasymin}\ {\isacharbraceleft}{\isadigit{1}}{\isacharcomma}\ {\isasymdots}\ n{\isacharbraceright}}) with result \isa{{\isasymtau}} is represented by + the iterated function space \isa{{\isasymtau}\isactrlsub {\isadigit{1}}\ {\isasymrightarrow}\ {\isasymdots}\ {\isasymrightarrow}\ {\isasymtau}\isactrlsub n\ {\isasymrightarrow}\ {\isasymtau}}. This is + isomorphic to the encoding via tuples \isa{{\isasymtau}\isactrlsub {\isadigit{1}}\ {\isasymtimes}\ {\isasymdots}\ {\isasymtimes}\ {\isasymtau}\isactrlsub n\ {\isasymrightarrow}\ {\isasymtau}}, but + the curried version fits more smoothly into the basic + calculus.\footnote{The difference is even more significant in + higher-order logic, because additional the redundant tuple structure + needs to be accommodated by formal reasoning.} + + Currying gives some flexiblity due to \emph{partial application}. A + function \isa{f{\isacharcolon}\ {\isasymtau}\isactrlsub {\isadigit{1}}\ {\isasymrightarrow}\ {\isasymtau}\isactrlbsub {\isadigit{2}}\isactrlesub \ {\isasymrightarrow}\ {\isasymtau}} can be applied to \isa{x{\isacharcolon}\ {\isasymtau}\isactrlsub {\isadigit{1}}} + and the remaining \isa{{\isacharparenleft}f\ x{\isacharparenright}{\isacharcolon}\ {\isasymtau}\isactrlsub {\isadigit{2}}\ {\isasymrightarrow}\ {\isasymtau}} passed to other functions + etc. How well this works in practice depends on the order of + arguments. In the worst case, arguments are arranged erratically, + and using a function in a certain situation always requires some + glue code. Thus we would get exponentially many oppurtunities to + decorate the code with meaningless permutations of arguments. + + This can be avoided by \emph{canonical argument order}, which + observes certain standard patterns of function definitions and + minimizes adhoc permutations in their application. In Isabelle/ML, + large portions of non-trivial source code are written without ever + using the infamous function \isa{swap{\isacharcolon}\ {\isasymalpha}\ {\isasymtimes}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}\ {\isasymtimes}\ {\isasymalpha}} or the + combinator \isa{C{\isacharcolon}\ {\isacharparenleft}{\isasymalpha}\ {\isasymrightarrow}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymgamma}{\isacharparenright}\ {\isasymrightarrow}\ {\isacharparenleft}{\isasymbeta}\ {\isasymrightarrow}\ {\isasymalpha}\ {\isasymrightarrow}\ {\isasymgamma}{\isacharparenright}}, which is not even + defined in our library. + + \medskip The basic idea is that arguments that vary less are moved + further to the left than those that vary more. Two particularly + important categories of functions are \emph{selectors} and + \emph{updates}. + + The subsequent scheme is based on a hypothetical set-like container + of type \isa{{\isasymbeta}} that manages elements of type \isa{{\isasymalpha}}. Both + the names and types of the associated operations are canonical for + Isabelle/ML. + + \medskip + \begin{tabular}{ll} + kind & canonical name and type \\\hline + selector & \isa{member{\isacharcolon}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymalpha}\ {\isasymrightarrow}\ bool} \\ + update & \isa{insert{\isacharcolon}\ {\isasymalpha}\ {\isasymrightarrow}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}} \\ + \end{tabular} + \medskip + + Given a container \isa{B{\isacharcolon}\ {\isasymbeta}}, the partially applied \isa{member\ B} is a predicate over elements \isa{{\isasymalpha}\ {\isasymrightarrow}\ bool}, and + thus represents the intended denotation directly. It is customary + to pass the abstract predicate to further operations, not the + concrete container. The argument order makes it easy to use other + combinators: \isa{forall\ {\isacharparenleft}member\ B{\isacharparenright}\ list} will check a list of + elements for membership in \isa{B} etc. Often the explicit + \isa{list} is pointless and can be contracted in the application + \isa{forall\ {\isacharparenleft}member\ B{\isacharparenright}} to get directly a predicate again. + + The update operation naturally ``varies'' the container, so it moves + to the right: \isa{insert\ a} is a function \isa{{\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}} to + insert a value \isa{a}. These can be composed naturally as + follows: \isa{insert\ c\ {\isasymcirc}\ insert\ b\ {\isasymcirc}\ insert\ a}. This works well, + apart from some awkwardness due to conventional mathematical + function composition, which can be easily amended as follows.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Forward application and composition% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Regular function application and infix notation works best for + relatively deeply structured expressions, e.g.\ \isa{h\ {\isacharparenleft}f\ x\ y\ {\isacharplus}\ g\ z{\isacharparenright}}. The important special case if \emph{linear transformation} + applies a cascade of functions as follows: \isa{f\isactrlsub n\ {\isacharparenleft}{\isasymdots}\ {\isacharparenleft}f\isactrlsub {\isadigit{1}}\ x{\isacharparenright}{\isacharparenright}}. + This becomes hard to read and maintain if the functions are + themselves complex expressions. The notation can be significantly + improved by introducing \emph{forward} versions of application and + composition as follows: + + \medskip + \begin{tabular}{lll} + \isa{x\ {\isacharbar}{\isachargreater}\ f} & \isa{{\isasymequiv}} & \isa{f\ x} \\ + \isa{f\ {\isacharhash}{\isachargreater}\ g} & \isa{{\isasymequiv}} & \isa{x\ {\isacharbar}{\isachargreater}\ f\ {\isacharbar}{\isachargreater}\ g} \\ + \end{tabular} + \medskip + + This enables to write conveniently \isa{x\ {\isacharbar}{\isachargreater}\ f\isactrlsub {\isadigit{1}}\ {\isacharbar}{\isachargreater}\ {\isasymdots}\ {\isacharbar}{\isachargreater}\ f\isactrlsub n} or + \isa{f\isactrlsub {\isadigit{1}}\ {\isacharhash}{\isachargreater}\ {\isasymdots}\ {\isacharhash}{\isachargreater}\ f\isactrlsub n} for its functional abstraction over \isa{x}. + + \medskip There is an additional set of combinators to accommodate + multiple results (via pairs) that are passed on as multiple + arguments (via currying). + + \medskip + \begin{tabular}{lll} + \isa{{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isacharbar}{\isacharminus}{\isachargreater}\ f} & \isa{{\isasymequiv}} & \isa{f\ x\ y} \\ + \isa{f\ {\isacharhash}{\isacharminus}{\isachargreater}\ g} & \isa{{\isasymequiv}} & \isa{x\ {\isacharbar}{\isachargreater}\ f\ {\isacharbar}{\isacharminus}{\isachargreater}\ g} \\ + \end{tabular} + \medskip% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{$\mid$$>$}\verb|op |\verb,|,\verb|> : 'a * ('a -> 'b) -> 'b| \\ + \indexdef{}{ML}{$\mid$-$>$}\verb|op |\verb,|,\verb|-> : ('c * 'a) * ('c -> 'a -> 'b) -> 'b| \\ + \indexdef{}{ML}{\#$>$}\verb|op #> : ('a -> 'b) * ('b -> 'c) -> 'a -> 'c| \\ + \indexdef{}{ML}{\#-$>$}\verb|op #-> : ('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd| \\ + \end{mldecls} + + %FIXME description!?% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsubsection{Canonical iteration% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +As explained above, a function \isa{f{\isacharcolon}\ {\isasymalpha}\ {\isasymrightarrow}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}} can be + understood as update on a configuration of type \isa{{\isasymbeta}} that is + parametrized by arguments of type \isa{{\isasymalpha}}. Given \isa{a{\isacharcolon}\ {\isasymalpha}} + the partial application \isa{{\isacharparenleft}f\ a{\isacharparenright}{\isacharcolon}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}} operates + homogeneously on \isa{{\isasymbeta}}. This can be iterated naturally over a + list of parameters \isa{{\isacharbrackleft}a\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ a\isactrlsub n{\isacharbrackright}} as \isa{f\ a\isactrlsub {\isadigit{1}}\ {\isacharhash}{\isachargreater}\ {\isasymdots}\ {\isacharhash}{\isachargreater}\ f\ a\isactrlbsub n\isactrlesub \isactrlbsub \isactrlesub }. The latter expression is again a function \isa{{\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}}. + It can be applied to an initial configuration \isa{b{\isacharcolon}\ {\isasymbeta}} to + start the iteration over the given list of arguments: each \isa{a} in \isa{a\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ a\isactrlsub n} is applied consecutively by updating a + cumulative configuration. + + The \isa{fold} combinator in Isabelle/ML lifts a function \isa{f} as above to its iterated version over a list of arguments. + Lifting can be repeated, e.g.\ \isa{{\isacharparenleft}fold\ {\isasymcirc}\ fold{\isacharparenright}\ f} iterates + over a list of lists as expected. + + The variant \isa{fold{\isacharunderscore}rev} works inside-out over the list of + arguments, such that \isa{fold{\isacharunderscore}rev\ f\ {\isasymequiv}\ fold\ f\ {\isasymcirc}\ rev} holds. + + The \isa{fold{\isacharunderscore}map} combinator essentially performs \isa{fold} and \isa{map} simultaneously: each application of \isa{f} produces an updated configuration together with a side-result; + the iteration collects all such side-results as a separate list.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{fold}\verb|fold: ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b| \\ + \indexdef{}{ML}{fold\_rev}\verb|fold_rev: ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b| \\ + \indexdef{}{ML}{fold\_map}\verb|fold_map: ('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b| \\ + \end{mldecls} + + \begin{description} + + \item \verb|fold|~\isa{f} lifts the parametrized update function + \isa{f} to a list of parameters. + + \item \verb|fold_rev|~\isa{f} is similar to \verb|fold|~\isa{f}, but works inside-out. + + \item \verb|fold_map|~\isa{f} lifts the parametrized update + function \isa{f} (with side-result) to a list of parameters and + cumulative side-results. + + \end{description} + + \begin{warn} + The literature on functional programming provides a multitude of + combinators called \isa{foldl}, \isa{foldr} etc. SML97 + provides its own variations as \verb|List.foldl| and \verb|List.foldr|, while the classic Isabelle library still has the + slightly more convenient historic \verb|Library.foldl| and \verb|Library.foldr|. To avoid further confusion, all of this should be + ignored, and \verb|fold| (or \verb|fold_rev|) used exclusively. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isadelimmlex +% +\endisadelimmlex +% +\isatagmlex +% +\begin{isamarkuptext}% +The following example shows how to fill a text buffer + incrementally by adding strings, either individually or from a given + list.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlex +{\isafoldmlex}% +% +\isadelimmlex +% +\endisadelimmlex +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ val\ s\ {\isacharequal}\isanewline +\ \ \ \ Buffer{\isachardot}empty\isanewline +\ \ \ \ {\isacharbar}{\isachargreater}\ Buffer{\isachardot}add\ {\isachardoublequote}digits{\isacharcolon}\ {\isachardoublequote}\isanewline +\ \ \ \ {\isacharbar}{\isachargreater}\ fold\ {\isacharparenleft}Buffer{\isachardot}add\ o\ string{\isacharunderscore}of{\isacharunderscore}int{\isacharparenright}\ {\isacharparenleft}{\isadigit{0}}\ upto\ {\isadigit{9}}{\isacharparenright}\isanewline +\ \ \ \ {\isacharbar}{\isachargreater}\ Buffer{\isachardot}content{\isacharsemicolon}\isanewline +\isanewline +\ \ % +\isaantiq +assert% +\endisaantiq +\ {\isacharparenleft}s\ {\isacharequal}\ {\isachardoublequote}digits{\isacharcolon}\ {\isadigit{0}}{\isadigit{1}}{\isadigit{2}}{\isadigit{3}}{\isadigit{4}}{\isadigit{5}}{\isadigit{6}}{\isadigit{7}}{\isadigit{8}}{\isadigit{9}}{\isachardoublequote}{\isacharparenright}{\isacharsemicolon}\isanewline +{\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +Note how \verb|fold (Buffer.add o string_of_int)| above saves + an extra \verb|map| over the given list. This kind of peephole + optimization reduces both the code size and the tree structures in + memory (``deforestation''), but requires some practice to read and + write it fluently. + + \medskip The next example elaborates this idea and demonstrates fast + accumulation of tree content using a text buffer.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ datatype\ tree\ {\isacharequal}\ Text\ of\ string\ {\isacharbar}\ Elem\ of\ string\ {\isacharasterisk}\ tree\ list{\isacharsemicolon}\isanewline +\isanewline +\ \ fun\ slow{\isacharunderscore}content\ {\isacharparenleft}Text\ txt{\isacharparenright}\ {\isacharequal}\ txt\isanewline +\ \ \ \ {\isacharbar}\ slow{\isacharunderscore}content\ {\isacharparenleft}Elem\ {\isacharparenleft}name{\isacharcomma}\ ts{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ \ \ \ {\isachardoublequote}{\isacharless}{\isachardoublequote}\ {\isacharcircum}\ name\ {\isacharcircum}\ {\isachardoublequote}{\isachargreater}{\isachardoublequote}\ {\isacharcircum}\isanewline +\ \ \ \ \ \ \ \ implode\ {\isacharparenleft}map\ slow{\isacharunderscore}content\ ts{\isacharparenright}\ {\isacharcircum}\isanewline +\ \ \ \ \ \ \ \ {\isachardoublequote}{\isacharless}{\isacharslash}{\isachardoublequote}\ {\isacharcircum}\ name\ {\isacharcircum}\ {\isachardoublequote}{\isachargreater}{\isachardoublequote}\isanewline +\isanewline +\ \ fun\ add{\isacharunderscore}content\ {\isacharparenleft}Text\ txt{\isacharparenright}\ {\isacharequal}\ Buffer{\isachardot}add\ txt\isanewline +\ \ \ \ {\isacharbar}\ add{\isacharunderscore}content\ {\isacharparenleft}Elem\ {\isacharparenleft}name{\isacharcomma}\ ts{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ \ \ \ Buffer{\isachardot}add\ {\isacharparenleft}{\isachardoublequote}{\isacharless}{\isachardoublequote}\ {\isacharcircum}\ name\ {\isacharcircum}\ {\isachardoublequote}{\isachargreater}{\isachardoublequote}{\isacharparenright}\ {\isacharhash}{\isachargreater}\isanewline +\ \ \ \ \ \ \ \ fold\ add{\isacharunderscore}content\ ts\ {\isacharhash}{\isachargreater}\isanewline +\ \ \ \ \ \ \ \ Buffer{\isachardot}add\ {\isacharparenleft}{\isachardoublequote}{\isacharless}{\isacharslash}{\isachardoublequote}\ {\isacharcircum}\ name\ {\isacharcircum}\ {\isachardoublequote}{\isachargreater}{\isachardoublequote}{\isacharparenright}{\isacharsemicolon}\isanewline +\isanewline +\ \ fun\ fast{\isacharunderscore}content\ tree\ {\isacharequal}\isanewline +\ \ \ \ Buffer{\isachardot}empty\ {\isacharbar}{\isachargreater}\ add{\isacharunderscore}content\ tree\ {\isacharbar}{\isachargreater}\ Buffer{\isachardot}content{\isacharsemicolon}\isanewline +{\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +The slow part of \verb|slow_content| is the \verb|implode| of + the recursive results, because it copies previously produced strings + afresh. + + The incremental \verb|add_content| avoids this by operating on a + buffer that is passed-through in a linear fashion. Using \verb|#>| and contraction over the actual \verb|Buffer.T| argument + saves some additional boiler-plate, but requires again some + practice. Of course, the two \verb|Buffer.add| invocations with + concatenated strings could have been split into smaller parts, but + this would have obfuscated the source without making a big + difference in allocations. Here we have done some + peephole-optimization for the sake of readability. + + Another benefit of \verb|add_content| is its ``open'' form as a + function on \verb|Buffer.T| that can be continued in further + linear transformations, folding etc. Thus it is more compositional + than the naive \verb|slow_content|. As a realistic example, compare + the slightly old-style \verb|Term.maxidx_of_term: term -> int| with + the newer \verb|Term.maxidx_term: term -> int -> int| in + Isabelle/Pure. + + Note that \verb|fast_content| above is only defined for completeness + of the example. In many practical situations, it is customary to + defined the incremental \verb|add_content| only and leave the + initialization and termination to the concrete application by the + user.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Message output channels \label{sec:message-channels}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isabelle provides output channels for different kinds of + messages: regular output, high-volume tracing information, warnings, + and errors. + + Depending on the user interface involved, these messages may appear + in different text styles or colours. The standard output for + terminal sessions prefixes each line of warnings by \verb|###| and errors by \verb|***|, but leaves anything else + unchanged. + + Messages are associated with the transaction context of the running + Isar command. This enables the front-end to manage commands and + resulting messages together. For example, after deleting a command + from a given theory document version, the corresponding message + output can be retracted from the display.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{writeln}\verb|writeln: string -> unit| \\ + \indexdef{}{ML}{tracing}\verb|tracing: string -> unit| \\ + \indexdef{}{ML}{warning}\verb|warning: string -> unit| \\ + \indexdef{}{ML}{error}\verb|error: string -> 'a| \\ + \end{mldecls} + + \begin{description} + + \item \verb|writeln|~\isa{text} outputs \isa{text} as regular + message. This is the primary message output operation of Isabelle + and should be used by default. + + \item \verb|tracing|~\isa{text} outputs \isa{text} as special + tracing message, indicating potential high-volume output to the + front-end (hundreds or thousands of messages issued by a single + command). The idea is to allow the user-interface to downgrade the + quality of message display to achieve higher throughput. + + Note that the user might have to take special actions to see tracing + output, e.g.\ switch to a different output window. So this channel + should not be used for regular output. + + \item \verb|warning|~\isa{text} outputs \isa{text} as + warning, which typically means some extra emphasis on the front-end + side (color highlighting, icon). + + \item \verb|error|~\isa{text} raises exception \verb|ERROR|~\isa{text} and thus lets the Isar toplevel print \isa{text} on the + error channel, which typically means some extra emphasis on the + front-end side (color highlighting, icon). + + This assumes that the exception is not handled before the command + terminates. Handling exception \verb|ERROR|~\isa{text} is a + perfectly legal alternative: it means that the error is absorbed + without any message output. + + \begin{warn} + The actual error channel is accessed via \verb|Output.error_msg|, but + the interaction protocol of Proof~General \emph{crashes} if that + function is used in regular ML code: error output and toplevel + command failure always need to coincide here. + \end{warn} + + \end{description} + + \begin{warn} + Regular Isabelle/ML code should output messages exclusively by the + official channels. Using raw I/O on \emph{stdout} or \emph{stderr} + instead (e.g.\ via \verb|TextIO.output|) is apt to cause problems in + the presence of parallel and asynchronous processing of Isabelle + theories. Such raw output might be displayed by the front-end in + some system console log, with a low chance that the user will ever + see it. Moreover, as a genuine side-effect on global process + channels, there is no proper way to retract output when Isar command + transactions are reset. + \end{warn} + + \begin{warn} + The message channels should be used in a message-oriented manner. + This means that multi-line output that logically belongs together + needs to be issued by a \emph{single} invocation of \verb|writeln| + etc. with the functional concatenation of all message constituents. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isadelimmlex +% +\endisadelimmlex +% +\isatagmlex +% +\begin{isamarkuptext}% +The following example demonstrates a multi-line + warning. Note that in some situations the user sees only the first + line, so the most important point should be made first.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlex +{\isafoldmlex}% +% +\isadelimmlex +% +\endisadelimmlex +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML{\isacharunderscore}command}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ warning\ {\isacharparenleft}cat{\isacharunderscore}lines\isanewline +\ \ \ {\isacharbrackleft}{\isachardoublequote}Beware\ the\ Jabberwock{\isacharcomma}\ my\ son{\isacharbang}{\isachardoublequote}{\isacharcomma}\isanewline +\ \ \ \ {\isachardoublequote}The\ jaws\ that\ bite{\isacharcomma}\ the\ claws\ that\ catch{\isacharbang}{\isachardoublequote}{\isacharcomma}\isanewline +\ \ \ \ {\isachardoublequote}Beware\ the\ Jubjub\ Bird{\isacharcomma}\ and\ shun{\isachardoublequote}{\isacharcomma}\isanewline +\ \ \ \ {\isachardoublequote}The\ frumious\ Bandersnatch{\isacharbang}{\isachardoublequote}{\isacharbrackright}{\isacharparenright}{\isacharsemicolon}\isanewline +{\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\isamarkupsection{Exceptions \label{sec:exceptions}% } \isamarkuptrue% % \begin{isamarkuptext}% -Recent versions of Poly/ML (5.2.1 or later) support robust - multithreaded execution, based on native operating system threads of - the underlying platform. Thus threads will actually be executed in - parallel on multi-core systems. A speedup-factor of approximately - 1.5--3 can be expected on a regular 4-core machine.\footnote{There - is some inherent limitation of the speedup factor due to garbage - collection, which is still sequential. It helps to provide initial - heap space generously, using the \texttt{-H} option of Poly/ML.} - Threads also help to organize advanced operations of the system, - with explicit communication between sub-components, real-time - conditions, time-outs etc. +The Standard ML semantics of strict functional evaluation + together with exceptions is rather well defined, but some delicate + points need to be observed to avoid that ML programs go wrong + despite static type-checking. Exceptions in Isabelle/ML are + subsequently categorized as follows. + + \paragraph{Regular user errors.} These are meant to provide + informative feedback about malformed input etc. + + The \emph{error} function raises the corresponding \emph{ERROR} + exception, with a plain text message as argument. \emph{ERROR} + exceptions can be handled internally, in order to be ignored, turned + into other exceptions, or cascaded by appending messages. If the + corresponding Isabelle/Isar command terminates with an \emph{ERROR} + exception state, the toplevel will print the result on the error + channel (see \secref{sec:message-channels}). + + It is considered bad style to refer to internal function names or + values in ML source notation in user error messages. + + Grammatical correctness of error messages can be improved by + \emph{omitting} final punctuation: messages are often concatenated + or put into a larger context (e.g.\ augmented with source position). + By not insisting in the final word at the origin of the error, the + system can perform its administrative tasks more easily and + robustly. + + \paragraph{Program failures.} There is a handful of standard + exceptions that indicate general failure situations, or failures of + core operations on logical entities (types, terms, theorems, + theories, see \chref{ch:logic}). + + These exceptions indicate a genuine breakdown of the program, so the + main purpose is to determine quickly what has happened where. + Traditionally, the (short) exception message would include the name + of an ML function, although this no longer necessary, because the ML + runtime system prints a detailed source position of the + corresponding \verb|raise| keyword. + + \medskip User modules can always introduce their own custom + exceptions locally, e.g.\ to organize internal failures robustly + without overlapping with existing exceptions. Exceptions that are + exposed in module signatures require extra care, though, and should + \emph{not} be introduced by default. Surprise by end-users or ML + users of a module can be often minimized by using plain user errors. + + \paragraph{Interrupts.} These indicate arbitrary system events: + both the ML runtime system and the Isabelle/ML infrastructure signal + various exceptional situations by raising the special + \emph{Interrupt} exception in user code. + + This is the one and only way that physical events can intrude an + Isabelle/ML program. Such an interrupt can mean out-of-memory, + stack overflow, timeout, internal signaling of threads, or the user + producing a console interrupt manually etc. An Isabelle/ML program + that intercepts interrupts becomes dependent on physical effects of + the environment. Even worse, exception handling patterns that are + too general by accident, e.g.\ by mispelled exception constructors, + will cover interrupts unintentionally, and thus render the program + semantics ill-defined. + + Note that the Interrupt exception dates back to the original SML90 + language definition. It was excluded from the SML97 version to + avoid its malign impact on ML program semantics, but without + providing a viable alternative. Isabelle/ML recovers physical + interruptibility (which an indispensable tool to implement managed + evaluation of Isar command transactions), but requires user code to + be strictly transparent wrt.\ interrupts. + + \begin{warn} + Isabelle/ML user code needs to terminate promptly on interruption, + without guessing at its meaning to the system infrastructure. + Temporary handling of interrupts for cleanup of global resources + etc.\ needs to be followed immediately by re-raising of the original + exception. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{try}\verb|try: ('a -> 'b) -> 'a -> 'b option| \\ + \indexdef{}{ML}{can}\verb|can: ('a -> 'b) -> 'a -> bool| \\ + \indexdef{}{ML}{ERROR}\verb|ERROR: string -> exn| \\ + \indexdef{}{ML}{Fail}\verb|Fail: string -> exn| \\ + \indexdef{}{ML}{Exn.is\_interrupt}\verb|Exn.is_interrupt: exn -> bool| \\ + \indexdef{}{ML}{reraise}\verb|reraise: exn -> 'a| \\ + \indexdef{}{ML}{exception\_trace}\verb|exception_trace: (unit -> 'a) -> 'a| \\ + \end{mldecls} + + \begin{description} + + \item \verb|try|~\isa{f\ x} makes the partiality of evaluating + \isa{f\ x} explicit via the option datatype. Interrupts are + \emph{not} handled here, i.e.\ this form serves as safe replacement + for the \emph{unsafe} version \verb|(SOME|~\isa{f\ x}~\verb|handle _ => NONE)| that is occasionally seen in + books about SML. + + \item \verb|can| is similar to \verb|try| with more abstract result. + + \item \verb|ERROR|~\isa{msg} represents user errors; this + exception is always raised via the \verb|error| function (see + \secref{sec:message-channels}). + + \item \verb|Fail|~\isa{msg} represents general program failures. + + \item \verb|Exn.is_interrupt| identifies interrupts robustly, without + mentioning concrete exception constructors in user code. Handled + interrupts need to be re-raised promptly! + + \item \verb|reraise|~\isa{exn} raises exception \isa{exn} + while preserving its implicit position information (if possible, + depending on the ML platform). + + \item \verb|exception_trace|~\verb|(fn _ =>|~\isa{e}\verb|)| evaluates expression \isa{e} while printing + a full trace of its stack of nested exceptions (if possible, + depending on the ML platform).\footnote{In various versions of + Poly/ML the trace will appear on raw stdout of the Isabelle + process.} + + Inserting \verb|exception_trace| into ML code temporarily is useful + for debugging, but not suitable for production code. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isadelimmlantiq +% +\endisadelimmlantiq +% +\isatagmlantiq +% +\begin{isamarkuptext}% +\begin{matharray}{rcl} + \indexdef{}{ML antiquotation}{assert}\hypertarget{ML antiquotation.assert}{\hyperlink{ML antiquotation.assert}{\mbox{\isa{assert}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \end{matharray} + + \begin{description} + + \item \isa{{\isacharat}{\isacharbraceleft}assert{\isacharbraceright}} inlines a function \isa{bool\ {\isasymRightarrow}\ unit} + that raises \verb|Fail| if the argument is \verb|false|. Due to + inlining the source position of failed assertions is included in the + error output. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlantiq +{\isafoldmlantiq}% +% +\isadelimmlantiq +% +\endisadelimmlantiq +% +\isamarkupsection{Basic data types% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The basis library proposal of SML97 need to be treated with + caution. Many of its operations simply do not fit with important + Isabelle/ML conventions (like ``canonical argument order'', see + \secref{sec:canonical-argument-order}), others can cause serious + problems with the parallel evaluation model of Isabelle/ML (such as + \verb|TextIO.print| or \verb|OS.Process.system|). + + Subsequently we give a brief overview of important operations on + basic ML data types.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Characters% +} +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML type}{char}\verb|type char| \\ + \end{mldecls} + + \begin{description} + + \item Type \verb|char| is \emph{not} used. The smallest textual + unit in Isabelle is represented a ``symbol'' (see + \secref{sec:symbols}). + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsubsection{Integers% +} +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML type}{int}\verb|type int| \\ + \end{mldecls} + + \begin{description} + + \item Type \verb|int| represents regular mathematical integers, + which are \emph{unbounded}. Overflow never happens in + practice.\footnote{The size limit for integer bit patterns in memory + is 64\,MB for 32-bit Poly/ML, and much higher for 64-bit systems.} + This works uniformly for all supported ML platforms (Poly/ML and + SML/NJ). + + Literal integers in ML text (e.g.\ \verb|123456789012345678901234567890|) are forced to be of this one true + integer type --- overloading of SML97 is disabled. + + Structure \verb|IntInf| of SML97 is obsolete, always use + \verb|Int|. Structure \verb|Integer| in \hyperlink{file.~~/src/Pure/General/integer.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}General{\isacharslash}integer{\isachardot}ML}}}} provides some additional + operations. - Threads lack the memory protection of separate processes, and - operate concurrently on shared heap memory. This has the advantage - that results of independent computations are immediately available - to other threads, without requiring untyped character streams, - awkward serialization etc. + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsubsection{Options% +} +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{Option.map}\verb|Option.map: ('a -> 'b) -> 'a option -> 'b option| \\ + \indexdef{}{ML}{is\_some}\verb|is_some: 'a option -> bool| \\ + \indexdef{}{ML}{is\_none}\verb|is_none: 'a option -> bool| \\ + \indexdef{}{ML}{the}\verb|the: 'a option -> 'a| \\ + \indexdef{}{ML}{these}\verb|these: 'a list option -> 'a list| \\ + \indexdef{}{ML}{the\_list}\verb|the_list: 'a option -> 'a list| \\ + \indexdef{}{ML}{the\_default}\verb|the_default: 'a -> 'a option -> 'a| \\ + \end{mldecls}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\begin{isamarkuptext}% +Apart from \verb|Option.map| most operations defined in + structure \verb|Option| are alien to Isabelle/ML. The + operations shown above are defined in \hyperlink{file.~~/src/Pure/General/basics.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}General{\isacharslash}basics{\isachardot}ML}}}}, among others.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Lists% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Lists are ubiquitous in ML as simple and light-weight + ``collections'' for many everyday programming tasks. Isabelle/ML + provides important additions and improvements over operations that + are predefined in the SML97 library.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{cons}\verb|cons: 'a -> 'a list -> 'a list| \\ + \indexdef{}{ML}{member}\verb|member: ('b * 'a -> bool) -> 'a list -> 'b -> bool| \\ + \indexdef{}{ML}{insert}\verb|insert: ('a * 'a -> bool) -> 'a -> 'a list -> 'a list| \\ + \indexdef{}{ML}{remove}\verb|remove: ('b * 'a -> bool) -> 'b -> 'a list -> 'a list| \\ + \indexdef{}{ML}{update}\verb|update: ('a * 'a -> bool) -> 'a -> 'a list -> 'a list| \\ + \end{mldecls} + + \begin{description} + + \item \verb|cons|~\isa{x\ xs} evaluates to \isa{x\ {\isacharcolon}{\isacharcolon}\ xs}. + + Tupled infix operators are a historical accident in Standard ML. + The curried \verb|cons| amends this, but it should be only used when + partial application is required. + + \item \verb|member|, \verb|insert|, \verb|remove|, \verb|update| treat + lists as a set-like container that maintains the order of elements. + See \hyperlink{file.~~/src/Pure/library.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}library{\isachardot}ML}}}} for the full specifications + (written in ML). There are some further derived operations like + \verb|union| or \verb|inter|. + + Note that \verb|insert| is conservative about elements that are + already a \verb|member| of the list, while \verb|update| ensures that + the last entry is always put in front. The latter discipline is + often more appropriate in declarations of context data + (\secref{sec:context-data}) that are issued by the user in Isar + source: more recent declarations normally take precedence over + earlier ones. - On the other hand, some programming guidelines need to be observed - in order to make unprotected parallelism work out smoothly. While - the ML system implementation is responsible to maintain basic - integrity of the representation of ML values in memory, the - application programmer needs to ensure that multithreaded execution - does not break the intended semantics. + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isadelimmlex +% +\endisadelimmlex +% +\isatagmlex +% +\begin{isamarkuptext}% +Using canonical \verb|fold| together with canonical \verb|cons|, or similar standard operations, alternates the orientation of + data. The is quite natural and should not altered forcible by + inserting extra applications \verb|rev|. The alternative \verb|fold_rev| can be used in the relatively few situations, where + alternation should be prevented.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlex +{\isafoldmlex}% +% +\isadelimmlex +% +\endisadelimmlex +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ val\ items\ {\isacharequal}\ {\isacharbrackleft}{\isadigit{1}}{\isacharcomma}\ {\isadigit{2}}{\isacharcomma}\ {\isadigit{3}}{\isacharcomma}\ {\isadigit{4}}{\isacharcomma}\ {\isadigit{5}}{\isacharcomma}\ {\isadigit{6}}{\isacharcomma}\ {\isadigit{7}}{\isacharcomma}\ {\isadigit{8}}{\isacharcomma}\ {\isadigit{9}}{\isacharcomma}\ {\isadigit{1}}{\isadigit{0}}{\isacharbrackright}{\isacharsemicolon}\isanewline +\isanewline +\ \ val\ list{\isadigit{1}}\ {\isacharequal}\ fold\ cons\ items\ {\isacharbrackleft}{\isacharbrackright}{\isacharsemicolon}\isanewline +\ \ % +\isaantiq +assert% +\endisaantiq +\ {\isacharparenleft}list{\isadigit{1}}\ {\isacharequal}\ rev\ items{\isacharparenright}{\isacharsemicolon}\isanewline +\isanewline +\ \ val\ list{\isadigit{2}}\ {\isacharequal}\ fold{\isacharunderscore}rev\ cons\ items\ {\isacharbrackleft}{\isacharbrackright}{\isacharsemicolon}\isanewline +\ \ % +\isaantiq +assert% +\endisaantiq +\ {\isacharparenleft}list{\isadigit{2}}\ {\isacharequal}\ items{\isacharparenright}{\isacharsemicolon}\isanewline +{\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +The subsequent example demonstrates how to \emph{merge} two + lists in a natural way.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ fun\ merge{\isacharunderscore}lists\ eq\ {\isacharparenleft}xs{\isacharcomma}\ ys{\isacharparenright}\ {\isacharequal}\ fold{\isacharunderscore}rev\ {\isacharparenleft}insert\ eq{\isacharparenright}\ ys\ xs{\isacharsemicolon}\isanewline +{\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +Here the first list is treated conservatively: only the new + elements from the second list are inserted. The inside-out order of + insertion via \verb|fold_rev| attempts to preserve the order of + elements in the result. + + This way of merging lists is typical for context data + (\secref{sec:context-data}). See also \verb|merge| as defined in + \hyperlink{file.~~/src/Pure/library.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}library{\isachardot}ML}}}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Association lists% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The operations for association lists interpret a concrete list + of pairs as a finite function from keys to values. Redundant + representations with multiple occurrences of the same key are + implicitly normalized: lookup and update only take the first + occurrence into account.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{AList.lookup}\verb|AList.lookup: ('a * 'b -> bool) -> ('b * 'c) list -> 'a -> 'c option| \\ + \indexdef{}{ML}{AList.defined}\verb|AList.defined: ('a * 'b -> bool) -> ('b * 'c) list -> 'a -> bool| \\ + \indexdef{}{ML}{AList.update}\verb|AList.update: ('a * 'a -> bool) -> 'a * 'b -> ('a * 'b) list -> ('a * 'b) list| \\ + \end{mldecls} + + \begin{description} + + \item \verb|AList.lookup|, \verb|AList.defined|, \verb|AList.update| + implement the main ``framework operations'' for mappings in + Isabelle/ML, with standard conventions for names and types. + + Note that a function called \isa{lookup} is obliged to express its + partiality via an explicit option element. There is no choice to + raise an exception, without changing the name to something like + \isa{the{\isacharunderscore}element} or \isa{get}. + + The \isa{defined} operation is essentially a contraction of \verb|is_some| and \isa{lookup}, but this is sufficiently frequent to + justify its independent existence. This also gives the + implementation some opportunity for peep-hole optimization. + + \end{description} - \medskip \paragraph{Critical shared resources.} Actually only those - parts outside the purely functional world of ML are critical. In - particular, this covers + Association lists are adequate as simple and light-weight + implementation of finite mappings in many practical situations. A + more heavy-duty table structure is defined in \hyperlink{file.~~/src/Pure/General/table.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}General{\isacharslash}table{\isachardot}ML}}}}; that version scales easily to + thousands or millions of elements.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Unsynchronized references% +} +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML type}{Unsynchronized.ref}\verb|type 'a Unsynchronized.ref| \\ + \indexdef{}{ML}{Unsynchronized.ref}\verb|Unsynchronized.ref: 'a -> 'a Unsynchronized.ref| \\ + \indexdef{}{ML}{!}\verb|! : 'a Unsynchronized.ref -> 'a| \\ + \indexdef{}{ML}{:=}\verb|op := : 'a Unsynchronized.ref * 'a -> unit| \\ + \end{mldecls}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\begin{isamarkuptext}% +Due to ubiquitous parallelism in Isabelle/ML (see also + \secref{sec:multi-threading}), the mutable reference cells of + Standard ML are notorious for causing problems. In a highly + parallel system, both correctness \emph{and} performance are easily + degraded when using mutable data. + + The unwieldy name of \verb|Unsynchronized.ref| for the constructor + for references in Isabelle/ML emphasizes the inconveniences caused by + mutability. Existing operations \verb|!| and \verb|op :=| are + unchanged, but should be used with special precautions, say in a + strictly local situation that is guaranteed to be restricted to + sequential evaluation --- now and in the future.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Thread-safe programming \label{sec:multi-threading}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Multi-threaded execution has become an everyday reality in + Isabelle since Poly/ML 5.2.1 and Isabelle2008. Isabelle/ML provides + implicit and explicit parallelism by default, and there is no way + for user-space tools to ``opt out''. ML programs that are purely + functional, output messages only via the official channels + (\secref{sec:message-channels}), and do not intercept interrupts + (\secref{sec:exceptions}) can participate in the multi-threaded + environment immediately without further ado. + + More ambitious tools with more fine-grained interaction with the + environment need to observe the principles explained below.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Multi-threading with shared memory% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Multiple threads help to organize advanced operations of the + system, such as real-time conditions on command transactions, + sub-components with explicit communication, general asynchronous + interaction etc. Moreover, parallel evaluation is a prerequisite to + make adequate use of the CPU resources that are available on + multi-core systems.\footnote{Multi-core computing does not mean that + there are ``spare cycles'' to be wasted. It means that the + continued exponential speedup of CPU performance due to ``Moore's + Law'' follows different rules: clock frequency has reached its peak + around 2005, and applications need to be parallelized in order to + avoid a perceived loss of performance. See also + \cite{Sutter:2005}.} + + Isabelle/Isar exploits the inherent structure of theories and proofs + to support \emph{implicit parallelism} to a large extent. LCF-style + theorem provides almost ideal conditions for that; see also + \cite{Wenzel:2009}. This means, significant parts of theory and + proof checking is parallelized by default. A maximum speedup-factor + of 3.0 on 4 cores and 5.0 on 8 cores can be + expected.\footnote{Further scalability is limited due to garbage + collection, which is still sequential in Poly/ML 5.2/5.3/5.4. It + helps to provide initial heap space generously, using the + \texttt{-H} option. Initial heap size needs to be scaled-up + together with the number of CPU cores: approximately 1--2\,GB per + core..} + + \medskip ML threads lack the memory protection of separate + processes, and operate concurrently on shared heap memory. This has + the advantage that results of independent computations are + immediately available to other threads: abstract values can be + passed between threads without copying or awkward serialization that + is typically required for explicit message passing between separate + processes. + + To make shared-memory multi-threading work robustly and efficiently, + some programming guidelines need to be observed. While the ML + system is responsible to maintain basic integrity of the + representation of ML values in memory, the application programmer + needs to ensure that multi-threaded execution does not break the + intended semantics. + + \begin{warn} + To participate in implicit parallelism, tools need to be + thread-safe. A single ill-behaved tool can affect the stability and + performance of the whole system. + \end{warn} + + Apart from observing the principles of thread-safeness passively, + advanced tools may also exploit parallelism actively, e.g.\ by using + ``future values'' (\secref{sec:futures}) or the more basic library + functions for parallel list operations (\secref{sec:parlist}). + + \begin{warn} + Parallel computing resources are managed centrally by the + Isabelle/ML infrastructure. User programs must not fork their own + ML threads to perform computations. + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Critical shared resources% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Thread-safeness is mainly concerned about concurrent + read/write access to shared resources, which are outside the purely + functional world of ML. This covers the following in particular. \begin{itemize} - \item global references (or arrays), i.e.\ those that persist over - several invocations of associated operations,\footnote{This is - independent of the visibility of such mutable values in the toplevel - scope.} + \item Global references (or arrays), i.e.\ mutable memory cells that + persist over several invocations of associated + operations.\footnote{This is independent of the visibility of such + mutable values in the toplevel scope.} - \item direct I/O on shared channels, notably \isa{stdin}, \isa{stdout}, \isa{stderr}. + \item Global state of the running Isabelle/ML process, i.e.\ raw I/O + channels, environment variables, current working directory. + + \item Writable resources in the file-system that are shared among + different threads or other processes. \end{itemize} - The majority of tools implemented within the Isabelle/Isar framework - will not require any of these critical elements: nothing special - needs to be observed when staying in the purely functional fragment - of ML. Note that output via the official Isabelle channels does not - count as direct I/O, so the operations \verb|writeln|, \verb|warning|, \verb|tracing| etc.\ are safe. - - Moreover, ML bindings within the toplevel environment (\verb|type|, \verb|val|, \verb|structure| etc.) due to - run-time invocation of the compiler are also safe, because - Isabelle/Isar manages this as part of the theory or proof context. - - \paragraph{Multithreading in Isabelle/Isar.} The theory loader - automatically exploits the overall parallelism of independent nodes - in the development graph, as well as the inherent irrelevance of - proofs for goals being fully specified in advance. This means, - checking of individual Isar proofs is parallelized by default. - Beyond that, very sophisticated proof tools may use local - parallelism internally, via the general programming model of - ``future values'' (see also \hyperlink{file.~~/src/Pure/Concurrent/future.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Concurrent{\isacharslash}future{\isachardot}ML}}}}). - - Any ML code that works relatively to the present background theory - is already safe. Contextual data may be easily stored within the - theory or proof context, thanks to the generic data concept of - Isabelle/Isar (see \secref{sec:context-data}). This greatly - diminishes the demand for global state information in the first - place. - - \medskip In rare situations where actual mutable content needs to be - manipulated, Isabelle provides a single \emph{critical section} that - may be entered while preventing any other thread from doing the - same. Entering the critical section without contention is very - fast, and several basic system operations do so frequently. This - also means that each thread should leave the critical section - quickly, otherwise parallel execution performance may degrade - significantly. - - Despite this potential bottle-neck, centralized locking is - convenient, because it prevents deadlocks without demanding special - precautions. Explicit communication demands other means, though. - The high-level abstraction of synchronized variables \hyperlink{file.~~/src/Pure/Concurrent/synchronized.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Concurrent{\isacharslash}synchronized{\isachardot}ML}}}} enables parallel - components to communicate via shared state; see also \hyperlink{file.~~/src/Pure/Concurrent/mailbox.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Concurrent{\isacharslash}mailbox{\isachardot}ML}}}} as canonical example. - - \paragraph{Good conduct of impure programs.} The following - guidelines enable non-functional programs to participate in - multithreading. + Isabelle/ML provides various mechanisms to avoid critical shared + resources in most practical situations. As last resort there are + some mechanisms for explicit synchronization. The following + guidelines help to make Isabelle/ML programs work smoothly in a + concurrent environment. \begin{itemize} - \item Minimize global state information. Using proper theory and - proof context data will actually return to functional update of - values, without any special precautions for multithreading. Apart - from the fully general functors for theory and proof data (see - \secref{sec:context-data}) there are drop-in replacements that - emulate primitive references for common cases of \emph{configuration - options} for type \verb|bool|/\verb|int|/\verb|string| (see structure \verb|Config| and \verb|Attrib.config_bool| etc.), and lists of theorems (see functor - \verb|Named_Thms|). + \item Avoid global references altogether. Isabelle/Isar maintains a + uniform context that incorporates arbitrary data declared by user + programs (\secref{sec:context-data}). This context is passed as + plain value and user tools can get/map their own data in a purely + functional manner. Configuration options within the context + (\secref{sec:config-options}) provide simple drop-in replacements + for formerly writable flags. + + \item Keep components with local state information re-entrant. + Instead of poking initial values into (private) global references, a + new state record can be created on each invocation, and passed + through any auxiliary functions of the component. The state record + may well contain mutable references, without requiring any special + synchronizations, as long as each invocation gets its own copy. + + \item Avoid raw output on \isa{stdout} or \isa{stderr}. The + Poly/ML library is thread-safe for each individual output operation, + but the ordering of parallel invocations is arbitrary. This means + raw output will appear on some system console with unpredictable + interleaving of atomic chunks. + + Note that this does not affect regular message output channels + (\secref{sec:message-channels}). An official message is associated + with the command transaction from where it originates, independently + of other transactions. This means each running Isar command has + effectively its own set of message channels, and interleaving can + only happen when commands use parallelism internally (and only at + message boundaries). + + \item Treat environment variables and the current working directory + of the running process as strictly read-only. - \item Keep components with local state information - \emph{re-entrant}. Instead of poking initial values into (private) - global references, create a new state record on each invocation, and - pass that through any auxiliary functions of the component. The - state record may well contain mutable references, without requiring - any special synchronizations, as long as each invocation sees its - own copy. Occasionally, one might even return to plain functional - updates on non-mutable record values here. + \item Restrict writing to the file-system to unique temporary files. + Isabelle already provides a temporary directory that is unique for + the running process, and there is a centralized source of unique + serial numbers in Isabelle/ML. Thus temporary files that are passed + to to some external process will be always disjoint, and thus + thread-safe. + + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{File.tmp\_path}\verb|File.tmp_path: Path.T -> Path.T| \\ + \indexdef{}{ML}{serial\_string}\verb|serial_string: unit -> string| \\ + \end{mldecls} + + \begin{description} + + \item \verb|File.tmp_path|~\isa{path} relocates the base + component of \isa{path} into the unique temporary directory of + the running Isabelle/ML process. + + \item \verb|serial_string|~\isa{{\isacharparenleft}{\isacharparenright}} creates a new serial number + that is unique over the runtime of the Isabelle/ML process. - \item Isolate process configuration flags. The main legitimate - application of global references is to configure the whole process - in a certain way, essentially affecting all threads. A typical - example is the \verb|show_types| flag, which tells the pretty printer - to output explicit type information for terms. Such flags usually - do not affect the functionality of the core system, but only the - view being presented to the user. - - Occasionally, such global process flags are treated like implicit - arguments to certain operations, by using the \verb|setmp_CRITICAL| combinator - for safe temporary assignment. Its traditional purpose was to - ensure proper recovery of the original value when exceptions are - raised in the body, now the functionality is extended to enter the - \emph{critical section} (with its usual potential of degrading - parallelism). + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isadelimmlex +% +\endisadelimmlex +% +\isatagmlex +% +\begin{isamarkuptext}% +The following example shows how to create unique + temporary file names.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlex +{\isafoldmlex}% +% +\isadelimmlex +% +\endisadelimmlex +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ val\ tmp{\isadigit{1}}\ {\isacharequal}\ File{\isachardot}tmp{\isacharunderscore}path\ {\isacharparenleft}Path{\isachardot}basic\ {\isacharparenleft}{\isachardoublequote}foo{\isachardoublequote}\ {\isacharcircum}\ serial{\isacharunderscore}string\ {\isacharparenleft}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharsemicolon}\isanewline +\ \ val\ tmp{\isadigit{2}}\ {\isacharequal}\ File{\isachardot}tmp{\isacharunderscore}path\ {\isacharparenleft}Path{\isachardot}basic\ {\isacharparenleft}{\isachardoublequote}foo{\isachardoublequote}\ {\isacharcircum}\ serial{\isacharunderscore}string\ {\isacharparenleft}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharsemicolon}\isanewline +\ \ % +\isaantiq +assert% +\endisaantiq +\ {\isacharparenleft}tmp{\isadigit{1}}\ {\isacharless}{\isachargreater}\ tmp{\isadigit{2}}{\isacharparenright}{\isacharsemicolon}\isanewline +{\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\isamarkupsubsection{Explicit synchronization% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isabelle/ML also provides some explicit synchronization + mechanisms, for the rare situations where mutable shared resources + are really required. These are based on the synchronizations + primitives of Poly/ML, which have been adapted to the specific + assumptions of the concurrent Isabelle/ML environment. User code + must not use the Poly/ML primitives directly! - Note that recovery of plain value passing semantics via \verb|setmp_CRITICAL|~\isa{ref\ value} assumes that this \isa{ref} is - exclusively manipulated within the critical section. In particular, - any persistent global assignment of \isa{ref\ {\isacharcolon}{\isacharequal}\ value} needs to - be marked critical as well, to prevent intruding another threads - local view, and a lost-update in the global scope, too. + \medskip The most basic synchronization concept is a single + \emph{critical section} (also called ``monitor'' in the literature). + A thread that enters the critical section prevents all other threads + from doing the same. A thread that is already within the critical + section may re-enter it in an idempotent manner. - \end{itemize} + Such centralized locking is convenient, because it prevents + deadlocks by construction. - Recall that in an open ``LCF-style'' system like Isabelle/Isar, the - user participates in constructing the overall environment. This - means that state-based facilities offered by one component will - require special caution later on. So minimizing critical elements, - by staying within the plain value-oriented view relative to theory - or proof contexts most of the time, will also reduce the chance of - mishaps occurring to end-users.% + \medskip More fine-grained locking works via \emph{synchronized + variables}. An explicit state component is associated with + mechanisms for locking and signaling. There are operations to + await a condition, change the state, and signal the change to all + other waiting threads. + + Here the synchronized access to the state variable is \emph{not} + re-entrant: direct or indirect nesting within the same thread will + cause a deadlock!% \end{isamarkuptext}% \isamarkuptrue% % @@ -277,215 +2190,49 @@ \begin{mldecls} \indexdef{}{ML}{NAMED\_CRITICAL}\verb|NAMED_CRITICAL: string -> (unit -> 'a) -> 'a| \\ \indexdef{}{ML}{CRITICAL}\verb|CRITICAL: (unit -> 'a) -> 'a| \\ - \indexdef{}{ML}{setmp\_CRITICAL}\verb|setmp_CRITICAL: 'a Unsynchronized.ref -> 'a -> ('b -> 'c) -> 'b -> 'c| \\ + \end{mldecls} + \begin{mldecls} + \indexdef{}{ML type}{Synchronized.var}\verb|type 'a Synchronized.var| \\ + \indexdef{}{ML}{Synchronized.var}\verb|Synchronized.var: string -> 'a -> 'a Synchronized.var| \\ + \indexdef{}{ML}{Synchronized.guarded\_access}\verb|Synchronized.guarded_access: 'a Synchronized.var ->|\isasep\isanewline% +\verb| ('a -> ('b * 'a) option) -> 'b| \\ \end{mldecls} \begin{description} - \item \verb|NAMED_CRITICAL|~\isa{name\ f} evaluates \isa{f\ {\isacharparenleft}{\isacharparenright}} - while staying within the critical section of Isabelle/Isar. No - other thread may do so at the same time, but non-critical parallel - execution will continue. The \isa{name} argument serves for - diagnostic purposes and might help to spot sources of congestion. + \item \verb|NAMED_CRITICAL|~\isa{name\ e} evaluates \isa{e\ {\isacharparenleft}{\isacharparenright}} + within the central critical section of Isabelle/ML. No other thread + may do so at the same time, but non-critical parallel execution will + continue. The \isa{name} argument is used for tracing and might + help to spot sources of congestion. + + Entering the critical section without contention is very fast, and + several basic system operations do so frequently. Each thread + should leave the critical section quickly, otherwise parallel + performance may degrade. \item \verb|CRITICAL| is the same as \verb|NAMED_CRITICAL| with empty name argument. - \item \verb|setmp_CRITICAL|~\isa{ref\ value\ f\ x} evaluates \isa{f\ x} - while staying within the critical section and having \isa{ref\ {\isacharcolon}{\isacharequal}\ value} assigned temporarily. This recovers a value-passing - semantics involving global references, regardless of exceptions or - concurrency. + \item Type \verb|'a Synchronized.var| represents synchronized + variables with state of type \verb|'a|. - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isamarkupchapter{Basic library functions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Beyond the proposal of the SML/NJ basis library, Isabelle comes - with its own library, from which selected parts are given here. - These should encourage a study of the Isabelle sources, - in particular files \emph{Pure/library.ML} and \emph{Pure/General/*.ML}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Linear transformations \label{sec:ML-linear-trans}% -} -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexdef{}{ML}{op $\mid$$>$ }\verb|op |\verb,|,\verb|> : 'a * ('a -> 'b) -> 'b| \\ - \end{mldecls}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\isadelimML -% -\endisadelimML -% -\isatagML -% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -% -\begin{isamarkuptext}% -\noindent Many problems in functional programming can be thought of - as linear transformations, i.e.~a caluclation starts with a - particular value \verb|x : foo| which is then transformed - by application of a function \verb|f : foo -> foo|, - continued by an application of a function \verb|g : foo -> bar|, - and so on. As a canoncial example, take functions enriching - a theory by constant declararion and primitive definitions: - - \smallskip\begin{mldecls} - \verb|Sign.declare_const: (binding * typ) * mixfix|\isasep\isanewline% -\verb| -> theory -> term * theory| \\ - \verb|Thm.add_def: bool -> bool -> binding * term -> theory -> (string * thm) * theory| - \end{mldecls} - - \noindent Written with naive application, an addition of constant - \isa{bar} with type \isa{foo\ {\isasymRightarrow}\ foo} and - a corresponding definition \isa{bar\ {\isasymequiv}\ {\isasymlambda}x{\isachardot}\ x} would look like: - - \smallskip\begin{mldecls} - \verb|(fn (t, thy) => Thm.add_def false false|\isasep\isanewline% -\verb| (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})) thy)|\isasep\isanewline% -\verb| (Sign.declare_const|\isasep\isanewline% -\verb| ((Binding.name "bar", @{typ "foo => foo"}), NoSyn) thy)| - \end{mldecls} + \item \verb|Synchronized.var|~\isa{name\ x} creates a synchronized + variable that is initialized with value \isa{x}. The \isa{name} is used for tracing. - \noindent With increasing numbers of applications, this code gets quite - unreadable. Further, it is unintuitive that things are - written down in the opposite order as they actually ``happen''.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\noindent At this stage, Isabelle offers some combinators which allow - for more convenient notation, most notably reverse application: - - \smallskip\begin{mldecls} -\verb|thy|\isasep\isanewline% -\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline% -\verb||\verb,|,\verb|> (fn (t, thy) => thy|\isasep\isanewline% -\verb||\verb,|,\verb|> Thm.add_def false false|\isasep\isanewline% -\verb| (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})))| - \end{mldecls}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexdef{}{ML}{op $\mid$-$>$ }\verb|op |\verb,|,\verb|-> : ('c * 'a) * ('c -> 'a -> 'b) -> 'b| \\ - \indexdef{}{ML}{op $\mid$$>$$>$ }\verb|op |\verb,|,\verb|>> : ('a * 'c) * ('a -> 'b) -> 'b * 'c| \\ - \indexdef{}{ML}{op $\mid$$\mid$$>$ }\verb|op |\verb,|,\verb||\verb,|,\verb|> : ('c * 'a) * ('a -> 'b) -> 'c * 'b| \\ - \indexdef{}{ML}{op $\mid$$\mid$$>$$>$ }\verb|op |\verb,|,\verb||\verb,|,\verb|>> : ('c * 'a) * ('a -> 'd * 'b) -> ('c * 'd) * 'b| \\ - \end{mldecls}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\begin{isamarkuptext}% -\noindent Usually, functions involving theory updates also return - side results; to use these conveniently, yet another - set of combinators is at hand, most notably \verb|op |\verb,|,\verb|->| - which allows curried access to side results: + \item \verb|Synchronized.guarded_access|~\isa{var\ f} lets the + function \isa{f} operate within a critical section on the state + \isa{x} as follows: if \isa{f\ x} produces \verb|NONE|, we + continue to wait on the internal condition variable, expecting that + some other thread will eventually change the content in a suitable + manner; if \isa{f\ x} produces \verb|SOME|~\isa{{\isacharparenleft}y{\isacharcomma}\ x{\isacharprime}{\isacharparenright}} we + are satisfied and assign the new state value \isa{x{\isacharprime}}, broadcast + a signal to all waiting threads on the associated condition + variable, and return the result \isa{y}. - \smallskip\begin{mldecls} -\verb|thy|\isasep\isanewline% -\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline% -\verb||\verb,|,\verb|-> (fn t => Thm.add_def false false|\isasep\isanewline% -\verb| (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})))|\isasep\isanewline% - - \end{mldecls} - - \noindent \verb|op |\verb,|,\verb|>>| allows for processing side results on their own: - - \smallskip\begin{mldecls} -\verb|thy|\isasep\isanewline% -\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline% -\verb||\verb,|,\verb|>> (fn t => Logic.mk_equals (t, @{term "%x. x"}))|\isasep\isanewline% -\verb||\verb,|,\verb|-> (fn def => Thm.add_def false false (Binding.name "bar_def", def))|\isasep\isanewline% - - \end{mldecls} - - \noindent Orthogonally, \verb|op |\verb,|,\verb||\verb,|,\verb|>| applies a transformation - in the presence of side results which are left unchanged: + \end{description} - \smallskip\begin{mldecls} -\verb|thy|\isasep\isanewline% -\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline% -\verb||\verb,|,\verb||\verb,|,\verb|> Sign.add_path "foobar"|\isasep\isanewline% -\verb||\verb,|,\verb|-> (fn t => Thm.add_def false false|\isasep\isanewline% -\verb| (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})))|\isasep\isanewline% -\verb||\verb,|,\verb||\verb,|,\verb|> Sign.restore_naming thy|\isasep\isanewline% - - \end{mldecls} - - \noindent \verb|op |\verb,|,\verb||\verb,|,\verb|>>| accumulates side results: - - \smallskip\begin{mldecls} -\verb|thy|\isasep\isanewline% -\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline% -\verb||\verb,|,\verb||\verb,|,\verb|>> Sign.declare_const ((Binding.name "foobar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline% -\verb||\verb,|,\verb|-> (fn (t1, t2) => Thm.add_def false false|\isasep\isanewline% -\verb| (Binding.name "bar_def", Logic.mk_equals (t1, t2)))|\isasep\isanewline% - - \end{mldecls}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexdef{}{ML}{fold}\verb|fold: ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b| \\ - \indexdef{}{ML}{fold\_map}\verb|fold_map: ('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b| \\ - \end{mldecls}% + There are some further variants of the general \verb|Synchronized.guarded_access| combinator, see \hyperlink{file.~~/src/Pure/Concurrent/synchronized.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Concurrent{\isacharslash}synchronized{\isachardot}ML}}}} for details.% \end{isamarkuptext}% \isamarkuptrue% % @@ -496,283 +2243,68 @@ % \endisadelimmlref % -\begin{isamarkuptext}% -\noindent This principles naturally lift to \emph{lists} using - the \verb|fold| and \verb|fold_map| combinators. - The first lifts a single function - \begin{quote}\footnotesize - \verb|f : 'a -> 'b -> 'b| to \verb|'a list -> 'b -> 'b| - \end{quote} - such that - \begin{quote}\footnotesize - \verb|y |\verb,|,\verb|> fold f [x1, x2, ..., x_n]| \\ - \hspace*{2ex}\isa{{\isasymleadsto}} \verb|y |\verb,|,\verb|> f x1 |\verb,|,\verb|> f x2 |\verb,|,\verb|> ... |\verb,|,\verb|> f x_n| - \end{quote} - \noindent The second accumulates side results in a list by lifting - a single function - \begin{quote}\footnotesize - \verb|f : 'a -> 'b -> 'c * 'b| to \verb|'a list -> 'b -> 'c list * 'b| - \end{quote} - such that - \begin{quote}\footnotesize - \verb|y |\verb,|,\verb|> fold_map f [x1, x2, ..., x_n]| \\ - \hspace*{2ex}\isa{{\isasymleadsto}} \verb|y |\verb,|,\verb|> f x1 |\verb,|,\verb||\verb,|,\verb|>> f x2 |\verb,|,\verb||\verb,|,\verb|>> ... |\verb,|,\verb||\verb,|,\verb|>> f x_n| \\ - \hspace*{6ex}\verb||\verb,|,\verb||\verb,|,\verb|> (fn ((z1, z2), ..., z_n) => [z1, z2, ..., z_n])| - \end{quote} - - \noindent Example: - - \smallskip\begin{mldecls} -\verb|let|\isasep\isanewline% -\verb| val consts = ["foo", "bar"];|\isasep\isanewline% -\verb|in|\isasep\isanewline% -\verb| thy|\isasep\isanewline% -\verb| |\verb,|,\verb|> fold_map (fn const => Sign.declare_const|\isasep\isanewline% -\verb| ((Binding.name const, @{typ "foo => foo"}), NoSyn)) consts|\isasep\isanewline% -\verb| |\verb,|,\verb|>> map (fn t => Logic.mk_equals (t, @{term "%x. x"}))|\isasep\isanewline% -\verb| |\verb,|,\verb|-> (fn defs => fold_map (fn def =>|\isasep\isanewline% -\verb| Thm.add_def false false (Binding.empty, def)) defs)|\isasep\isanewline% -\verb|end| - \end{mldecls}% -\end{isamarkuptext}% -\isamarkuptrue% +\isadelimmlex % -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref +\endisadelimmlex % -\begin{isamarkuptext}% -\begin{mldecls} - \indexdef{}{ML}{op \#$>$ }\verb|op #> : ('a -> 'b) * ('b -> 'c) -> 'a -> 'c| \\ - \indexdef{}{ML}{op \#-$>$ }\verb|op #-> : ('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd| \\ - \indexdef{}{ML}{op \#$>$$>$ }\verb|op #>> : ('a -> 'c * 'b) * ('c -> 'd) -> 'a -> 'd * 'b| \\ - \indexdef{}{ML}{op \#\#$>$ }\verb|op ##> : ('a -> 'c * 'b) * ('b -> 'd) -> 'a -> 'c * 'd| \\ - \indexdef{}{ML}{op \#\#$>$$>$ }\verb|op ##>> : ('a -> 'c * 'b) * ('b -> 'e * 'd) -> 'a -> ('c * 'e) * 'd| \\ - \end{mldecls}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref +\isatagmlex % \begin{isamarkuptext}% -\noindent All those linear combinators also exist in higher-order - variants which do not expect a value on the left hand side - but a function.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexdef{}{ML}{op ` }\verb|op ` : ('b -> 'a) -> 'b -> 'a * 'b| \\ - \indexdef{}{ML}{tap}\verb|tap: ('b -> 'a) -> 'b -> 'b| \\ - \end{mldecls}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref -% -\begin{isamarkuptext}% -\noindent These operators allow to ``query'' a context - in a series of context transformations: +See \hyperlink{file.~~/src/Pure/Concurrent/mailbox.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Concurrent{\isacharslash}mailbox{\isachardot}ML}}}} how to + implement a mailbox as synchronized variable over a purely + functional queue. - \smallskip\begin{mldecls} -\verb|thy|\isasep\isanewline% -\verb||\verb,|,\verb|> tap (fn _ => writeln "now adding constant")|\isasep\isanewline% -\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline% -\verb||\verb,|,\verb||\verb,|,\verb|>> `(fn thy => Sign.declared_const thy|\isasep\isanewline% -\verb| (Sign.full_name thy (Binding.name "foobar")))|\isasep\isanewline% -\verb||\verb,|,\verb|-> (fn (t, b) => if b then I|\isasep\isanewline% -\verb| else Sign.declare_const|\isasep\isanewline% -\verb| ((Binding.name "foobar", @{typ "foo => foo"}), NoSyn) #> snd)|\isasep\isanewline% - - \end{mldecls}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Options and partiality% -} -\isamarkuptrue% -% -\isadelimmlref -% -\endisadelimmlref -% -\isatagmlref -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexdef{}{ML}{is\_some}\verb|is_some: 'a option -> bool| \\ - \indexdef{}{ML}{is\_none}\verb|is_none: 'a option -> bool| \\ - \indexdef{}{ML}{the}\verb|the: 'a option -> 'a| \\ - \indexdef{}{ML}{these}\verb|these: 'a list option -> 'a list| \\ - \indexdef{}{ML}{the\_list}\verb|the_list: 'a option -> 'a list| \\ - \indexdef{}{ML}{the\_default}\verb|the_default: 'a -> 'a option -> 'a| \\ - \indexdef{}{ML}{try}\verb|try: ('a -> 'b) -> 'a -> 'b option| \\ - \indexdef{}{ML}{can}\verb|can: ('a -> 'b) -> 'a -> bool| \\ - \end{mldecls}% + \medskip The following example implements a counter that produces + positive integers that are unique over the runtime of the Isabelle + process:% \end{isamarkuptext}% \isamarkuptrue% % -\endisatagmlref -{\isafoldmlref}% -% -\isadelimmlref -% -\endisadelimmlref +\endisatagmlex +{\isafoldmlex}% % -\begin{isamarkuptext}% -Standard selector functions on \isa{option}s are provided. The - \verb|try| and \verb|can| functions provide a convenient interface for - handling exceptions -- both take as arguments a function \verb|f| - together with a parameter \verb|x| and handle any exception during - the evaluation of the application of \verb|f| to \verb|x|, either - return a lifted result (\verb|NONE| on failure) or a boolean value - (\verb|false| on failure).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Common data structures% -} -\isamarkuptrue% -% -\isamarkupsubsection{Lists (as set-like data structures)% -} -\isamarkuptrue% +\isadelimmlex % -\begin{isamarkuptext}% -\begin{mldecls} - \indexdef{}{ML}{member}\verb|member: ('b * 'a -> bool) -> 'a list -> 'b -> bool| \\ - \indexdef{}{ML}{insert}\verb|insert: ('a * 'a -> bool) -> 'a -> 'a list -> 'a list| \\ - \indexdef{}{ML}{remove}\verb|remove: ('b * 'a -> bool) -> 'b -> 'a list -> 'a list| \\ - \indexdef{}{ML}{merge}\verb|merge: ('a * 'a -> bool) -> 'a list * 'a list -> 'a list| \\ - \end{mldecls}% -\end{isamarkuptext}% -\isamarkuptrue% +\endisadelimmlex % -\begin{isamarkuptext}% -Lists are often used as set-like data structures -- set-like in - the sense that they support a notion of \verb|member|-ship, - \verb|insert|-ing and \verb|remove|-ing, but are order-sensitive. - This is convenient when implementing a history-like mechanism: - \verb|insert| adds an element \emph{to the front} of a list, - if not yet present; \verb|remove| removes \emph{all} occurences - of a particular element. Correspondingly \verb|merge| implements a - a merge on two lists suitable for merges of context data - (\secref{sec:context-theory}). - - Functions are parametrized by an explicit equality function - to accomplish overloaded equality; in most cases of monomorphic - equality, writing \verb|op =| should suffice.% -\end{isamarkuptext}% -\isamarkuptrue% +\isadelimML % -\isamarkupsubsection{Association lists% -} -\isamarkuptrue% +\endisadelimML % -\begin{isamarkuptext}% -\begin{mldecls} - \indexdef{}{ML exception}{AList.DUP}\verb|exception AList.DUP| \\ - \indexdef{}{ML}{AList.lookup}\verb|AList.lookup: ('a * 'b -> bool) -> ('b * 'c) list -> 'a -> 'c option| \\ - \indexdef{}{ML}{AList.defined}\verb|AList.defined: ('a * 'b -> bool) -> ('b * 'c) list -> 'a -> bool| \\ - \indexdef{}{ML}{AList.update}\verb|AList.update: ('a * 'a -> bool) -> ('a * 'b) -> ('a * 'b) list -> ('a * 'b) list| \\ - \indexdef{}{ML}{AList.default}\verb|AList.default: ('a * 'a -> bool) -> ('a * 'b) -> ('a * 'b) list -> ('a * 'b) list| \\ - \indexdef{}{ML}{AList.delete}\verb|AList.delete: ('a * 'b -> bool) -> 'a -> ('b * 'c) list -> ('b * 'c) list| \\ - \indexdef{}{ML}{AList.map\_entry}\verb|AList.map_entry: ('a * 'b -> bool) -> 'a|\isasep\isanewline% -\verb| -> ('c -> 'c) -> ('b * 'c) list -> ('b * 'c) list| \\ - \indexdef{}{ML}{AList.map\_default}\verb|AList.map_default: ('a * 'a -> bool) -> 'a * 'b -> ('b -> 'b)|\isasep\isanewline% -\verb| -> ('a * 'b) list -> ('a * 'b) list| \\ - \indexdef{}{ML}{AList.join}\verb|AList.join: ('a * 'a -> bool) -> ('a -> 'b * 'b -> 'b) (*exception DUP*)|\isasep\isanewline% -\verb| -> ('a * 'b) list * ('a * 'b) list -> ('a * 'b) list (*exception AList.DUP*)| \\ - \indexdef{}{ML}{AList.merge}\verb|AList.merge: ('a * 'a -> bool) -> ('b * 'b -> bool)|\isasep\isanewline% -\verb| -> ('a * 'b) list * ('a * 'b) list -> ('a * 'b) list (*exception AList.DUP*)| - \end{mldecls}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -Association lists can be seens as an extension of set-like lists: - on the one hand, they may be used to implement finite mappings, - on the other hand, they remain order-sensitive and allow for - multiple key-value-pair with the same key: \verb|AList.lookup| - returns the \emph{first} value corresponding to a particular - key, if present. \verb|AList.update| updates - the \emph{first} occurence of a particular key; if no such - key exists yet, the key-value-pair is added \emph{to the front}. - \verb|AList.delete| only deletes the \emph{first} occurence of a key. - \verb|AList.merge| provides an operation suitable for merges of context data - (\secref{sec:context-theory}), where an equality parameter on - values determines whether a merge should be considered a conflict. - A slightly generalized operation if implementend by the \verb|AList.join| - function which allows for explicit conflict resolution.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Tables% -} -\isamarkuptrue% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ local\isanewline +\ \ \ \ val\ counter\ {\isacharequal}\ Synchronized{\isachardot}var\ {\isachardoublequote}counter{\isachardoublequote}\ {\isadigit{0}}{\isacharsemicolon}\isanewline +\ \ in\isanewline +\ \ \ \ fun\ next\ {\isacharparenleft}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ \ Synchronized{\isachardot}guarded{\isacharunderscore}access\ counter\isanewline +\ \ \ \ \ \ \ \ {\isacharparenleft}fn\ i\ {\isacharequal}{\isachargreater}\isanewline +\ \ \ \ \ \ \ \ \ \ let\ val\ j\ {\isacharequal}\ i\ {\isacharplus}\ {\isadigit{1}}\isanewline +\ \ \ \ \ \ \ \ \ \ in\ SOME\ {\isacharparenleft}j{\isacharcomma}\ j{\isacharparenright}\ end{\isacharparenright}{\isacharsemicolon}\isanewline +\ \ end{\isacharsemicolon}\isanewline +{\isacharverbatimclose}\isanewline +\isanewline +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ val\ a\ {\isacharequal}\ next\ {\isacharparenleft}{\isacharparenright}{\isacharsemicolon}\isanewline +\ \ val\ b\ {\isacharequal}\ next\ {\isacharparenleft}{\isacharparenright}{\isacharsemicolon}\isanewline +\ \ % +\isaantiq +assert% +\endisaantiq +\ {\isacharparenleft}a\ {\isacharless}{\isachargreater}\ b{\isacharparenright}{\isacharsemicolon}\isanewline +{\isacharverbatimclose}% +\endisatagML +{\isafoldML}% % -\begin{isamarkuptext}% -\begin{mldecls} - \indexdef{}{ML type}{'a Symtab.table}\verb|type 'a Symtab.table| \\ - \indexdef{}{ML exception}{Symtab.DUP}\verb|exception Symtab.DUP of string| \\ - \indexdef{}{ML exception}{Symtab.SAME}\verb|exception Symtab.SAME| \\ - \indexdef{}{ML exception}{Symtab.UNDEF}\verb|exception Symtab.UNDEF of string| \\ - \indexdef{}{ML}{Symtab.empty}\verb|Symtab.empty: 'a Symtab.table| \\ - \indexdef{}{ML}{Symtab.lookup}\verb|Symtab.lookup: 'a Symtab.table -> string -> 'a option| \\ - \indexdef{}{ML}{Symtab.defined}\verb|Symtab.defined: 'a Symtab.table -> string -> bool| \\ - \indexdef{}{ML}{Symtab.update}\verb|Symtab.update: (string * 'a) -> 'a Symtab.table -> 'a Symtab.table| \\ - \indexdef{}{ML}{Symtab.default}\verb|Symtab.default: string * 'a -> 'a Symtab.table -> 'a Symtab.table| \\ - \indexdef{}{ML}{Symtab.delete}\verb|Symtab.delete: string|\isasep\isanewline% -\verb| -> 'a Symtab.table -> 'a Symtab.table (*exception Symtab.UNDEF*)| \\ - \indexdef{}{ML}{Symtab.map\_entry}\verb|Symtab.map_entry: string -> ('a -> 'a)|\isasep\isanewline% -\verb| -> 'a Symtab.table -> 'a Symtab.table| \\ - \indexdef{}{ML}{Symtab.map\_default}\verb|Symtab.map_default: (string * 'a) -> ('a -> 'a)|\isasep\isanewline% -\verb| -> 'a Symtab.table -> 'a Symtab.table| \\ - \indexdef{}{ML}{Symtab.join}\verb|Symtab.join: (string -> 'a * 'a -> 'a) (*exception Symtab.DUP/Symtab.SAME*)|\isasep\isanewline% -\verb| -> 'a Symtab.table * 'a Symtab.table|\isasep\isanewline% -\verb| -> 'a Symtab.table (*exception Symtab.DUP*)| \\ - \indexdef{}{ML}{Symtab.merge}\verb|Symtab.merge: ('a * 'a -> bool)|\isasep\isanewline% -\verb| -> 'a Symtab.table * 'a Symtab.table|\isasep\isanewline% -\verb| -> 'a Symtab.table (*exception Symtab.DUP*)| - \end{mldecls}% -\end{isamarkuptext}% -\isamarkuptrue% +\isadelimML % -\begin{isamarkuptext}% -Tables are an efficient representation of finite mappings without - any notion of order; due to their efficiency they should be used - whenever such pure finite mappings are neccessary. - - The key type of tables must be given explicitly by instantiating - the \verb|Table| functor which takes the key type - together with its \verb|order|; for convience, we restrict - here to the \verb|Symtab| instance with \verb|string| - as key type. - - Most table functions correspond to those of association lists.% -\end{isamarkuptext}% -\isamarkuptrue% +\endisadelimML +\isanewline % \isadelimtheory +\isanewline % \endisadelimtheory % diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/document/Prelim.tex --- a/doc-src/IsarImplementation/Thy/document/Prelim.tex Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/document/Prelim.tex Mon Oct 25 08:08:08 2010 +0200 @@ -159,13 +159,13 @@ \medskip There is a separate notion of \emph{theory reference} for maintaining a live link to an evolving theory context: updates on - drafts are propagated automatically. Dynamic updating stops after - an explicit \isa{end} only. + drafts are propagated automatically. Dynamic updating stops when + the next \isa{checkpoint} is reached. Derived entities may store a theory reference in order to indicate - the context they belong to. This implicitly assumes monotonic - reasoning, because the referenced context may become larger without - further notice.% + the formal context from which they are derived. This implicitly + assumes monotonic reasoning, because the referenced context may + become larger without further notice.% \end{isamarkuptext}% \isamarkuptrue% % @@ -178,11 +178,14 @@ \begin{isamarkuptext}% \begin{mldecls} \indexdef{}{ML type}{theory}\verb|type theory| \\ + \indexdef{}{ML}{Theory.eq\_thy}\verb|Theory.eq_thy: theory * theory -> bool| \\ \indexdef{}{ML}{Theory.subthy}\verb|Theory.subthy: theory * theory -> bool| \\ \indexdef{}{ML}{Theory.checkpoint}\verb|Theory.checkpoint: theory -> theory| \\ \indexdef{}{ML}{Theory.copy}\verb|Theory.copy: theory -> theory| \\ \indexdef{}{ML}{Theory.merge}\verb|Theory.merge: theory * theory -> theory| \\ \indexdef{}{ML}{Theory.begin\_theory}\verb|Theory.begin_theory: string -> theory list -> theory| \\ + \indexdef{}{ML}{Theory.parents\_of}\verb|Theory.parents_of: theory -> theory list| \\ + \indexdef{}{ML}{Theory.ancestors\_of}\verb|Theory.ancestors_of: theory -> theory list| \\ \end{mldecls} \begin{mldecls} \indexdef{}{ML type}{theory\_ref}\verb|type theory_ref| \\ @@ -192,10 +195,15 @@ \begin{description} - \item \verb|theory| represents theory contexts. This is - essentially a linear type, with explicit runtime checking! Most - internal theory operations destroy the original version, which then - becomes ``stale''. + \item Type \verb|theory| represents theory contexts. This is + essentially a linear type, with explicit runtime checking. + Primitive theory operations destroy the original version, which then + becomes ``stale''. This can be prevented by explicit checkpointing, + which the system does at least at the boundary of toplevel command + transactions \secref{sec:isar-toplevel}. + + \item \verb|Theory.eq_thy|~\isa{{\isacharparenleft}thy\isactrlsub {\isadigit{1}}{\isacharcomma}\ thy\isactrlsub {\isadigit{2}}{\isacharparenright}} check strict + identity of two theories. \item \verb|Theory.subthy|~\isa{{\isacharparenleft}thy\isactrlsub {\isadigit{1}}{\isacharcomma}\ thy\isactrlsub {\isadigit{2}}{\isacharparenright}} compares theories according to the intrinsic graph structure of the construction. @@ -216,11 +224,17 @@ This version of ad-hoc theory merge fails for unrelated theories! \item \verb|Theory.begin_theory|~\isa{name\ parents} constructs - a new theory based on the given parents. This {\ML} function is + a new theory based on the given parents. This ML function is normally not invoked directly. - \item \verb|theory_ref| represents a sliding reference to an - always valid theory; updates on the original are propagated + \item \verb|Theory.parents_of|~\isa{thy} returns the direct + ancestors of \isa{thy}. + + \item \verb|Theory.ancestors_of|~\isa{thy} returns all + ancestors of \isa{thy} (not including \isa{thy} itself). + + \item Type \verb|theory_ref| represents a sliding reference to + an always valid theory; updates on the original are propagated automatically. \item \verb|Theory.deref|~\isa{thy{\isacharunderscore}ref} turns a \verb|theory_ref| into an \verb|theory| value. As the referenced @@ -239,14 +253,54 @@ % \endisadelimmlref % +\isadelimmlantiq +% +\endisadelimmlantiq +% +\isatagmlantiq +% +\begin{isamarkuptext}% +\begin{matharray}{rcl} + \indexdef{}{ML antiquotation}{theory}\hypertarget{ML antiquotation.theory}{\hyperlink{ML antiquotation.theory}{\mbox{\isa{theory}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{theory\_ref}\hypertarget{ML antiquotation.theory-ref}{\hyperlink{ML antiquotation.theory-ref}{\mbox{\isa{theory{\isacharunderscore}ref}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \end{matharray} + + \begin{rail} + ('theory' | 'theory\_ref') nameref? + ; + \end{rail} + + \begin{description} + + \item \isa{{\isacharat}{\isacharbraceleft}theory{\isacharbraceright}} refers to the background theory of the + current context --- as abstract value. + + \item \isa{{\isacharat}{\isacharbraceleft}theory\ A{\isacharbraceright}} refers to an explicitly named ancestor + theory \isa{A} of the background theory of the current context + --- as abstract value. + + \item \isa{{\isacharat}{\isacharbraceleft}theory{\isacharunderscore}ref{\isacharbraceright}} is similar to \isa{{\isacharat}{\isacharbraceleft}theory{\isacharbraceright}}, but + produces a \verb|theory_ref| via \verb|Theory.check_thy| as + explained above. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlantiq +{\isafoldmlantiq}% +% +\isadelimmlantiq +% +\endisadelimmlantiq +% \isamarkupsubsection{Proof context \label{sec:context-proof}% } \isamarkuptrue% % \begin{isamarkuptext}% A proof context is a container for pure data with a - back-reference to the theory it belongs to. The \isa{init} - operation creates a proof context from a given theory. + back-reference to the theory from which it is derived. The \isa{init} operation creates a proof context from a given theory. Modifications to draft theories are propagated to the proof context as usual, but there is also an explicit \isa{transfer} operation to force resynchronization with more substantial updates to the @@ -289,9 +343,9 @@ \begin{description} - \item \verb|Proof.context| represents proof contexts. Elements - of this type are essentially pure values, with a sliding reference - to the background theory. + \item Type \verb|Proof.context| represents proof contexts. + Elements of this type are essentially pure values, with a sliding + reference to the background theory. \item \verb|ProofContext.init_global|~\isa{thy} produces a proof context derived from \isa{thy}, initializing all data. @@ -314,6 +368,37 @@ % \endisadelimmlref % +\isadelimmlantiq +% +\endisadelimmlantiq +% +\isatagmlantiq +% +\begin{isamarkuptext}% +\begin{matharray}{rcl} + \indexdef{}{ML antiquotation}{context}\hypertarget{ML antiquotation.context}{\hyperlink{ML antiquotation.context}{\mbox{\isa{context}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \end{matharray} + + \begin{description} + + \item \isa{{\isacharat}{\isacharbraceleft}context{\isacharbraceright}} refers to \emph{the} context at + compile-time --- as abstract value. Independently of (local) theory + or proof mode, this always produces a meaningful result. + + This is probably the most common antiquotation in interactive + experimentation with ML inside Isar. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlantiq +{\isafoldmlantiq}% +% +\isadelimmlantiq +% +\endisadelimmlantiq +% \isamarkupsubsection{Generic contexts \label{sec:generic-context}% } \isamarkuptrue% @@ -348,7 +433,7 @@ \begin{description} - \item \verb|Context.generic| is the direct sum of \verb|theory| and \verb|Proof.context|, with the datatype + \item Type \verb|Context.generic| is the direct sum of \verb|theory| and \verb|Proof.context|, with the datatype constructors \verb|Context.Theory| and \verb|Context.Proof|. \item \verb|Context.theory_of|~\isa{context} always produces a @@ -391,8 +476,9 @@ \end{tabular} \medskip - \noindent The \isa{empty} value acts as initial default for - \emph{any} theory that does not declare actual data content; \isa{extend} is acts like a unitary version of \isa{merge}. + The \isa{empty} value acts as initial default for \emph{any} + theory that does not declare actual data content; \isa{extend} + is acts like a unitary version of \isa{merge}. Implementing \isa{merge} can be tricky. The general idea is that \isa{merge\ {\isacharparenleft}data\isactrlsub {\isadigit{1}}{\isacharcomma}\ data\isactrlsub {\isadigit{2}}{\isacharparenright}} inserts those parts of \isa{data\isactrlsub {\isadigit{2}}} into \isa{data\isactrlsub {\isadigit{1}}} that are not yet present, while @@ -413,19 +499,24 @@ \end{tabular} \medskip - \noindent The \isa{init} operation is supposed to produce a pure - value from the given background theory and should be somehow + The \isa{init} operation is supposed to produce a pure value + from the given background theory and should be somehow ``immediate''. Whenever a proof context is initialized, which happens frequently, the the system invokes the \isa{init} - operation of \emph{all} theory data slots ever declared. + operation of \emph{all} theory data slots ever declared. This also + means that one needs to be economic about the total number of proof + data declarations in the system, i.e.\ each ML module should declare + at most one, sometimes two data slots for its internal use. + Repeated data declarations to simulate a record type should be + avoided! \paragraph{Generic data} provides a hybrid interface for both theory and proof data. The \isa{init} operation for proof contexts is predefined to select the current data value from the background theory. - \bigskip Any of these data declaration over type \isa{T} result - in an ML structure with the following signature: + \bigskip Any of the above data declarations over type \isa{T} + result in an ML structure with the following signature: \medskip \begin{tabular}{ll} @@ -435,12 +526,12 @@ \end{tabular} \medskip - \noindent These other operations provide exclusive access for the - particular kind of context (theory, proof, or generic context). - This interface fully observes the ML discipline for types and - scopes: there is no other way to access the corresponding data slot - of a context. By keeping these operations private, an Isabelle/ML - module may maintain abstract values authentically.% + These other operations provide exclusive access for the particular + kind of context (theory, proof, or generic context). This interface + observes the ML discipline for types and scopes: there is no other + way to access the corresponding data slot of a context. By keeping + these operations private, an Isabelle/ML module may maintain + abstract values authentically.% \end{isamarkuptext}% \isamarkuptrue% % @@ -522,9 +613,9 @@ \endisadelimML % \begin{isamarkuptext}% -\noindent The implementation uses private theory data - internally, and only exposes an operation that involves explicit - argument checking wrt.\ the given theory.% +The implementation uses private theory data internally, and + only exposes an operation that involves explicit argument checking + wrt.\ the given theory.% \end{isamarkuptext}% \isamarkuptrue% % @@ -545,13 +636,16 @@ \ \ \ \ val\ extend\ {\isacharequal}\ I{\isacharsemicolon}\isanewline \ \ \ \ fun\ merge\ {\isacharparenleft}ts{\isadigit{1}}{\isacharcomma}\ ts{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\isanewline \ \ \ \ \ \ Ord{\isacharunderscore}List{\isachardot}union\ Term{\isacharunderscore}Ord{\isachardot}fast{\isacharunderscore}term{\isacharunderscore}ord\ ts{\isadigit{1}}\ ts{\isadigit{2}}{\isacharsemicolon}\isanewline -\ \ {\isacharparenright}\isanewline +\ \ {\isacharparenright}{\isacharsemicolon}\isanewline \isanewline \ \ val\ get\ {\isacharequal}\ Terms{\isachardot}get{\isacharsemicolon}\isanewline \isanewline \ \ fun\ add\ raw{\isacharunderscore}t\ thy\ {\isacharequal}\isanewline -\ \ \ \ let\ val\ t\ {\isacharequal}\ Sign{\isachardot}cert{\isacharunderscore}term\ thy\ raw{\isacharunderscore}t\isanewline -\ \ \ \ in\ Terms{\isachardot}map\ {\isacharparenleft}Ord{\isacharunderscore}List{\isachardot}insert\ Term{\isacharunderscore}Ord{\isachardot}fast{\isacharunderscore}term{\isacharunderscore}ord\ t{\isacharparenright}\ thy\ end{\isacharsemicolon}\isanewline +\ \ \ \ let\isanewline +\ \ \ \ \ \ val\ t\ {\isacharequal}\ Sign{\isachardot}cert{\isacharunderscore}term\ thy\ raw{\isacharunderscore}t{\isacharsemicolon}\isanewline +\ \ \ \ in\isanewline +\ \ \ \ \ \ Terms{\isachardot}map\ {\isacharparenleft}Ord{\isacharunderscore}List{\isachardot}insert\ Term{\isacharunderscore}Ord{\isachardot}fast{\isacharunderscore}term{\isacharunderscore}ord\ t{\isacharparenright}\ thy\isanewline +\ \ \ \ end{\isacharsemicolon}\isanewline \isanewline \ \ end{\isacharsemicolon}\isanewline {\isacharverbatimclose}% @@ -563,17 +657,17 @@ \endisadelimML % \begin{isamarkuptext}% -We use \verb|term Ord_List.T| for reasonably efficient - representation of a set of terms: all operations are linear in the - number of stored elements. Here we assume that our users do not - care about the declaration order, since that data structure forces - its own arrangement of elements. +Type \verb|term Ord_List.T| is used for reasonably + efficient representation of a set of terms: all operations are + linear in the number of stored elements. Here we assume that users + of this module do not care about the declaration order, since that + data structure forces its own arrangement of elements. Observe how the \verb|merge| operation joins the data slots of the two constituents: \verb|Ord_List.union| prevents duplication of common data from different branches, thus avoiding the danger of - exponential blowup. (Plain list append etc.\ must never be used for - theory data merges.) + exponential blowup. Plain list append etc.\ must never be used for + theory data merges! \medskip Our intended invariant is achieved as follows: \begin{enumerate} @@ -594,7 +688,7 @@ type-inference via \verb|Syntax.check_term| (cf.\ \secref{sec:term-check}) is not necessarily monotonic wrt.\ the background theory, since constraints of term constants can be - strengthened by later declarations, for example. + modified by later declarations, for example. In most cases, user-space context data does not have to take such invariants too seriously. The situation is different in the @@ -603,6 +697,333 @@ \end{isamarkuptext}% \isamarkuptrue% % +\isamarkupsubsection{Configuration options \label{sec:config-options}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +A \emph{configuration option} is a named optional value of + some basic type (Boolean, integer, string) that is stored in the + context. It is a simple application of general context data + (\secref{sec:context-data}) that is sufficiently common to justify + customized setup, which includes some concrete declarations for + end-users using existing notation for attributes (cf.\ + \secref{sec:attributes}). + + For example, the predefined configuration option \hyperlink{attribute.show-types}{\mbox{\isa{show{\isacharunderscore}types}}} controls output of explicit type constraints for + variables in printed terms (cf.\ \secref{sec:read-print}). Its + value can be modified within Isar text like this:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{declare}\isamarkupfalse% +\ {\isacharbrackleft}{\isacharbrackleft}show{\isacharunderscore}types\ {\isacharequal}\ false{\isacharbrackright}{\isacharbrackright}\isanewline +\ \ % +\isamarkupcmt{declaration within (local) theory context% +} +\isanewline +\isanewline +\isacommand{example{\isacharunderscore}proof}\isamarkupfalse% +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{note}\isamarkupfalse% +\ {\isacharbrackleft}{\isacharbrackleft}show{\isacharunderscore}types\ {\isacharequal}\ true{\isacharbrackright}{\isacharbrackright}\isanewline +\ \ \ \ % +\isamarkupcmt{declaration within proof (forward mode)% +} +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\ \ \isacommand{term}\isamarkupfalse% +\ x\isanewline +% +\isadelimproof +\isanewline +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\ \ \ \ \isacommand{using}\isamarkupfalse% +\ {\isacharbrackleft}{\isacharbrackleft}show{\isacharunderscore}types\ {\isacharequal}\ false{\isacharbrackright}{\isacharbrackright}\isanewline +\ \ \ \ \ \ % +\isamarkupcmt{declaration within proof (backward mode)% +} +\isanewline +\ \ \ \ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Configuration options that are not set explicitly hold a + default value that can depend on the application context. This + allows to retrieve the value from another slot within the context, + or fall back on a global preference mechanism, for example. + + The operations to declare configuration options and get/map their + values are modeled as direct replacements for historic global + references, only that the context is made explicit. This allows + easy configuration of tools, without relying on the execution order + as required for old-style mutable references.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{Config.get}\verb|Config.get: Proof.context -> 'a Config.T -> 'a| \\ + \indexdef{}{ML}{Config.map}\verb|Config.map: 'a Config.T -> ('a -> 'a) -> Proof.context -> Proof.context| \\ + \indexdef{}{ML}{Attrib.config\_bool}\verb|Attrib.config_bool: string -> (Context.generic -> bool) ->|\isasep\isanewline% +\verb| bool Config.T * (theory -> theory)| \\ + \indexdef{}{ML}{Attrib.config\_int}\verb|Attrib.config_int: string -> (Context.generic -> int) ->|\isasep\isanewline% +\verb| int Config.T * (theory -> theory)| \\ + \indexdef{}{ML}{Attrib.config\_string}\verb|Attrib.config_string: string -> (Context.generic -> string) ->|\isasep\isanewline% +\verb| string Config.T * (theory -> theory)| \\ + \end{mldecls} + + \begin{description} + + \item \verb|Config.get|~\isa{ctxt\ config} gets the value of + \isa{config} in the given context. + + \item \verb|Config.map|~\isa{config\ f\ ctxt} updates the context + by updating the value of \isa{config}. + + \item \isa{{\isacharparenleft}config{\isacharcomma}\ setup{\isacharparenright}\ {\isacharequal}}~\verb|Attrib.config_bool|~\isa{name\ default} creates a named configuration option of type + \verb|bool|, with the given \isa{default} depending on the + application context. The resulting \isa{config} can be used to + get/map its value in a given context. The \isa{setup} function + needs to be applied to the theory initially, in order to make + concrete declaration syntax available to the user. + + \item \verb|Attrib.config_int| and \verb|Attrib.config_string| work + like \verb|Attrib.config_bool|, but for types \verb|int| and + \verb|string|, respectively. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isadelimmlex +% +\endisadelimmlex +% +\isatagmlex +% +\begin{isamarkuptext}% +The following example shows how to declare and use a + Boolean configuration option called \isa{my{\isacharunderscore}flag} with constant + default value \verb|false|.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlex +{\isafoldmlex}% +% +\isadelimmlex +% +\endisadelimmlex +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ val\ {\isacharparenleft}my{\isacharunderscore}flag{\isacharcomma}\ my{\isacharunderscore}flag{\isacharunderscore}setup{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ Attrib{\isachardot}config{\isacharunderscore}bool\ {\isachardoublequote}my{\isacharunderscore}flag{\isachardoublequote}\ {\isacharparenleft}K\ false{\isacharparenright}\isanewline +{\isacharverbatimclose}\isanewline +\isacommand{setup}\isamarkupfalse% +\ my{\isacharunderscore}flag{\isacharunderscore}setup% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +Now the user can refer to \hyperlink{attribute.my-flag}{\mbox{\isa{my{\isacharunderscore}flag}}} in + declarations, while we can retrieve the current value from the + context via \verb|Config.get|.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML{\isacharunderscore}val}\isamarkupfalse% +\ {\isacharverbatimopen}\ % +\isaantiq +assert% +\endisaantiq +\ {\isacharparenleft}Config{\isachardot}get\ % +\isaantiq +context% +\endisaantiq +\ my{\isacharunderscore}flag\ {\isacharequal}\ false{\isacharparenright}\ {\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +\isanewline +\isanewline +\isacommand{declare}\isamarkupfalse% +\ {\isacharbrackleft}{\isacharbrackleft}my{\isacharunderscore}flag\ {\isacharequal}\ true{\isacharbrackright}{\isacharbrackright}\isanewline +% +\isadelimML +\isanewline +% +\endisadelimML +% +\isatagML +\isacommand{ML{\isacharunderscore}val}\isamarkupfalse% +\ {\isacharverbatimopen}\ % +\isaantiq +assert% +\endisaantiq +\ {\isacharparenleft}Config{\isachardot}get\ % +\isaantiq +context% +\endisaantiq +\ my{\isacharunderscore}flag\ {\isacharequal}\ true{\isacharparenright}\ {\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +\isanewline +% +\endisadelimML +\isanewline +\isacommand{example{\isacharunderscore}proof}\isamarkupfalse% +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{{\isacharbraceleft}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{note}\isamarkupfalse% +\ {\isacharbrackleft}{\isacharbrackleft}my{\isacharunderscore}flag\ {\isacharequal}\ false{\isacharbrackright}{\isacharbrackright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +% +\isadelimML +\ \ \ \ % +\endisadelimML +% +\isatagML +\isacommand{ML{\isacharunderscore}val}\isamarkupfalse% +\ {\isacharverbatimopen}\ % +\isaantiq +assert% +\endisaantiq +\ {\isacharparenleft}Config{\isachardot}get\ % +\isaantiq +context% +\endisaantiq +\ my{\isacharunderscore}flag\ {\isacharequal}\ false{\isacharparenright}\ {\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +\isanewline +% +\endisadelimML +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{{\isacharbraceright}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +% +\isadelimML +\ \ % +\endisadelimML +% +\isatagML +\isacommand{ML{\isacharunderscore}val}\isamarkupfalse% +\ {\isacharverbatimopen}\ % +\isaantiq +assert% +\endisaantiq +\ {\isacharparenleft}Config{\isachardot}get\ % +\isaantiq +context% +\endisaantiq +\ my{\isacharunderscore}flag\ {\isacharequal}\ true{\isacharparenright}\ {\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +\isanewline +% +\endisadelimML +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% \isamarkupsection{Names \label{sec:names}% } \isamarkuptrue% @@ -635,7 +1056,7 @@ \end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsubsection{Strings of symbols% +\isamarkupsubsection{Strings of symbols \label{sec:symbols}% } \isamarkuptrue% % @@ -669,11 +1090,11 @@ \end{enumerate} - \noindent The \isa{ident} syntax for symbol names is \isa{letter\ {\isacharparenleft}letter\ {\isacharbar}\ digit{\isacharparenright}\isactrlsup {\isacharasterisk}}, where \isa{letter\ {\isacharequal}\ A{\isachardot}{\isachardot}Za{\isachardot}{\isachardot}z} and \isa{digit\ {\isacharequal}\ {\isadigit{0}}{\isachardot}{\isachardot}{\isadigit{9}}}. There are infinitely many - regular symbols and control symbols, but a fixed collection of - standard symbols is treated specifically. For example, - ``\verb,\,\verb,,'' is classified as a letter, which means it - may occur within regular Isabelle identifiers. + The \isa{ident} syntax for symbol names is \isa{letter\ {\isacharparenleft}letter\ {\isacharbar}\ digit{\isacharparenright}\isactrlsup {\isacharasterisk}}, where \isa{letter\ {\isacharequal}\ A{\isachardot}{\isachardot}Za{\isachardot}{\isachardot}z} and \isa{digit\ {\isacharequal}\ {\isadigit{0}}{\isachardot}{\isachardot}{\isadigit{9}}}. There are infinitely many regular symbols and + control symbols, but a fixed collection of standard symbols is + treated specifically. For example, ``\verb,\,\verb,,'' is + classified as a letter, which means it may occur within regular + Isabelle identifiers. The character set underlying Isabelle symbols is 7-bit ASCII, but 8-bit character sequences are passed-through unchanged. Unicode/UCS @@ -718,22 +1139,22 @@ \begin{description} - \item \verb|Symbol.symbol| represents individual Isabelle + \item Type \verb|Symbol.symbol| represents individual Isabelle symbols. \item \verb|Symbol.explode|~\isa{str} produces a symbol list - from the packed form. This function supercedes \verb|String.explode| for virtually all purposes of manipulating text in + from the packed form. This function supersedes \verb|String.explode| for virtually all purposes of manipulating text in Isabelle!\footnote{The runtime overhead for exploded strings is mainly that of the list structure: individual symbols that happen to - be a singleton string --- which is the most common case --- do not - require extra memory in Poly/ML.} + be a singleton string do not require extra memory in Poly/ML.} \item \verb|Symbol.is_letter|, \verb|Symbol.is_digit|, \verb|Symbol.is_quasi|, \verb|Symbol.is_blank| classify standard symbols according to fixed syntactic conventions of Isabelle, cf.\ \cite{isabelle-isar-ref}. - \item \verb|Symbol.sym| is a concrete datatype that represents - the different kinds of symbols explicitly, with constructors \verb|Symbol.Char|, \verb|Symbol.Sym|, \verb|Symbol.UTF8|, \verb|Symbol.Ctrl|, \verb|Symbol.Raw|. + \item Type \verb|Symbol.sym| is a concrete datatype that + represents the different kinds of symbols explicitly, with + constructors \verb|Symbol.Char|, \verb|Symbol.Sym|, \verb|Symbol.UTF8|, \verb|Symbol.Ctrl|, \verb|Symbol.Raw|. \item \verb|Symbol.decode| converts the string representation of a symbol into the datatype version. @@ -823,8 +1244,8 @@ \item \verb|Name.skolem|~\isa{name} produces a Skolem name by adding two underscores. - \item \verb|Name.context| represents the context of already used - names; the initial value is \verb|Name.context|. + \item Type \verb|Name.context| represents the context of already + used names; the initial value is \verb|Name.context|. \item \verb|Name.declare|~\isa{name} enters a used name into the context. @@ -853,6 +1274,102 @@ % \endisadelimmlref % +\isadelimmlex +% +\endisadelimmlex +% +\isatagmlex +% +\begin{isamarkuptext}% +The following simple examples demonstrate how to produce + fresh names from the initial \verb|Name.context|.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlex +{\isafoldmlex}% +% +\isadelimmlex +% +\endisadelimmlex +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ val\ list{\isadigit{1}}\ {\isacharequal}\ Name{\isachardot}invents\ Name{\isachardot}context\ {\isachardoublequote}a{\isachardoublequote}\ {\isadigit{5}}{\isacharsemicolon}\isanewline +\ \ % +\isaantiq +assert% +\endisaantiq +\ {\isacharparenleft}list{\isadigit{1}}\ {\isacharequal}\ {\isacharbrackleft}{\isachardoublequote}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}b{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}c{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}d{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}e{\isachardoublequote}{\isacharbrackright}{\isacharparenright}{\isacharsemicolon}\isanewline +\isanewline +\ \ val\ list{\isadigit{2}}\ {\isacharequal}\isanewline +\ \ \ \ {\isacharhash}{\isadigit{1}}\ {\isacharparenleft}Name{\isachardot}variants\ {\isacharbrackleft}{\isachardoublequote}x{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}x{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}a{\isachardoublequote}{\isacharbrackright}\ Name{\isachardot}context{\isacharparenright}{\isacharsemicolon}\isanewline +\ \ % +\isaantiq +assert% +\endisaantiq +\ {\isacharparenleft}list{\isadigit{2}}\ {\isacharequal}\ {\isacharbrackleft}{\isachardoublequote}x{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}xa{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}aa{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}aa{\isachardoublequote}{\isacharbrackright}{\isacharparenright}{\isacharsemicolon}\isanewline +{\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +\medskip The same works reletively to the formal context as + follows.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{locale}\isamarkupfalse% +\ ex\ {\isacharequal}\ \isakeyword{fixes}\ a\ b\ c\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\isanewline +\isakeyword{begin}\isanewline +% +\isadelimML +\isanewline +% +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ val\ names\ {\isacharequal}\ Variable{\isachardot}names{\isacharunderscore}of\ % +\isaantiq +context% +\endisaantiq +{\isacharsemicolon}\isanewline +\isanewline +\ \ val\ list{\isadigit{1}}\ {\isacharequal}\ Name{\isachardot}invents\ names\ {\isachardoublequote}a{\isachardoublequote}\ {\isadigit{5}}{\isacharsemicolon}\isanewline +\ \ % +\isaantiq +assert% +\endisaantiq +\ {\isacharparenleft}list{\isadigit{1}}\ {\isacharequal}\ {\isacharbrackleft}{\isachardoublequote}d{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}e{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}f{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}g{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}h{\isachardoublequote}{\isacharbrackright}{\isacharparenright}{\isacharsemicolon}\isanewline +\isanewline +\ \ val\ list{\isadigit{2}}\ {\isacharequal}\isanewline +\ \ \ \ {\isacharhash}{\isadigit{1}}\ {\isacharparenleft}Name{\isachardot}variants\ {\isacharbrackleft}{\isachardoublequote}x{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}x{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}a{\isachardoublequote}{\isacharbrackright}\ names{\isacharparenright}{\isacharsemicolon}\isanewline +\ \ % +\isaantiq +assert% +\endisaantiq +\ {\isacharparenleft}list{\isadigit{2}}\ {\isacharequal}\ {\isacharbrackleft}{\isachardoublequote}x{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}xa{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}aa{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}ab{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}aa{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}ab{\isachardoublequote}{\isacharbrackright}{\isacharparenright}{\isacharsemicolon}\isanewline +{\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +\isanewline +% +\endisadelimML +\isanewline +\isacommand{end}\isamarkupfalse% +% \isamarkupsubsection{Indexed names \label{sec:indexname}% } \isamarkuptrue% @@ -901,15 +1418,14 @@ % \begin{isamarkuptext}% \begin{mldecls} - \indexdef{}{ML type}{indexname}\verb|type indexname| \\ + \indexdef{}{ML type}{indexname}\verb|type indexname = string * int| \\ \end{mldecls} \begin{description} - \item \verb|indexname| represents indexed names. This is an - abbreviation for \verb|string * int|. The second component is - usually non-negative, except for situations where \isa{{\isacharparenleft}x{\isacharcomma}\ {\isacharminus}{\isadigit{1}}{\isacharparenright}} - is used to inject basic names into this type. Other negative + \item Type \verb|indexname| represents indexed names. This is + an abbreviation for \verb|string * int|. The second component + is usually non-negative, except for situations where \isa{{\isacharparenleft}x{\isacharcomma}\ {\isacharminus}{\isadigit{1}}{\isacharparenright}} is used to inject basic names into this type. Other negative indexes should not be used. \end{description}% @@ -1017,13 +1533,12 @@ main equation of this ``chemical reaction'' when binding new entities in a context is as follows: - \smallskip + \medskip \begin{tabular}{l} \isa{binding\ {\isacharplus}\ naming\ {\isasymlongrightarrow}\ long\ name\ {\isacharplus}\ name\ space\ accesses} \end{tabular} - \smallskip - \medskip As a general principle, there is a separate name space for + \bigskip As a general principle, there is a separate name space for each kind of formal entity, e.g.\ fact, logical constant, type constructor, type class. It is usually clear from the occurrence in concrete syntax (or from the scope) which kind of entity a name @@ -1036,10 +1551,17 @@ constant \isa{c}. This technique of mapping names from one space into another requires some care in order to avoid conflicts. In particular, theorem names derived from a type constructor or type - class are better suffixed in addition to the usual qualification, - e.g.\ \isa{c{\isacharunderscore}type{\isachardot}intro} and \isa{c{\isacharunderscore}class{\isachardot}intro} for - theorems related to type \isa{c} and class \isa{c}, - respectively.% + class should get an additional suffix in addition to the usual + qualification. This leads to the following conventions for derived + names: + + \medskip + \begin{tabular}{ll} + logical entity & fact name \\\hline + constant \isa{c} & \isa{c{\isachardot}intro} \\ + type \isa{c} & \isa{c{\isacharunderscore}type{\isachardot}intro} \\ + class \isa{c} & \isa{c{\isacharunderscore}class{\isachardot}intro} \\ + \end{tabular}% \end{isamarkuptext}% \isamarkuptrue% % @@ -1078,13 +1600,14 @@ \begin{description} - \item \verb|binding| represents the abstract concept of name - bindings. + \item Type \verb|binding| represents the abstract concept of + name bindings. \item \verb|Binding.empty| is the empty binding. \item \verb|Binding.name|~\isa{name} produces a binding with base - name \isa{name}. + name \isa{name}. Note that this lacks proper source position + information; see also the ML antiquotation \hyperlink{ML antiquotation.binding}{\mbox{\isa{binding}}}. \item \verb|Binding.qualify|~\isa{mandatory\ name\ binding} prefixes qualifier \isa{name} to \isa{binding}. The \isa{mandatory} flag tells if this name component always needs to be @@ -1107,8 +1630,8 @@ representation for human-readable output, together with some formal markup that might get used in GUI front-ends, for example. - \item \verb|Name_Space.naming| represents the abstract concept of - a naming policy. + \item Type \verb|Name_Space.naming| represents the abstract + concept of a naming policy. \item \verb|Name_Space.default_naming| is the default naming policy. In a theory context, this is usually augmented by a path prefix @@ -1121,7 +1644,7 @@ name binding (usually a basic name) into the fully qualified internal name, according to the given naming policy. - \item \verb|Name_Space.T| represents name spaces. + \item Type \verb|Name_Space.T| represents name spaces. \item \verb|Name_Space.empty|~\isa{kind} and \verb|Name_Space.merge|~\isa{{\isacharparenleft}space\isactrlisub {\isadigit{1}}{\isacharcomma}\ space\isactrlisub {\isadigit{2}}{\isacharparenright}} are the canonical operations for maintaining name spaces according to theory data management @@ -1161,6 +1684,111 @@ % \endisadelimmlref % +\isadelimmlantiq +% +\endisadelimmlantiq +% +\isatagmlantiq +% +\begin{isamarkuptext}% +\begin{matharray}{rcl} + \indexdef{}{ML antiquotation}{binding}\hypertarget{ML antiquotation.binding}{\hyperlink{ML antiquotation.binding}{\mbox{\isa{binding}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \end{matharray} + + \begin{rail} + 'binding' name + ; + \end{rail} + + \begin{description} + + \item \isa{{\isacharat}{\isacharbraceleft}binding\ name{\isacharbraceright}} produces a binding with base name + \isa{name} and the source position taken from the concrete + syntax of this antiquotation. In many situations this is more + appropriate than the more basic \verb|Binding.name| function. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlantiq +{\isafoldmlantiq}% +% +\isadelimmlantiq +% +\endisadelimmlantiq +% +\isadelimmlex +% +\endisadelimmlex +% +\isatagmlex +% +\begin{isamarkuptext}% +The following example yields the source position of some + concrete binding inlined into the text.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlex +{\isafoldmlex}% +% +\isadelimmlex +% +\endisadelimmlex +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\ Binding{\isachardot}pos{\isacharunderscore}of\ % +\isaantiq +binding\ here% +\endisaantiq +\ {\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +\medskip That position can be also printed in a message as + follows.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML{\isacharunderscore}command}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ writeln\isanewline +\ \ \ \ {\isacharparenleft}{\isachardoublequote}Look\ here{\isachardoublequote}\ {\isacharcircum}\ Position{\isachardot}str{\isacharunderscore}of\ {\isacharparenleft}Binding{\isachardot}pos{\isacharunderscore}of\ % +\isaantiq +binding\ here% +\endisaantiq +{\isacharparenright}{\isacharparenright}\isanewline +{\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +This illustrates a key virtue of formalized bindings as + opposed to raw specifications of base names: the system can use this + additional information for advanced feedback given to the user.% +\end{isamarkuptext}% +\isamarkuptrue% +% \isadelimtheory % \endisadelimtheory diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/document/Proof.tex --- a/doc-src/IsarImplementation/Thy/document/Proof.tex Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/document/Proof.tex Mon Oct 25 08:08:08 2010 +0200 @@ -67,12 +67,87 @@ This twist of dependencies is also accommodated by the reverse operation of exporting results from a context: a type variable \isa{{\isasymalpha}} is considered fixed as long as it occurs in some fixed - term variable of the context. For example, exporting \isa{x{\isacharcolon}\ term{\isacharcomma}\ {\isasymalpha}{\isacharcolon}\ type\ {\isasymturnstile}\ x\isactrlisub {\isasymalpha}\ {\isacharequal}\ x\isactrlisub {\isasymalpha}} produces in the first step - \isa{x{\isacharcolon}\ term\ {\isasymturnstile}\ x\isactrlisub {\isasymalpha}\ {\isacharequal}\ x\isactrlisub {\isasymalpha}} for fixed \isa{{\isasymalpha}}, - and only in the second step \isa{{\isasymturnstile}\ {\isacharquery}x\isactrlisub {\isacharquery}\isactrlisub {\isasymalpha}\ {\isacharequal}\ {\isacharquery}x\isactrlisub {\isacharquery}\isactrlisub {\isasymalpha}} for schematic \isa{{\isacharquery}x} and \isa{{\isacharquery}{\isasymalpha}}. - - \medskip The Isabelle/Isar proof context manages the gory details of - term vs.\ type variables, with high-level principles for moving the + term variable of the context. For example, exporting \isa{x{\isacharcolon}\ term{\isacharcomma}\ {\isasymalpha}{\isacharcolon}\ type\ {\isasymturnstile}\ x\isactrlisub {\isasymalpha}\ {\isasymequiv}\ x\isactrlisub {\isasymalpha}} produces in the first step \isa{x{\isacharcolon}\ term\ {\isasymturnstile}\ x\isactrlisub {\isasymalpha}\ {\isasymequiv}\ x\isactrlisub {\isasymalpha}} for fixed \isa{{\isasymalpha}}, and only in the second step + \isa{{\isasymturnstile}\ {\isacharquery}x\isactrlisub {\isacharquery}\isactrlisub {\isasymalpha}\ {\isasymequiv}\ {\isacharquery}x\isactrlisub {\isacharquery}\isactrlisub {\isasymalpha}} for schematic \isa{{\isacharquery}x} and \isa{{\isacharquery}{\isasymalpha}}. + The following Isar source text illustrates this scenario.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{example{\isacharunderscore}proof}\isamarkupfalse% +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{{\isacharbraceleft}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\ \ % +\isamarkupcmt{all potential occurrences of some \isa{x{\isacharcolon}{\isacharcolon}{\isasymtau}} are fixed% +} +\isanewline +\ \ \ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x{\isacharcolon}{\isacharcolon}{\isacharprime}a\ {\isasymequiv}\ x{\isachardoublequoteclose}\ \ % +\isamarkupcmt{implicit type assigment by concrete occurrence% +} +\isanewline +\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}rule\ reflexive{\isacharparenright}\isanewline +\ \ \ \ \isacommand{{\isacharbraceright}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\ \ \ \ \isacommand{thm}\isamarkupfalse% +\ this\ \ % +\isamarkupcmt{result still with fixed type \isa{{\isacharprime}a}% +} +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{{\isacharbraceright}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\ \ \isacommand{thm}\isamarkupfalse% +\ this\ \ % +\isamarkupcmt{fully general result for arbitrary \isa{{\isacharquery}x{\isacharcolon}{\isacharcolon}{\isacharquery}{\isacharprime}a}% +} +\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +The Isabelle/Isar proof context manages the details of term + vs.\ type variables, with high-level principles for moving the frontier between fixed and schematic variables. The \isa{add{\isacharunderscore}fixes} operation explictly declares fixed @@ -176,9 +251,8 @@ \isatagmlex % \begin{isamarkuptext}% -The following example (in theory \hyperlink{theory.Pure}{\mbox{\isa{Pure}}}) shows - how to work with fixed term and type parameters work with - type-inference.% +The following example shows how to work with fixed term + and type parameters and with type-inference.% \end{isamarkuptext}% \isamarkuptrue% % @@ -188,14 +262,8 @@ \isadelimmlex % \endisadelimmlex -\isacommand{typedecl}\isamarkupfalse% -\ foo\ \ % -\isamarkupcmt{some basic type for testing purposes% -} -\isanewline % \isadelimML -\isanewline % \endisadelimML % @@ -217,11 +285,11 @@ \ \ val\ t{\isadigit{1}}{\isacharprime}\ {\isacharequal}\ singleton\ {\isacharparenleft}Variable{\isachardot}polymorphic\ ctxt{\isadigit{1}}{\isacharparenright}\ t{\isadigit{1}}{\isacharsemicolon}\isanewline \isanewline \ \ {\isacharparenleft}{\isacharasterisk}term\ u\ enforces\ specific\ type\ assignment{\isacharasterisk}{\isacharparenright}\isanewline -\ \ val\ u\ {\isacharequal}\ Syntax{\isachardot}read{\isacharunderscore}term\ ctxt{\isadigit{1}}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcolon}{\isacharcolon}foo{\isacharparenright}\ {\isasymequiv}\ y{\isachardoublequote}{\isacharsemicolon}\isanewline +\ \ val\ u\ {\isacharequal}\ Syntax{\isachardot}read{\isacharunderscore}term\ ctxt{\isadigit{1}}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ {\isasymequiv}\ y{\isachardoublequote}{\isacharsemicolon}\isanewline \isanewline \ \ {\isacharparenleft}{\isacharasterisk}official\ declaration\ of\ u\ {\isacharminus}{\isacharminus}\ propagates\ constraints\ etc{\isachardot}{\isacharasterisk}{\isacharparenright}\isanewline \ \ val\ ctxt{\isadigit{2}}\ {\isacharequal}\ ctxt{\isadigit{1}}\ {\isacharbar}{\isachargreater}\ Variable{\isachardot}declare{\isacharunderscore}term\ u{\isacharsemicolon}\isanewline -\ \ val\ t{\isadigit{2}}\ {\isacharequal}\ Syntax{\isachardot}read{\isacharunderscore}term\ ctxt{\isadigit{2}}\ {\isachardoublequote}x{\isachardoublequote}{\isacharsemicolon}\ \ {\isacharparenleft}{\isacharasterisk}x{\isacharcolon}{\isacharcolon}foo\ is\ enforced{\isacharasterisk}{\isacharparenright}\isanewline +\ \ val\ t{\isadigit{2}}\ {\isacharequal}\ Syntax{\isachardot}read{\isacharunderscore}term\ ctxt{\isadigit{2}}\ {\isachardoublequote}x{\isachardoublequote}{\isacharsemicolon}\ \ {\isacharparenleft}{\isacharasterisk}x{\isacharcolon}{\isacharcolon}nat\ is\ enforced{\isacharasterisk}{\isacharparenright}\isanewline {\isacharverbatimclose}% \endisatagML {\isafoldML}% @@ -263,31 +331,16 @@ \endisadelimML % \begin{isamarkuptext}% -\noindent Subsequent ML code can now work with the invented - names of \verb|x1|, \verb|x2|, \verb|x3|, without - depending on the details on the system policy for introducing these - variants. Recall that within a proof body the system always invents - fresh ``skolem constants'', e.g.\ as follows:% +The following ML code can now work with the invented names of + \verb|x1|, \verb|x2|, \verb|x3|, without depending on + the details on the system policy for introducing these variants. + Recall that within a proof body the system always invents fresh + ``skolem constants'', e.g.\ as follows:% \end{isamarkuptext}% \isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isachardoublequoteopen}PROP\ XXX{\isachardoublequoteclose}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isacharminus}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof +\isacommand{example{\isacharunderscore}proof}\isamarkupfalse% \isanewline % -\endisadelimproof -% \isadelimML \ \ % \endisadelimML @@ -318,7 +371,7 @@ \endisadelimML % \begin{isamarkuptext}% -\noindent In this situation \verb|Variable.add_fixes| and \verb|Variable.variant_fixes| are very similar, but identical name +In this situation \verb|Variable.add_fixes| and \verb|Variable.variant_fixes| are very similar, but identical name proposals given in a row are only accepted by the second version.% \end{isamarkuptext}% \isamarkuptrue% @@ -398,10 +451,11 @@ \begin{description} - \item \verb|Assumption.export| represents arbitrary export - rules, which is any function of type \verb|bool -> cterm list -> thm -> thm|, - where the \verb|bool| indicates goal mode, and the \verb|cterm list| the collection of assumptions to be discharged - simultaneously. + \item Type \verb|Assumption.export| represents arbitrary export + rules, which is any function of type \verb|bool -> cterm list|\isasep\isanewline% +\verb| -> thm -> thm|, where the \verb|bool| indicates goal mode, + and the \verb|cterm list| the collection of assumptions to be + discharged simultaneously. \item \verb|Assumption.assume|~\isa{A} turns proposition \isa{A} into a primitive assumption \isa{A\ {\isasymturnstile}\ A{\isacharprime}}, where the conclusion \isa{A{\isacharprime}} is in HHF normal form. @@ -487,10 +541,9 @@ \endisadelimML % \begin{isamarkuptext}% -\noindent Note that the variables of the resulting rule are - not generalized. This would have required to fix them properly in - the context beforehand, and export wrt.\ variables afterwards (cf.\ - \verb|Variable.export| or the combined \verb|ProofContext.export|).% +Note that the variables of the resulting rule are not + generalized. This would have required to fix them properly in the + context beforehand, and export wrt.\ variables afterwards (cf.\ \verb|Variable.export| or the combined \verb|ProofContext.export|).% \end{isamarkuptext}% \isamarkuptrue% % @@ -531,8 +584,8 @@ conclusion: the proof demonstrates that the context may be augmented by parameters and assumptions, without affecting any conclusions that do not mention these parameters. See also - \cite{isabelle-isar-ref} for the user-level \isa{{\isasymOBTAIN}} and - \isa{{\isasymGUESS}} elements. Final results, which may not refer to + \cite{isabelle-isar-ref} for the user-level \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}} and + \hyperlink{command.guess}{\mbox{\isa{\isacommand{guess}}}} elements. Final results, which may not refer to the parameters in the conclusion, need to exported explicitly into the original context.% \end{isamarkuptext}% @@ -546,10 +599,17 @@ % \begin{isamarkuptext}% \begin{mldecls} - \indexdef{}{ML}{SUBPROOF}\verb|SUBPROOF: (Subgoal.focus -> tactic) -> Proof.context -> int -> tactic| \\ - \indexdef{}{ML}{Subgoal.FOCUS}\verb|Subgoal.FOCUS: (Subgoal.focus -> tactic) -> Proof.context -> int -> tactic| \\ - \indexdef{}{ML}{Subgoal.FOCUS\_PREMS}\verb|Subgoal.FOCUS_PREMS: (Subgoal.focus -> tactic) -> Proof.context -> int -> tactic| \\ - \indexdef{}{ML}{Subgoal.FOCUS\_PARAMS}\verb|Subgoal.FOCUS_PARAMS: (Subgoal.focus -> tactic) -> Proof.context -> int -> tactic| \\ + \indexdef{}{ML}{SUBPROOF}\verb|SUBPROOF: (Subgoal.focus -> tactic) ->|\isasep\isanewline% +\verb| Proof.context -> int -> tactic| \\ + \indexdef{}{ML}{Subgoal.FOCUS}\verb|Subgoal.FOCUS: (Subgoal.focus -> tactic) ->|\isasep\isanewline% +\verb| Proof.context -> int -> tactic| \\ + \indexdef{}{ML}{Subgoal.FOCUS\_PREMS}\verb|Subgoal.FOCUS_PREMS: (Subgoal.focus -> tactic) ->|\isasep\isanewline% +\verb| Proof.context -> int -> tactic| \\ + \indexdef{}{ML}{Subgoal.FOCUS\_PARAMS}\verb|Subgoal.FOCUS_PARAMS: (Subgoal.focus -> tactic) ->|\isasep\isanewline% +\verb| Proof.context -> int -> tactic| \\ + \indexdef{}{ML}{Subgoal.focus}\verb|Subgoal.focus: Proof.context -> int -> thm -> Subgoal.focus * thm| \\ + \indexdef{}{ML}{Subgoal.focus\_prems}\verb|Subgoal.focus_prems: Proof.context -> int -> thm -> Subgoal.focus * thm| \\ + \indexdef{}{ML}{Subgoal.focus\_params}\verb|Subgoal.focus_params: Proof.context -> int -> thm -> Subgoal.focus * thm| \\ \end{mldecls} \begin{mldecls} @@ -559,8 +619,8 @@ \verb| ({prems: thm list, context: Proof.context} -> tactic) -> thm list| \\ \end{mldecls} \begin{mldecls} - \indexdef{}{ML}{Obtain.result}\verb|Obtain.result: (Proof.context -> tactic) ->|\isasep\isanewline% -\verb| thm list -> Proof.context -> ((string * cterm) list * thm list) * Proof.context| \\ + \indexdef{}{ML}{Obtain.result}\verb|Obtain.result: (Proof.context -> tactic) -> thm list ->|\isasep\isanewline% +\verb| Proof.context -> ((string * cterm) list * thm list) * Proof.context| \\ \end{mldecls} \begin{description} @@ -576,6 +636,11 @@ imported into the context, and the body tactic may introduce new subgoals and schematic variables. + \item \verb|Subgoal.focus|, \verb|Subgoal.focus_prems|, \verb|Subgoal.focus_params| extract the focus information from a goal + state in the same way as the corresponding tacticals above. This is + occasionally useful to experiment without writing actual tactics + yet. + \item \verb|Goal.prove|~\isa{ctxt\ xs\ As\ C\ tac} states goal \isa{C} in the context augmented by fixed variables \isa{xs} and assumptions \isa{As}, and applies tactic \isa{tac} to solve it. The latter may depend on the local assumptions being presented @@ -602,7 +667,157 @@ % \endisadelimmlref % +\isadelimmlex +% +\endisadelimmlex +% +\isatagmlex +% +\begin{isamarkuptext}% +The following minimal example illustrates how to access + the focus information of a structured goal state.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlex +{\isafoldmlex}% +% +\isadelimmlex +% +\endisadelimmlex +\isacommand{example{\isacharunderscore}proof}\isamarkupfalse% +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{fix}\isamarkupfalse% +\ A\ B\ C\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymAnd}x{\isachardot}\ A\ x\ {\isasymLongrightarrow}\ B\ x\ {\isasymLongrightarrow}\ C\ x{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +% +\isadelimML +\ \ \ \ % +\endisadelimML +% +\isatagML +\isacommand{ML{\isacharunderscore}val}\isamarkupfalse% +\isanewline +\ \ \ \ {\isacharverbatimopen}\isanewline +\ \ \ \ \ \ val\ {\isacharbraceleft}goal{\isacharcomma}\ context\ {\isacharequal}\ goal{\isacharunderscore}ctxt{\isacharcomma}\ {\isachardot}{\isachardot}{\isachardot}{\isacharbraceright}\ {\isacharequal}\ % +\isaantiq +Isar{\isachardot}goal% +\endisaantiq +{\isacharsemicolon}\isanewline +\ \ \ \ \ \ val\ {\isacharparenleft}focus\ as\ {\isacharbraceleft}params{\isacharcomma}\ asms{\isacharcomma}\ concl{\isacharcomma}\ {\isachardot}{\isachardot}{\isachardot}{\isacharbraceright}{\isacharcomma}\ goal{\isacharprime}{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ \ \ \ Subgoal{\isachardot}focus\ goal{\isacharunderscore}ctxt\ {\isadigit{1}}\ goal{\isacharsemicolon}\isanewline +\ \ \ \ \ \ val\ {\isacharbrackleft}A{\isacharcomma}\ B{\isacharbrackright}\ {\isacharequal}\ {\isacharhash}prems\ focus{\isacharsemicolon}\isanewline +\ \ \ \ \ \ val\ {\isacharbrackleft}{\isacharparenleft}{\isacharunderscore}{\isacharcomma}\ x{\isacharparenright}{\isacharbrackright}\ {\isacharequal}\ {\isacharhash}params\ focus{\isacharsemicolon}\isanewline +\ \ \ \ {\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +\isanewline +% +\endisadelimML +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{oops}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\medskip The next example demonstrates forward-elimination in + a local context, using \verb|Obtain.result|.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{example{\isacharunderscore}proof}\isamarkupfalse% +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{assume}\isamarkupfalse% +\ ex{\isacharcolon}\ {\isachardoublequoteopen}{\isasymexists}x{\isachardot}\ B\ x{\isachardoublequoteclose}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +% +\isadelimML +\isanewline +\ \ % +\endisadelimML +% +\isatagML +\isacommand{ML{\isacharunderscore}prf}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ \ \ val\ ctxt{\isadigit{0}}\ {\isacharequal}\ % +\isaantiq +context% +\endisaantiq +{\isacharsemicolon}\isanewline +\ \ \ \ val\ {\isacharparenleft}{\isacharparenleft}{\isacharbrackleft}{\isacharparenleft}{\isacharunderscore}{\isacharcomma}\ x{\isacharparenright}{\isacharbrackright}{\isacharcomma}\ {\isacharbrackleft}B{\isacharbrackright}{\isacharparenright}{\isacharcomma}\ ctxt{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ ctxt{\isadigit{0}}\isanewline +\ \ \ \ \ \ {\isacharbar}{\isachargreater}\ Obtain{\isachardot}result\ {\isacharparenleft}fn\ {\isacharunderscore}\ {\isacharequal}{\isachargreater}\ etac\ % +\isaantiq +thm\ exE% +\endisaantiq +\ {\isadigit{1}}{\isacharparenright}\ {\isacharbrackleft}% +\isaantiq +thm\ ex% +\endisaantiq +{\isacharbrackright}{\isacharsemicolon}\isanewline +\ \ {\isacharverbatimclose}\isanewline +\ \ \isacommand{ML{\isacharunderscore}prf}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ \ \ singleton\ {\isacharparenleft}ProofContext{\isachardot}export\ ctxt{\isadigit{1}}\ ctxt{\isadigit{0}}{\isacharparenright}\ % +\isaantiq +thm\ refl% +\endisaantiq +{\isacharsemicolon}\isanewline +\ \ {\isacharverbatimclose}\isanewline +\ \ \isacommand{ML{\isacharunderscore}prf}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ \ \ ProofContext{\isachardot}export\ ctxt{\isadigit{1}}\ ctxt{\isadigit{0}}\ {\isacharbrackleft}Thm{\isachardot}reflexive\ x{\isacharbrackright}\isanewline +\ \ \ \ \ \ handle\ ERROR\ msg\ {\isacharequal}{\isachargreater}\ {\isacharparenleft}warning\ msg{\isacharsemicolon}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isacharsemicolon}\isanewline +\ \ {\isacharverbatimclose}\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagML +{\isafoldML}% +% +\isadelimML +\isanewline +% +\endisadelimML +% \isadelimtheory +\isanewline % \endisadelimtheory % @@ -613,9 +828,9 @@ {\isafoldtheory}% % \isadelimtheory +\isanewline % \endisadelimtheory -\isanewline \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/document/Syntax.tex --- a/doc-src/IsarImplementation/Thy/document/Syntax.tex Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/document/Syntax.tex Mon Oct 25 08:08:08 2010 +0200 @@ -27,7 +27,46 @@ \end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsection{Parsing and printing% +\isamarkupsection{Reading and pretty printing \label{sec:read-print}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +FIXME% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{Syntax.read\_typ}\verb|Syntax.read_typ: Proof.context -> string -> typ| \\ + \indexdef{}{ML}{Syntax.read\_term}\verb|Syntax.read_term: Proof.context -> string -> term| \\ + \indexdef{}{ML}{Syntax.read\_prop}\verb|Syntax.read_prop: Proof.context -> string -> term| \\ + \indexdef{}{ML}{Syntax.pretty\_typ}\verb|Syntax.pretty_typ: Proof.context -> typ -> Pretty.T| \\ + \indexdef{}{ML}{Syntax.pretty\_term}\verb|Syntax.pretty_term: Proof.context -> term -> Pretty.T| \\ + \end{mldecls} + + \begin{description} + + \item FIXME + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsection{Parsing and unparsing \label{sec:parse-unparse}% } \isamarkuptrue% % @@ -36,10 +75,118 @@ \end{isamarkuptext}% \isamarkuptrue% % +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{Syntax.parse\_typ}\verb|Syntax.parse_typ: Proof.context -> string -> typ| \\ + \indexdef{}{ML}{Syntax.parse\_term}\verb|Syntax.parse_term: Proof.context -> string -> term| \\ + \indexdef{}{ML}{Syntax.parse\_prop}\verb|Syntax.parse_prop: Proof.context -> string -> term| \\ + \indexdef{}{ML}{Syntax.unparse\_typ}\verb|Syntax.unparse_typ: Proof.context -> typ -> Pretty.T| \\ + \indexdef{}{ML}{Syntax.unparse\_term}\verb|Syntax.unparse_term: Proof.context -> term -> Pretty.T| \\ + \end{mldecls} + + \begin{description} + + \item FIXME + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% \isamarkupsection{Checking and unchecking \label{sec:term-check}% } \isamarkuptrue% % +\begin{isamarkuptext}% +FIXME% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{Syntax.check\_typs}\verb|Syntax.check_typs: Proof.context -> typ list -> typ list| \\ + \indexdef{}{ML}{Syntax.check\_terms}\verb|Syntax.check_terms: Proof.context -> term list -> term list| \\ + \indexdef{}{ML}{Syntax.check\_props}\verb|Syntax.check_props: Proof.context -> term list -> term list| \\ + \indexdef{}{ML}{Syntax.uncheck\_typs}\verb|Syntax.uncheck_typs: Proof.context -> typ list -> typ list| \\ + \indexdef{}{ML}{Syntax.uncheck\_terms}\verb|Syntax.uncheck_terms: Proof.context -> term list -> term list| \\ + \end{mldecls} + + \begin{description} + + \item FIXME + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsection{Syntax translations% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +FIXME% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlantiq +% +\endisadelimmlantiq +% +\isatagmlantiq +% +\begin{isamarkuptext}% +\begin{matharray}{rcl} + \indexdef{}{ML antiquotation}{class\_syntax}\hypertarget{ML antiquotation.class-syntax}{\hyperlink{ML antiquotation.class-syntax}{\mbox{\isa{class{\isacharunderscore}syntax}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{type\_syntax}\hypertarget{ML antiquotation.type-syntax}{\hyperlink{ML antiquotation.type-syntax}{\mbox{\isa{type{\isacharunderscore}syntax}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{const\_syntax}\hypertarget{ML antiquotation.const-syntax}{\hyperlink{ML antiquotation.const-syntax}{\mbox{\isa{const{\isacharunderscore}syntax}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \indexdef{}{ML antiquotation}{syntax\_const}\hypertarget{ML antiquotation.syntax-const}{\hyperlink{ML antiquotation.syntax-const}{\mbox{\isa{syntax{\isacharunderscore}const}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\ + \end{matharray} + + \begin{rail} + ('class_syntax' | 'type_syntax' | 'const_syntax' | 'syntax_const') name + ; + \end{rail} + + \begin{description} + + \item FIXME + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlantiq +{\isafoldmlantiq}% +% +\isadelimmlantiq +% +\endisadelimmlantiq +% \isadelimtheory % \endisadelimtheory diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/Thy/document/Tactic.tex --- a/doc-src/IsarImplementation/Thy/document/Tactic.tex Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/Thy/document/Tactic.tex Mon Oct 25 08:08:08 2010 +0200 @@ -116,7 +116,7 @@ % \endisadelimmlref % -\isamarkupsection{Tactics% +\isamarkupsection{Tactics\label{sec:tactics}% } \isamarkuptrue% % @@ -220,13 +220,14 @@ \begin{description} - \item \verb|tactic| represents tactics. The well-formedness - conditions described above need to be observed. See also \hyperlink{file.~~/src/Pure/General/seq.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}General{\isacharslash}seq{\isachardot}ML}}}} for the underlying implementation of - lazy sequences. + \item Type \verb|tactic| represents tactics. The + well-formedness conditions described above need to be observed. See + also \hyperlink{file.~~/src/Pure/General/seq.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}General{\isacharslash}seq{\isachardot}ML}}}} for the underlying + implementation of lazy sequences. - \item \verb|int -> tactic| represents tactics with explicit - subgoal addressing, with well-formedness conditions as described - above. + \item Type \verb|int -> tactic| represents tactics with + explicit subgoal addressing, with well-formedness conditions as + described above. \item \verb|no_tac| is a tactic that always fails, returning the empty sequence. @@ -480,7 +481,10 @@ composition, disjunction (choice), iteration, or goal addressing. Various search strategies may be expressed via tacticals. - \medskip FIXME% + \medskip FIXME + + \medskip The chapter on tacticals in \cite{isabelle-ref} is still + applicable, despite a few outdated details.% \end{isamarkuptext}% \isamarkuptrue% % diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/implementation.tex --- a/doc-src/IsarImplementation/implementation.tex Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/implementation.tex Mon Oct 25 08:08:08 2010 +0200 @@ -4,6 +4,7 @@ \usepackage{../iman,../extra,../isar,../proof} \usepackage[nohyphen,strings]{../underscore} \usepackage{../isabelle,../isabellesym} +\usepackage{../ttbox,../rail,../railsetup} \usepackage{style} \usepackage{../pdfsetup} @@ -22,6 +23,11 @@ \makeindex +\railterm{lbrace,rbrace,atsign} +\railalias{lbracesym}{\isasymlbrace}\railterm{lbracesym} +\railalias{rbracesym}{\isasymrbrace}\railterm{rbracesym} +\railalias{dots}{\isasymdots}\railterm{dots} + \begin{document} @@ -76,18 +82,18 @@ \listoffigures \clearfirst +\setcounter{chapter}{-1} + +\input{Thy/document/ML.tex} \input{Thy/document/Prelim.tex} \input{Thy/document/Logic.tex} +\input{Thy/document/Syntax.tex} \input{Thy/document/Tactic.tex} \input{Thy/document/Proof.tex} -\input{Thy/document/Syntax.tex} \input{Thy/document/Isar.tex} \input{Thy/document/Local_Theory.tex} \input{Thy/document/Integration.tex} -\appendix -\input{Thy/document/ML.tex} - \begingroup \tocentry{\bibname} \bibliographystyle{abbrv} \small\raggedright\frenchspacing diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarImplementation/style.sty --- a/doc-src/IsarImplementation/style.sty Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarImplementation/style.sty Mon Oct 25 08:08:08 2010 +0200 @@ -20,17 +20,24 @@ \sloppy \binperiod +\parindent 0pt\parskip 0.5ex + \renewcommand{\isadigit}[1]{\isamath{#1}} \newenvironment{mldecls}{\par\noindent\begingroup\footnotesize\def\isanewline{\\}\begin{tabular}{l}}{\end{tabular}\smallskip\endgroup} \isafoldtag{FIXME} + \isakeeptag{mlref} -\renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Reference}\begingroup\def\isastyletext{\rm}\small} -\renewcommand{\endisatagmlref}{\endgroup} +\renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{ML}~~}Reference}} +\renewcommand{\endisatagmlref}{} + +\isakeeptag{mlantiq} +\renewcommand{\isatagmlantiq}{\subsection*{\makebox[0pt][r]{\fbox{ML}~~}Antiquotations}} +\renewcommand{\endisatagmlantiq}{} \isakeeptag{mlex} -\renewcommand{\isatagmlex}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Examples}} +\renewcommand{\isatagmlex}{\subsection*{\makebox[0pt][r]{\fbox{ML}~~}Examples}} \renewcommand{\endisatagmlex}{} \renewcommand{\isatagML}{\begingroup\isabellestyle{default}\isastyle\def\isadigit##1{##1}} diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarRef/Thy/Misc.thy --- a/doc-src/IsarRef/Thy/Misc.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarRef/Thy/Misc.thy Mon Oct 25 08:08:08 2010 +0200 @@ -143,7 +143,7 @@ \end{matharray} \begin{rail} - ('cd' | 'use\_thy' | 'update\_thy') name + ('cd' | 'use\_thy') name ; \end{rail} @@ -159,6 +159,13 @@ since loading of theories is done automatically as required. \end{description} + + %FIXME proper place (!?) + Isabelle file specification may contain path variables (e.g.\ + @{verbatim "$ISABELLE_HOME"}) that are expanded accordingly. Note + that @{verbatim "~"} abbreviates @{verbatim "$HOME"}, and @{verbatim + "~~"} abbreviates @{verbatim "$ISABELLE_HOME"}. The general syntax + for path specifications follows POSIX conventions. *} end diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarRef/Thy/Spec.thy --- a/doc-src/IsarRef/Thy/Spec.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarRef/Thy/Spec.thy Mon Oct 25 08:08:08 2010 +0200 @@ -837,11 +837,6 @@ @{command_def "attribute_setup"} & : & @{text "theory \ theory"} \\ \end{matharray} - \begin{mldecls} - @{index_ML bind_thms: "string * thm list -> unit"} \\ - @{index_ML bind_thm: "string * thm -> unit"} \\ - \end{mldecls} - \begin{rail} 'use' name ; @@ -918,20 +913,6 @@ let val context' = context in context' end)) *} "my declaration" -text {* - \begin{description} - - \item @{ML bind_thms}~@{text "(name, thms)"} stores a list of - theorems produced in ML both in the theory context and the ML - toplevel, associating it with the provided name. Theorems are put - into a global ``standard'' format before being stored. - - \item @{ML bind_thm} is similar to @{ML bind_thms} but refers to a - singleton theorem. - - \end{description} -*} - section {* Primitive specification elements *} @@ -973,7 +954,7 @@ Usually the default sort is only changed when defining a new object-logic. For example, the default sort in Isabelle/HOL is - @{text type}, the class of all HOL types. %FIXME sort antiq? + @{class type}, the class of all HOL types. When merging theories, the default sorts of the parents are logically intersected, i.e.\ the representations as lists of classes diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarRef/Thy/document/Misc.tex --- a/doc-src/IsarRef/Thy/document/Misc.tex Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarRef/Thy/document/Misc.tex Mon Oct 25 08:08:08 2010 +0200 @@ -164,7 +164,7 @@ \end{matharray} \begin{rail} - ('cd' | 'use\_thy' | 'update\_thy') name + ('cd' | 'use\_thy') name ; \end{rail} @@ -179,7 +179,13 @@ These system commands are scarcely used when working interactively, since loading of theories is done automatically as required. - \end{description}% + \end{description} + + %FIXME proper place (!?) + Isabelle file specification may contain path variables (e.g.\ + \verb|$ISABELLE_HOME|) that are expanded accordingly. Note + that \verb|~| abbreviates \verb|$HOME|, and \verb|~~| abbreviates \verb|$ISABELLE_HOME|. The general syntax + for path specifications follows POSIX conventions.% \end{isamarkuptext}% \isamarkuptrue% % diff -r 82873a6f2b81 -r 0d579da1902a doc-src/IsarRef/Thy/document/Spec.tex --- a/doc-src/IsarRef/Thy/document/Spec.tex Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/IsarRef/Thy/document/Spec.tex Mon Oct 25 08:08:08 2010 +0200 @@ -859,11 +859,6 @@ \indexdef{}{command}{attribute\_setup}\hypertarget{command.attribute-setup}{\hyperlink{command.attribute-setup}{\mbox{\isa{\isacommand{attribute{\isacharunderscore}setup}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\ \end{matharray} - \begin{mldecls} - \indexdef{}{ML}{bind\_thms}\verb|bind_thms: string * thm list -> unit| \\ - \indexdef{}{ML}{bind\_thm}\verb|bind_thm: string * thm -> unit| \\ - \end{mldecls} - \begin{rail} 'use' name ; @@ -951,21 +946,6 @@ % \endisadelimML % -\begin{isamarkuptext}% -\begin{description} - - \item \verb|bind_thms|~\isa{{\isachardoublequote}{\isacharparenleft}name{\isacharcomma}\ thms{\isacharparenright}{\isachardoublequote}} stores a list of - theorems produced in ML both in the theory context and the ML - toplevel, associating it with the provided name. Theorems are put - into a global ``standard'' format before being stored. - - \item \verb|bind_thm| is similar to \verb|bind_thms| but refers to a - singleton theorem. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% \isamarkupsection{Primitive specification elements% } \isamarkuptrue% @@ -1010,7 +990,7 @@ Usually the default sort is only changed when defining a new object-logic. For example, the default sort in Isabelle/HOL is - \isa{type}, the class of all HOL types. %FIXME sort antiq? + \isa{type}, the class of all HOL types. When merging theories, the default sorts of the parents are logically intersected, i.e.\ the representations as lists of classes diff -r 82873a6f2b81 -r 0d579da1902a doc-src/Ref/introduction.tex --- a/doc-src/Ref/introduction.tex Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/Ref/introduction.tex Mon Oct 25 08:08:08 2010 +0200 @@ -1,82 +1,5 @@ -\chapter{Basic Use of Isabelle}\index{sessions|(} - -\section{Basic interaction with Isabelle} -\index{starting up|bold}\nobreak -% -We assume that your local Isabelle administrator (this might be you!) has -already installed the Isabelle system together with appropriate object-logics. - -\medskip Let $\langle isabellehome \rangle$ denote the location where -the distribution has been installed. To run Isabelle from a the shell -prompt within an ordinary text terminal session, simply type -\begin{ttbox} -\({\langle}isabellehome{\rangle}\)/bin/isabelle -\end{ttbox} -This should start an interactive \ML{} session with the default object-logic -(usually HOL) already pre-loaded. - -Subsequently, we assume that the \texttt{isabelle} executable is determined -automatically by the shell, e.g.\ by adding {\tt \(\langle isabellehome - \rangle\)/bin} to your search path.\footnote{Depending on your installation, - there may be stand-alone binaries located in some global directory such as - \texttt{/usr/bin}. Do not attempt to copy {\tt \(\langle isabellehome - \rangle\)/bin/isabelle}, though! See \texttt{isabelle install} in - \emph{The Isabelle System Manual} of how to do this properly.} - -\medskip - -The object-logic image to load may be also specified explicitly as an argument -to the {\tt isabelle} command, e.g. -\begin{ttbox} -isabelle FOL -\end{ttbox} -This should put you into the world of polymorphic first-order logic (assuming -that an image of FOL has been pre-built). - -\index{saving your session|bold} Isabelle provides no means of storing -theorems or internal proof objects on files. Theorems are simply part of the -\ML{} state. To save your work between sessions, you may dump the \ML{} -system state to a file. This is done automatically when ending the session -normally (e.g.\ by typing control-D), provided that the image has been opened -\emph{writable} in the first place. The standard object-logic images are -usually read-only, so you have to create a private working copy first. For -example, the following shell command puts you into a writable Isabelle session -of name \texttt{Foo} that initially contains just plain HOL: -\begin{ttbox} -isabelle HOL Foo -\end{ttbox} -Ending the \texttt{Foo} session with control-D will cause the complete -\ML-world to be saved somewhere in your home directory\footnote{The default - location is in \texttt{\~\relax/isabelle/heaps}, but this depends on your - local configuration.}. Make sure there is enough space available! Then one -may later continue at exactly the same point by running -\begin{ttbox} -isabelle Foo -\end{ttbox} - -\medskip Saving the {\ML} state is not enough. Record, on a file, the -top-level commands that generate your theories and proofs. Such a record -allows you to replay the proofs whenever required, for instance after making -minor changes to the axioms. Ideally, these sources will be somewhat -intelligible to others as a formal description of your work. - -It is good practice to put all source files that constitute a separate -Isabelle session into an individual directory, together with an {\ML} file -called \texttt{ROOT.ML} that contains appropriate commands to load all other -files required. Running \texttt{isabelle} with option \texttt{-u} -automatically loads \texttt{ROOT.ML} on entering the session. The -\texttt{isabelle usedir} utility provides some more options to manage Isabelle -sessions, such as automatic generation of theory browsing information. - -\medskip More details about the \texttt{isabelle} and \texttt{isabelle} -commands may be found in \emph{The Isabelle System Manual}. - -\medskip There are more comfortable user interfaces than the bare-bones \ML{} -top-level run from a text terminal. The \texttt{Isabelle} executable (note -the capital I) runs one such interface, depending on your local configuration. -Again, see \emph{The Isabelle System Manual} for more information. - +\chapter{Basic Use of Isabelle} \section{Ending a session} \begin{ttbox} @@ -99,149 +22,6 @@ Typing control-D also finishes the session in essentially the same way as the sequence {\tt commit(); quit();} would. - -\section{Reading ML files} -\index{files!reading} -\begin{ttbox} -cd : string -> unit -pwd : unit -> string -use : string -> unit -time_use : string -> unit -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{cd} "{\it dir}";] changes the current directory to - {\it dir}. This is the default directory for reading files. - -\item[\ttindexbold{pwd}();] returns the full path of the current - directory. - -\item[\ttindexbold{use} "$file$";] -reads the given {\it file} as input to the \ML{} session. Reading a file -of Isabelle commands is the usual way of replaying a proof. - -\item[\ttindexbold{time_use} "$file$";] -performs {\tt use~"$file$"} and prints the total execution time. -\end{ttdescription} - -The $dir$ and $file$ specifications of the \texttt{cd} and \texttt{use} -commands may contain path variables (e.g.\ \texttt{\$ISABELLE_HOME}) that are -expanded appropriately. Note that \texttt{\~\relax} abbreviates -\texttt{\$HOME}, and \texttt{\~\relax\~\relax} abbreviates -\texttt{\$ISABELLE_HOME}\index{*\$ISABELLE_HOME}. The syntax for path -specifications follows Unix conventions. - - -\section{Reading theories}\label{sec:intro-theories} -\index{theories!reading} - -In Isabelle, any kind of declarations, definitions, etc.\ are organized around -named \emph{theory} objects. Logical reasoning always takes place within a -certain theory context, which may be switched at any time. Theory $name$ is -defined by a theory file $name$\texttt{.thy}, containing declarations of -\texttt{consts}, \texttt{types}, \texttt{defs}, etc.\ (see -\S\ref{sec:ref-defining-theories} for more details on concrete syntax). -Furthermore, there may be an associated {\ML} file $name$\texttt{.ML} with -proof scripts that are to be run in the context of the theory. - -\begin{ttbox} -context : theory -> unit -the_context : unit -> theory -theory : string -> theory -use_thy : string -> unit -time_use_thy : string -> unit -update_thy : string -> unit -\end{ttbox} - -\begin{ttdescription} - -\item[\ttindexbold{context} $thy$;] switches the current theory context. Any - subsequent command with ``implicit theory argument'' (e.g.\ \texttt{Goal}) - will refer to $thy$ as its theory. - -\item[\ttindexbold{the_context}();] obtains the current theory context, or - raises an error if absent. - -\item[\ttindexbold{theory} "$name$";] retrieves the theory called $name$ from - the internal data\-base of loaded theories, raising an error if absent. - -\item[\ttindexbold{use_thy} "$name$";] reads theory $name$ from the file - system, looking for $name$\texttt{.thy} and $name$\texttt{.ML} (the latter - being optional). It also ensures that all parent theories are loaded as - well. In case some older versions have already been present, - \texttt{use_thy} only tries to reload $name$ itself, but is content with any - version of its ancestors. - -\item[\ttindexbold{time_use_thy} "$name$";] same as \texttt{use_thy}, but - reports the time taken to process the actual theory parts and {\ML} files - separately. - -\item[\ttindexbold{update_thy} "$name$";] is similar to \texttt{use_thy}, but - ensures that theory $name$ is fully up-to-date with respect to the file - system --- apart from theory $name$ itself, any of its ancestors may be - reloaded as well. - -\end{ttdescription} - -Note that theories of pre-built logic images (e.g.\ HOL) are marked as -\emph{finished} and cannot be updated any more. See \S\ref{sec:more-theories} -for further information on Isabelle's theory loader. - - -\section{Setting flags} -\begin{ttbox} -set : bool ref -> bool -reset : bool ref -> bool -toggle : bool ref -> bool -\end{ttbox}\index{*set}\index{*reset}\index{*toggle} -These are some shorthands for manipulating boolean references. The new -value is returned. - - -\section{Diagnostic messages} -\index{error messages} -\index{warnings} - -Isabelle conceptually provides three output channels for different kinds of -messages: ordinary text, warnings, errors. Depending on the user interface -involved, these messages may appear in different text styles or colours. - -The default setup of an \texttt{isabelle} terminal session is as -follows: plain output of ordinary text, warnings prefixed by -\texttt{\#\#\#}'s, errors prefixed by \texttt{***}'s. For example, a -typical warning would look like this: -\begin{ttbox} -\#\#\# Beware the Jabberwock, my son! -\#\#\# The jaws that bite, the claws that catch! -\#\#\# Beware the Jubjub Bird, and shun -\#\#\# The frumious Bandersnatch! -\end{ttbox} - -\texttt{ML} programs may output diagnostic messages using the -following functions: -\begin{ttbox} -writeln : string -> unit -warning : string -> unit -error : string -> 'a -\end{ttbox} -Note that \ttindex{error} fails by raising exception \ttindex{ERROR} -after having output the text, while \ttindex{writeln} and -\ttindex{warning} resume normal program execution. - - -\section{Timing} -\index{timing statistics}\index{proofs!timing} -\begin{ttbox} -timing: bool ref \hfill{\bf initially false} -\end{ttbox} - -\begin{ttdescription} -\item[set \ttindexbold{timing};] enables global timing in Isabelle. - This information is compiler-dependent. -\end{ttdescription} - -\index{sessions|)} - - %%% Local Variables: %%% mode: latex %%% TeX-master: "ref" diff -r 82873a6f2b81 -r 0d579da1902a doc-src/Ref/theories.tex --- a/doc-src/Ref/theories.tex Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/Ref/theories.tex Mon Oct 25 08:08:08 2010 +0200 @@ -7,37 +7,6 @@ Isabelle's theory loader manages dependencies of the internal graph of theory nodes (the \emph{theory database}) and the external view of the file system. -See \S\ref{sec:intro-theories} for its most basic commands, such as -\texttt{use_thy}. There are a few more operations available. - -\begin{ttbox} -use_thy_only : string -> unit -update_thy_only : string -> unit -touch_thy : string -> unit -remove_thy : string -> unit -delete_tmpfiles : bool ref \hfill\textbf{initially true} -\end{ttbox} - -\begin{ttdescription} -\item[\ttindexbold{use_thy_only} "$name$";] is similar to \texttt{use_thy}, - but processes the actual theory file $name$\texttt{.thy} only, ignoring - $name$\texttt{.ML}. This might be useful in replaying proof scripts by hand - from the very beginning, starting with the fresh theory. - -\item[\ttindexbold{update_thy_only} "$name$";] is similar to - \texttt{update_thy}, but processes the actual theory file - $name$\texttt{.thy} only, ignoring $name$\texttt{.ML}. - -\item[\ttindexbold{touch_thy} "$name$";] marks theory node $name$ of the - internal graph as outdated. While the theory remains usable, subsequent - operations such as \texttt{use_thy} may cause a reload. - -\item[\ttindexbold{remove_thy} "$name$";] deletes theory node $name$, - including \emph{all} of its descendants. Beware! This is a quick way to - dispose a large number of theories at once. Note that {\ML} bindings to - theorems etc.\ of removed theories may still persist. - -\end{ttdescription} \medskip Theory and {\ML} files are located by skimming through the directories listed in Isabelle's internal load path, which merely contains the @@ -83,133 +52,18 @@ \subsection{*Theory inclusion} \begin{ttbox} -subthy : theory * theory -> bool -eq_thy : theory * theory -> bool transfer : theory -> thm -> thm -transfer_sg : Sign.sg -> thm -> thm \end{ttbox} -Inclusion and equality of theories is determined by unique -identification stamps that are created when declaring new components. -Theorems contain a reference to the theory (actually to its signature) -they have been derived in. Transferring theorems to super theories -has no logical significance, but may affect some operations in subtle -ways (e.g.\ implicit merges of signatures when applying rules, or -pretty printing of theorems). +Transferring theorems to super theories has no logical significance, +but may affect some operations in subtle ways (e.g.\ implicit merges +of signatures when applying rules, or pretty printing of theorems). \begin{ttdescription} -\item[\ttindexbold{subthy} ($thy@1$, $thy@2$)] determines if $thy@1$ - is included in $thy@2$ wrt.\ identification stamps. - -\item[\ttindexbold{eq_thy} ($thy@1$, $thy@2$)] determines if $thy@1$ - is exactly the same as $thy@2$. - \item[\ttindexbold{transfer} $thy$ $thm$] transfers theorem $thm$ to theory $thy$, provided the latter includes the theory of $thm$. -\item[\ttindexbold{transfer_sg} $sign$ $thm$] is similar to - \texttt{transfer}, but identifies the super theory via its - signature. - -\end{ttdescription} - - -\subsection{*Primitive theories} -\begin{ttbox} -ProtoPure.thy : theory -Pure.thy : theory -CPure.thy : theory -\end{ttbox} -\begin{description} -\item[\ttindexbold{ProtoPure.thy}, \ttindexbold{Pure.thy}, - \ttindexbold{CPure.thy}] contain the syntax and signature of the - meta-logic. There are basically no axioms: meta-level inferences - are carried out by \ML\ functions. \texttt{Pure} and \texttt{CPure} - just differ in their concrete syntax of prefix function application: - $t(u@1, \ldots, u@n)$ in \texttt{Pure} vs.\ $t\,u@1,\ldots\,u@n$ in - \texttt{CPure}. \texttt{ProtoPure} is their common parent, - containing no syntax for printing prefix applications at all! - -%% FIXME -%\item [\ttindexbold{extend_theory} $thy$ {\tt"}$T${\tt"} $\cdots$] extends -% the theory $thy$ with new types, constants, etc. $T$ identifies the theory -% internally. When a theory is redeclared, say to change an incorrect axiom, -% bindings to the old axiom may persist. Isabelle ensures that the old and -% new theories are not involved in the same proof. Attempting to combine -% different theories having the same name $T$ yields the fatal error -%extend_theory : theory -> string -> \(\cdots\) -> theory -%\begin{ttbox} -%Attempt to merge different versions of theory: \(T\) -%\end{ttbox} -\end{description} - -%% FIXME -%\item [\ttindexbold{extend_theory} $thy$ {\tt"}$T${\tt"} -% ($classes$, $default$, $types$, $arities$, $consts$, $sextopt$) $rules$] -%\hfill\break %%% include if line is just too short -%is the \ML{} equivalent of the following theory definition: -%\begin{ttbox} -%\(T\) = \(thy\) + -%classes \(c\) < \(c@1\),\(\dots\),\(c@m\) -% \dots -%default {\(d@1,\dots,d@r\)} -%types \(tycon@1\),\dots,\(tycon@i\) \(n\) -% \dots -%arities \(tycon@1'\),\dots,\(tycon@j'\) :: (\(s@1\),\dots,\(s@n\))\(c\) -% \dots -%consts \(b@1\),\dots,\(b@k\) :: \(\tau\) -% \dots -%rules \(name\) \(rule\) -% \dots -%end -%\end{ttbox} -%where -%\begin{tabular}[t]{l@{~=~}l} -%$classes$ & \tt[("$c$",["$c@1$",\dots,"$c@m$"]),\dots] \\ -%$default$ & \tt["$d@1$",\dots,"$d@r$"]\\ -%$types$ & \tt[([$tycon@1$,\dots,$tycon@i$], $n$),\dots] \\ -%$arities$ & \tt[([$tycon'@1$,\dots,$tycon'@j$], ([$s@1$,\dots,$s@n$],$c$)),\dots] -%\\ -%$consts$ & \tt[([$b@1$,\dots,$b@k$],$\tau$),\dots] \\ -%$rules$ & \tt[("$name$",$rule$),\dots] -%\end{tabular} - - -\subsection{Inspecting a theory}\label{sec:inspct-thy} -\index{theories!inspecting|bold} -\begin{ttbox} -print_syntax : theory -> unit -print_theory : theory -> unit -parents_of : theory -> theory list -ancestors_of : theory -> theory list -sign_of : theory -> Sign.sg -Sign.stamp_names_of : Sign.sg -> string list -\end{ttbox} -These provide means of viewing a theory's components. -\begin{ttdescription} -\item[\ttindexbold{print_syntax} $thy$] prints the syntax of $thy$ - (grammar, macros, translation functions etc., see - page~\pageref{pg:print_syn} for more details). - -\item[\ttindexbold{print_theory} $thy$] prints the logical parts of - $thy$, excluding the syntax. - -\item[\ttindexbold{parents_of} $thy$] returns the direct ancestors - of~$thy$. - -\item[\ttindexbold{ancestors_of} $thy$] returns all ancestors of~$thy$ - (not including $thy$ itself). - -\item[\ttindexbold{sign_of} $thy$] returns the signature associated - with~$thy$. It is useful with functions like {\tt - read_instantiate_sg}, which take a signature as an argument. - -\item[\ttindexbold{Sign.stamp_names_of} $sg$]\index{signatures} - returns the names of the identification \rmindex{stamps} of ax - signature. These coincide with the names of its full ancestry - including that of $sg$ itself. - \end{ttdescription} diff -r 82873a6f2b81 -r 0d579da1902a doc-src/antiquote_setup.ML --- a/doc-src/antiquote_setup.ML Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/antiquote_setup.ML Mon Oct 25 08:08:08 2010 +0200 @@ -44,18 +44,25 @@ local fun ml_val (txt1, "") = "fn _ => (" ^ txt1 ^ ");" - | ml_val (txt1, txt2) = "fn _ => (" ^ txt1 ^ ": " ^ txt2 ^ ");"; + | ml_val (txt1, txt2) = "fn _ => (" ^ txt1 ^ " : " ^ txt2 ^ ");"; fun ml_type (txt1, "") = "val _ = NONE : (" ^ txt1 ^ ") option;" | ml_type (txt1, txt2) = "val _ = [NONE : (" ^ txt1 ^ ") option, NONE : (" ^ txt2 ^ ") option];"; -fun ml_exn (txt1, "") = "fn _ => (" ^ txt1 ^ ": exn);" - | ml_exn (txt1, txt2) = "fn _ => (" ^ txt1 ^ ": " ^ txt2 ^ " -> exn);"; +fun ml_exn (txt1, "") = "fn _ => (" ^ txt1 ^ " : exn);" + | ml_exn (txt1, txt2) = "fn _ => (" ^ txt1 ^ " : " ^ txt2 ^ " -> exn);"; fun ml_structure (txt, _) = "functor XXX() = struct structure XX = " ^ txt ^ " end;"; fun ml_functor (txt, _) = "ML_Env.check_functor " ^ ML_Syntax.print_string txt; +val is_name = ML_Lex.kind_of #> (fn kind => kind = ML_Lex.Ident orelse kind = ML_Lex.LongIdent); + +fun ml_name txt = + (case filter is_name (ML_Lex.tokenize txt) of + toks as [_] => ML_Lex.flatten toks + | _ => error ("Single ML name expected in input: " ^ quote txt)); + fun index_ml name kind ml = Thy_Output.antiquotation name (Scan.lift (Args.name -- Scan.optional (Args.colon |-- Args.name) "")) (fn {context = ctxt, ...} => fn (txt1, txt2) => @@ -64,12 +71,13 @@ if txt2 = "" then txt1 else if kind = "type" then txt1 ^ " = " ^ txt2 else if kind = "exception" then txt1 ^ " of " ^ txt2 - else txt1 ^ ": " ^ txt2; + else if Syntax.is_identifier (Long_Name.base_name (ml_name txt1)) then txt1 ^ ": " ^ txt2 + else txt1 ^ " : " ^ txt2; val txt' = if kind = "" then txt else kind ^ " " ^ txt; val _ = ML_Context.eval_text_in (SOME ctxt) false Position.none (ml (txt1, txt2)); (* ML_Lex.read (!?) *) val kind' = if kind = "" then "ML" else "ML " ^ kind; in - "\\indexdef{}{" ^ kind' ^ "}{" ^ clean_string txt1 ^ "}" ^ + "\\indexdef{}{" ^ kind' ^ "}{" ^ clean_string (ml_name txt1) ^ "}" ^ (txt' |> (if Config.get ctxt Thy_Output.quotes then quote else I) |> (if Config.get ctxt Thy_Output.display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}" @@ -184,6 +192,7 @@ val _ = entity_antiqs (K check_tool) "isatt" "tool"; val _ = entity_antiqs (K (File.exists o Path.explode)) "isatt" "file"; val _ = entity_antiqs (K (can Thy_Info.get_theory)) "" "theory"; +val _ = entity_antiqs no_check "" "ML antiquotation"; (* FIXME proper check *) end; diff -r 82873a6f2b81 -r 0d579da1902a doc-src/iman.sty --- a/doc-src/iman.sty Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/iman.sty Mon Oct 25 08:08:08 2010 +0200 @@ -82,7 +82,8 @@ \def\dbend{\vtop to 0pt{\vss\hbox{\Huge\bf!}\vss}} \newenvironment{warn}{\medskip\medbreak\begingroup \clubpenalty=10000 \small %%WAS\baselineskip=0.9\baselineskip - \noindent \hangindent\parindent \hangafter=-2 + \noindent \ifdim\parindent > 0pt\hangindent\parindent\else\hangindent1.5em\fi + \hangafter=-2 \hbox to0pt{\hskip-\hangindent\dbend\hfill}\ignorespaces}% {\par\endgroup\medbreak} diff -r 82873a6f2b81 -r 0d579da1902a doc-src/manual.bib --- a/doc-src/manual.bib Fri Oct 22 18:38:59 2010 +0200 +++ b/doc-src/manual.bib Mon Oct 25 08:08:08 2010 +0200 @@ -319,6 +319,18 @@ %C +@InProceedings{Chaieb-Wenzel:2007, + author = {Amine Chaieb and Makarius Wenzel}, + title = {Context aware Calculation and Deduction --- + Ring Equalities via {Gr\"obner Bases} in {Isabelle}}, + booktitle = {Towards Mechanized Mathematical Assistants (CALCULEMUS 2007)}, + editor = {Manuel Kauers and Manfred Kerber and Robert Miner and Wolfgang Windsteiger}, + series = LNAI, + volume = 4573, + year = 2007, + publisher = Springer +} + @TechReport{camilleri92, author = {J. Camilleri and T. F. Melham}, title = {Reasoning with Inductively Defined Relations in the @@ -1053,7 +1065,7 @@ @manual{isabelle-intro, author = {Lawrence C. Paulson}, - title = {Introduction to {Isabelle}}, + title = {Old Introduction to {Isabelle}}, institution = CUCL, note = {\url{http://isabelle.in.tum.de/doc/intro.pdf}}} @@ -1065,7 +1077,7 @@ @manual{isabelle-ref, author = {Lawrence C. Paulson}, - title = {The {Isabelle} Reference Manual}, + title = {The Old {Isabelle} Reference Manual}, institution = CUCL, note = {\url{http://isabelle.in.tum.de/doc/ref.pdf}}} @@ -1393,6 +1405,14 @@ year = 2000, publisher = Springer} +@Article{Sutter:2005, + author = {H. Sutter}, + title = {The Free Lunch Is Over --- A Fundamental Turn Toward Concurrency in Software}, + journal = {Dr. Dobb's Journal}, + year = 2005, + volume = 30, + number = 3} + @InCollection{szasz93, author = {Nora Szasz}, title = {A Machine Checked Proof that {Ackermann's} Function is not @@ -1569,6 +1589,24 @@ note = {\url{http://www.in.tum.de/~wenzelm/papers/isar-framework.pdf}} } +@InProceedings{Wenzel-Chaieb:2007b, + author = {Makarius Wenzel and Amine Chaieb}, + title = {{SML} with antiquotations embedded into {Isabelle/Isar}}, + booktitle = {Workshop on Programming Languages for Mechanized Mathematics + (satellite of CALCULEMUS 2007). Hagenberg, Austria}, + editor = {Jacques Carette and Freek Wiedijk}, + month = {June}, + year = {2007} +} + +@InProceedings{Wenzel:2009, + author = {M. Wenzel}, + title = {Parallel Proof Checking in {Isabelle/Isar}}, + booktitle = {ACM SIGSAM Workshop on Programming Languages for Mechanized Mathematics Systems (PLMMS 2009)}, + year = 2009, + editor = {Dos Reis, G. and L. Th\'ery}, + publisher = {ACM Digital Library}} + @book{principia, author = {A. N. Whitehead and B. Russell}, title = {Principia Mathematica}, diff -r 82873a6f2b81 -r 0d579da1902a src/HOL/Decision_Procs/Commutative_Ring.thy --- a/src/HOL/Decision_Procs/Commutative_Ring.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOL/Decision_Procs/Commutative_Ring.thy Mon Oct 25 08:08:08 2010 +0200 @@ -266,7 +266,7 @@ proof cases assume "even l" then have "Suc l div 2 = l div 2" - by (simp add: nat_number even_nat_plus_one_div_two) + by (simp add: eval_nat_numeral even_nat_plus_one_div_two) moreover from Suc have "l < k" by simp with 1 have "\P. Ipol ls (pow l P) = Ipol ls P ^ l" by simp diff -r 82873a6f2b81 -r 0d579da1902a src/HOL/Library/Char_nat.thy --- a/src/HOL/Library/Char_nat.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOL/Library/Char_nat.thy Mon Oct 25 08:08:08 2010 +0200 @@ -171,7 +171,7 @@ proof - fix n m have nat_of_nibble_less_eq_15: "\n. nat_of_nibble n \ 15" - using Suc_leI [OF nat_of_nibble_less_16] by (auto simp add: nat_number) + using Suc_leI [OF nat_of_nibble_less_16] by (auto simp add: eval_nat_numeral) have less_eq_240: "nat_of_nibble n * 16 \ 240" using nat_of_nibble_less_eq_15 by auto have "nat_of_nibble n * 16 + nat_of_nibble m \ 240 + 15" diff -r 82873a6f2b81 -r 0d579da1902a src/HOL/Library/Ramsey.thy --- a/src/HOL/Library/Ramsey.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOL/Library/Ramsey.thy Mon Oct 25 08:08:08 2010 +0200 @@ -218,7 +218,7 @@ "\Y t. Y \ Z & infinite Y & t < s & (\x\Y. \y\Y. x\y --> f{x,y} = t)" proof - have part2: "\X. X \ Z & finite X & card X = 2 --> f X < s" - using part by (fastsimp simp add: nat_number card_Suc_eq) + using part by (fastsimp simp add: eval_nat_numeral card_Suc_eq) obtain Y t where "Y \ Z" "infinite Y" "t < s" "(\X. X \ Y & finite X & card X = 2 --> f X = t)" diff -r 82873a6f2b81 -r 0d579da1902a src/HOL/List.thy --- a/src/HOL/List.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOL/List.thy Mon Oct 25 08:08:08 2010 +0200 @@ -262,9 +262,9 @@ @{lemma "[a,b,c,d][2 := x] = [a,b,x,d]" by simp}\\ @{lemma "sublist [a,b,c,d,e] {0,2,3} = [a,c,d]" by (simp add:sublist_def)}\\ @{lemma "rotate1 [a,b,c,d] = [b,c,d,a]" by (simp add:rotate1_def)}\\ -@{lemma "rotate 3 [a,b,c,d] = [d,a,b,c]" by (simp add:rotate1_def rotate_def nat_number')}\\ -@{lemma "replicate 4 a = [a,a,a,a]" by (simp add:nat_number')}\\ -@{lemma "[2..<5] = [2,3,4]" by (simp add:nat_number')}\\ +@{lemma "rotate 3 [a,b,c,d] = [d,a,b,c]" by (simp add:rotate1_def rotate_def eval_nat_numeral)}\\ +@{lemma "replicate 4 a = [a,a,a,a]" by (simp add:eval_nat_numeral)}\\ +@{lemma "[2..<5] = [2,3,4]" by (simp add:eval_nat_numeral)}\\ @{lemma "listsum [1,2,3::nat] = 6" by (simp add: listsum_def)} \end{tabular}} \caption{Characteristic examples} diff -r 82873a6f2b81 -r 0d579da1902a src/HOL/MicroJava/BV/BVExample.thy --- a/src/HOL/MicroJava/BV/BVExample.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOL/MicroJava/BV/BVExample.thy Mon Oct 25 08:08:08 2010 +0200 @@ -258,7 +258,7 @@ lemma bounded_append [simp]: "check_bounded append_ins [(Suc 0, 2, 8, Xcpt NullPointer)]" apply (simp add: check_bounded_def) - apply (simp add: nat_number append_ins_def) + apply (simp add: eval_nat_numeral append_ins_def) apply (rule allI, rule impI) apply (elim pc_end pc_next pc_0) apply auto @@ -327,7 +327,7 @@ lemma bounded_makelist [simp]: "check_bounded make_list_ins []" apply (simp add: check_bounded_def) - apply (simp add: nat_number make_list_ins_def) + apply (simp add: eval_nat_numeral make_list_ins_def) apply (rule allI, rule impI) apply (elim pc_end pc_next pc_0) apply auto @@ -343,7 +343,7 @@ "wt_method E test_name [] (PrimT Void) 3 2 make_list_ins [] \\<^sub>m" apply (simp add: wt_method_def) apply (simp add: make_list_ins_def phi_makelist_def) - apply (simp add: wt_start_def nat_number) + apply (simp add: wt_start_def eval_nat_numeral) apply (simp add: wt_instr_def) apply clarify apply (elim pc_end pc_next pc_0) diff -r 82873a6f2b81 -r 0d579da1902a src/HOL/Multivariate_Analysis/Determinants.thy --- a/src/HOL/Multivariate_Analysis/Determinants.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOL/Multivariate_Analysis/Determinants.thy Mon Oct 25 08:08:08 2010 +0200 @@ -1042,9 +1042,9 @@ lemma setprod_1: "setprod f {(1::nat)..1} = f 1" by simp lemma setprod_2: "setprod f {(1::nat)..2} = f 1 * f 2" - by (simp add: nat_number setprod_numseg mult_commute) + by (simp add: eval_nat_numeral setprod_numseg mult_commute) lemma setprod_3: "setprod f {(1::nat)..3} = f 1 * f 2 * f 3" - by (simp add: nat_number setprod_numseg mult_commute) + by (simp add: eval_nat_numeral setprod_numseg mult_commute) lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1" by (simp add: det_def sign_id UNIV_1) diff -r 82873a6f2b81 -r 0d579da1902a src/HOL/Nat_Numeral.thy --- a/src/HOL/Nat_Numeral.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOL/Nat_Numeral.thy Mon Oct 25 08:08:08 2010 +0200 @@ -693,11 +693,7 @@ apply (simp only: nat_add_distrib, simp) done -lemmas nat_number = - nat_number_of_Pls nat_number_of_Min - nat_number_of_Bit0 nat_number_of_Bit1 - -lemmas nat_number' = +lemmas eval_nat_numeral = nat_number_of_Bit0 nat_number_of_Bit1 lemmas nat_arith = diff -r 82873a6f2b81 -r 0d579da1902a src/HOL/Tools/Function/function.ML --- a/src/HOL/Tools/Function/function.ML Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOL/Tools/Function/function.ML Mon Oct 25 08:08:08 2010 +0200 @@ -108,8 +108,9 @@ lthy |> addsmps (conceal_partial o Binding.qualify false "partial") "psimps" conceal_partial psimp_attribs psimps - ||> fold_option (snd oo addsmps I "simps" I simp_attribs) trsimps - ||> fold_option (Spec_Rules.add Spec_Rules.Equational o pair fs) trsimps + ||> (case trsimps of NONE => I | SOME thms => + addsmps I "simps" I simp_attribs thms #> snd + #> Spec_Rules.add Spec_Rules.Equational (fs, thms)) ||>> Local_Theory.note ((conceal_partial (qualify "pinduct"), [Attrib.internal (K (Rule_Cases.case_names cnames)), Attrib.internal (K (Rule_Cases.consumes 1)), @@ -117,7 +118,8 @@ ||>> Local_Theory.note ((Binding.conceal (qualify "termination"), []), [termination]) ||> (snd o Local_Theory.note ((qualify "cases", [Attrib.internal (K (Rule_Cases.case_names cnames))]), [cases])) - ||> fold_option (snd oo curry Local_Theory.note (qualify "domintros", [])) domintros + ||> (case domintros of NONE => I | SOME thms => + Local_Theory.note ((qualify "domintros", []), thms) #> snd) val info = { add_simps=addsmps, case_names=cnames, psimps=psimps', pinducts=snd pinducts', simps=NONE, inducts=NONE, termination=termination', diff -r 82873a6f2b81 -r 0d579da1902a src/HOL/Tools/Function/function_lib.ML --- a/src/HOL/Tools/Function/function_lib.ML Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOL/Tools/Function/function_lib.ML Mon Oct 25 08:08:08 2010 +0200 @@ -8,29 +8,15 @@ structure Function_Lib = (* FIXME proper signature *) struct -fun map_option f NONE = NONE - | map_option f (SOME x) = SOME (f x); - -fun fold_option f NONE y = y - | fold_option f (SOME x) y = f x y; - -(* Ex: "The variable" ^ plural " is" "s are" vs *) +(* "The variable" ^ plural " is" "s are" vs *) fun plural sg pl [x] = sg | plural sg pl _ = pl -fun dest_all (Const ("all", _) $ Abs (a as (_,T,_))) = - let - val (n, body) = Term.dest_abs a - in - (Free (n, T), body) - end - | dest_all _ = raise Match - (* Removes all quantifiers from a term, replacing bound variables by frees. *) fun dest_all_all (t as (Const ("all",_) $ _)) = let - val (v,b) = dest_all t + val (v,b) = Logic.dest_all t |> apfst Free val (vs, b') = dest_all_all b in (v :: vs, b') @@ -56,18 +42,10 @@ (ctx, [], t) -fun map3 _ [] [] [] = [] - | map3 f (x :: xs) (y :: ys) (z :: zs) = f x y z :: map3 f xs ys zs - | map3 _ _ _ _ = raise Library.UnequalLengths; - fun map4 _ [] [] [] [] = [] | map4 f (x :: xs) (y :: ys) (z :: zs) (u :: us) = f x y z u :: map4 f xs ys zs us | map4 _ _ _ _ _ = raise Library.UnequalLengths; -fun map6 _ [] [] [] [] [] [] = [] - | map6 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) = f x y z u v w :: map6 f xs ys zs us vs ws - | map6 _ _ _ _ _ _ _ = raise Library.UnequalLengths; - fun map7 _ [] [] [] [] [] [] [] = [] | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) (b :: bs) = f x y z u v w b :: map7 f xs ys zs us vs ws bs | map7 _ _ _ _ _ _ _ _ = raise Library.UnequalLengths; diff -r 82873a6f2b81 -r 0d579da1902a src/HOL/Tools/Function/mutual.ML --- a/src/HOL/Tools/Function/mutual.ML Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOL/Tools/Function/mutual.ML Mon Oct 25 08:08:08 2010 +0200 @@ -267,10 +267,10 @@ val rew_ss = HOL_basic_ss addsimps all_f_defs val mpsimps = map2 mk_mpsimp fqgars psimps - val mtrsimps = map_option (map2 mk_mpsimp fqgars) trsimps + val mtrsimps = Option.map (map2 mk_mpsimp fqgars) trsimps val minducts = mutual_induct_rules lthy simple_pinduct all_f_defs m val mtermination = full_simplify rew_ss termination - val mdomintros = map_option (map (full_simplify rew_ss)) domintros + val mdomintros = Option.map (map (full_simplify rew_ss)) domintros in FunctionResult { fs=fs, G=G, R=R, psimps=mpsimps, simple_pinducts=minducts, diff -r 82873a6f2b81 -r 0d579da1902a src/HOL/Tools/semiring_normalizer.ML --- a/src/HOL/Tools/semiring_normalizer.ML Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOL/Tools/semiring_normalizer.ML Mon Oct 25 08:08:08 2010 +0200 @@ -826,7 +826,7 @@ end; val nat_exp_ss = - HOL_basic_ss addsimps (@{thms nat_number} @ @{thms nat_arith} @ @{thms arith_simps} @ @{thms rel_simps}) + HOL_basic_ss addsimps (@{thms eval_nat_numeral} @ @{thms nat_arith} @ @{thms arith_simps} @ @{thms rel_simps}) addsimps [@{thm Let_def}, @{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc}]; fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS; diff -r 82873a6f2b81 -r 0d579da1902a src/HOL/ex/HarmonicSeries.thy --- a/src/HOL/ex/HarmonicSeries.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOL/ex/HarmonicSeries.thy Mon Oct 25 08:08:08 2010 +0200 @@ -153,7 +153,7 @@ (auto simp: atLeastSucAtMost_greaterThanAtMost) also have "\ = ((\n\{2..2::nat}. 1/real n) + 1/(real (1::nat)))" - by (simp add: nat_number) + by (simp add: eval_nat_numeral) also have "\ = 1/(real (2::nat)) + 1/(real (1::nat))" by simp finally have diff -r 82873a6f2b81 -r 0d579da1902a src/HOL/ex/NatSum.thy --- a/src/HOL/ex/NatSum.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOL/ex/NatSum.thy Mon Oct 25 08:08:08 2010 +0200 @@ -70,9 +70,9 @@ have "(\i = 0..Suc n. i)^2 = (\i = 0..n. i^3) + (2*(\i = 0..n. i)*(n+1) + (n+1)^2)" (is "_ = ?A + ?B") - using Suc by(simp add:nat_number) + using Suc by(simp add:eval_nat_numeral) also have "?B = (n+1)^3" - using sum_of_naturals by(simp add:nat_number) + using sum_of_naturals by(simp add:eval_nat_numeral) also have "?A + (n+1)^3 = (\i=0..Suc n. i^3)" by simp finally show ?case . qed diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/Bifinite.thy --- a/src/HOLCF/Bifinite.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/Bifinite.thy Mon Oct 25 08:08:08 2010 +0200 @@ -5,7 +5,7 @@ header {* Bifinite domains *} theory Bifinite -imports Algebraic Cprod Sprod Ssum Up Lift One Tr +imports Algebraic Cprod Sprod Ssum Up Lift One Tr Countable begin subsection {* Class of bifinite domains *} diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/Cfun.thy --- a/src/HOLCF/Cfun.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/Cfun.thy Mon Oct 25 08:08:08 2010 +0200 @@ -95,12 +95,6 @@ lemma UU_CFun: "\ \ CFun" by (simp add: CFun_def inst_fun_pcpo) -instance cfun :: (finite_po, finite_po) finite_po -by (rule typedef_finite_po [OF type_definition_CFun]) - -instance cfun :: (finite_po, chfin) chfin -by (rule typedef_chfin [OF type_definition_CFun below_CFun_def]) - instance cfun :: (cpo, discrete_cpo) discrete_cpo by intro_classes (simp add: below_CFun_def Rep_CFun_inject) @@ -550,7 +544,10 @@ lemma strict2 [simp]: "x \ \ \ strict\x = ID" by (simp add: strict_conv_if) - definition +lemma strict3 [simp]: "strict\x\\ = \" +by (simp add: strict_conv_if) + +definition strictify :: "('a \ 'b) \ 'a \ 'b" where "strictify = (\ f x. strict\x\(f\x))" diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/Discrete.thy --- a/src/HOLCF/Discrete.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/Discrete.thy Mon Oct 25 08:08:08 2010 +0200 @@ -44,22 +44,6 @@ "!!S::nat=>('a::type)discr. chain(S) ==> range(S) = {S 0}" by (fast elim: discr_chain0) -instance discr :: (finite) finite_po -proof - have "finite (Discr ` (UNIV :: 'a set))" - by (rule finite_imageI [OF finite]) - also have "(Discr ` (UNIV :: 'a set)) = UNIV" - by (auto, case_tac x, auto) - finally show "finite (UNIV :: 'a discr set)" . -qed - -instance discr :: (type) chfin -apply intro_classes -apply (rule_tac x=0 in exI) -apply (unfold max_in_chain_def) -apply (clarify, erule discr_chain0 [symmetric]) -done - subsection {* \emph{undiscr} *} definition diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/FOCUS/Fstream.thy --- a/src/HOLCF/FOCUS/Fstream.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/FOCUS/Fstream.thy Mon Oct 25 08:08:08 2010 +0200 @@ -42,11 +42,7 @@ lemma Def_maximal: "a = Def d \ a\b \ b = Def d" -apply (rule flat_eq [THEN iffD1, symmetric]) -apply (rule Def_not_UU) -apply (drule antisym_less_inverse) -apply (erule (1) conjunct2 [THEN trans_less]) -done +by simp section "fscons" diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/Fun_Cpo.thy --- a/src/HOLCF/Fun_Cpo.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/Fun_Cpo.thy Mon Oct 25 08:08:08 2010 +0200 @@ -110,26 +110,6 @@ from Yij Yik show "Y j = Y k" by auto qed -instance "fun" :: (finite, chfin) chfin -proof - fix Y :: "nat \ 'a \ 'b" - let ?n = "\x. LEAST n. max_in_chain n (\i. Y i x)" - assume "chain Y" - hence "\x. chain (\i. Y i x)" - by (rule ch2ch_fun) - hence "\x. \n. max_in_chain n (\i. Y i x)" - by (rule chfin) - hence "\x. max_in_chain (?n x) (\i. Y i x)" - by (rule LeastI_ex) - hence "\x. max_in_chain (Max (range ?n)) (\i. Y i x)" - by (rule maxinch_mono [OF _ Max_ge], simp_all) - hence "max_in_chain (Max (range ?n)) Y" - by (rule maxinch2maxinch_lambda) - thus "\n. max_in_chain n Y" .. -qed - -instance "fun" :: (finite, finite_po) finite_po .. - instance "fun" :: (type, discrete_cpo) discrete_cpo proof fix f g :: "'a \ 'b" diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/Library/Sum_Cpo.thy --- a/src/HOLCF/Library/Sum_Cpo.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/Library/Sum_Cpo.thy Mon Oct 25 08:08:08 2010 +0200 @@ -213,8 +213,6 @@ apply (case_tac "\i. Y i", simp_all) done -instance sum :: (finite_po, finite_po) finite_po .. - instance sum :: (discrete_cpo, discrete_cpo) discrete_cpo by intro_classes (simp add: below_sum_def split: sum.split) diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/Lift.thy --- a/src/HOLCF/Lift.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/Lift.thy Mon Oct 25 08:08:08 2010 +0200 @@ -5,17 +5,14 @@ header {* Lifting types of class type to flat pcpo's *} theory Lift -imports Discrete Up Countable +imports Discrete Up begin default_sort type -pcpodef 'a lift = "UNIV :: 'a discr u set" +pcpodef (open) 'a lift = "UNIV :: 'a discr u set" by simp_all -instance lift :: (finite) finite_po -by (rule typedef_finite_po [OF type_definition_lift]) - lemmas inst_lift_pcpo = Abs_lift_strict [symmetric] definition @@ -33,7 +30,7 @@ done rep_datatype "\\'a lift" Def - by (erule lift_induct) (simp_all add: Def_def Abs_lift_inject lift_def inst_lift_pcpo) + by (erule lift_induct) (simp_all add: Def_def Abs_lift_inject inst_lift_pcpo) lemmas lift_distinct1 = lift.distinct(1) lemmas lift_distinct2 = lift.distinct(2) @@ -65,7 +62,7 @@ by simp lemma Def_below_Def: "Def x \ Def y \ x = y" -by (simp add: below_lift_def Def_def Abs_lift_inverse lift_def) +by (simp add: below_lift_def Def_def Abs_lift_inverse) lemma Def_below_iff [simp]: "Def x \ y \ Def x = y" by (induct y, simp, simp add: Def_below_Def) @@ -80,6 +77,18 @@ by (induct x) auto qed +subsection {* Continuity of @{const lift_case} *} + +lemma lift_case_eq: "lift_case \ f x = fup\(\ y. f (undiscr y))\(Rep_lift x)" +apply (induct x, unfold lift.cases) +apply (simp add: Rep_lift_strict) +apply (simp add: Def_def Abs_lift_inverse) +done + +lemma cont2cont_lift_case [simp]: + "\\y. cont (\x. f x y); cont g\ \ cont (\x. lift_case \ (f x) (g x))" +unfolding lift_case_eq by (simp add: cont_Rep_lift [THEN cont_compose]) + subsection {* Further operations *} definition @@ -90,59 +99,14 @@ flift2 :: "('a \ 'b) \ ('a lift \ 'b lift)" where "flift2 f = (FLIFT x. Def (f x))" -subsection {* Continuity Proofs for flift1, flift2 *} - -text {* Need the instance of @{text flat}. *} - -lemma cont_lift_case1: "cont (\f. lift_case a f x)" -apply (induct x) -apply simp -apply simp -apply (rule cont_id [THEN cont2cont_fun]) -done - -lemma cont_lift_case2: "cont (\x. lift_case \ f x)" -apply (rule flatdom_strict2cont) -apply simp -done - -lemma cont_flift1: "cont flift1" -unfolding flift1_def -apply (rule cont2cont_LAM) -apply (rule cont_lift_case2) -apply (rule cont_lift_case1) -done - -lemma FLIFT_mono: - "(\x. f x \ g x) \ (FLIFT x. f x) \ (FLIFT x. g x)" -apply (rule monofunE [where f=flift1]) -apply (rule cont2mono [OF cont_flift1]) -apply (simp add: fun_below_iff) -done - -lemma cont2cont_flift1 [simp, cont2cont]: - "\\y. cont (\x. f x y)\ \ cont (\x. FLIFT y. f x y)" -apply (rule cont_flift1 [THEN cont_compose]) -apply simp -done - -lemma cont2cont_lift_case [simp]: - "\\y. cont (\x. f x y); cont g\ \ cont (\x. lift_case UU (f x) (g x))" -apply (subgoal_tac "cont (\x. (FLIFT y. f x y)\(g x))") -apply (simp add: flift1_def cont_lift_case2) -apply simp -done - -text {* rewrites for @{term flift1}, @{term flift2} *} - lemma flift1_Def [simp]: "flift1 f\(Def x) = (f x)" -by (simp add: flift1_def cont_lift_case2) +by (simp add: flift1_def) lemma flift2_Def [simp]: "flift2 f\(Def x) = Def (f x)" by (simp add: flift2_def) lemma flift1_strict [simp]: "flift1 f\\ = \" -by (simp add: flift1_def cont_lift_case2) +by (simp add: flift1_def) lemma flift2_strict [simp]: "flift2 f\\ = \" by (simp add: flift2_def) @@ -153,4 +117,12 @@ lemma flift2_defined_iff [simp]: "(flift2 f\x = \) = (x = \)" by (cases x, simp_all) +lemma FLIFT_mono: + "(\x. f x \ g x) \ (FLIFT x. f x) \ (FLIFT x. g x)" +by (rule cfun_belowI, case_tac x, simp_all) + +lemma cont2cont_flift1 [simp, cont2cont]: + "\\y. cont (\x. f x y)\ \ cont (\x. FLIFT y. f x y)" +by (simp add: flift1_def cont2cont_LAM) + end diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/Pcpo.thy --- a/src/HOLCF/Pcpo.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/Pcpo.thy Mon Oct 25 08:08:08 2010 +0200 @@ -262,18 +262,6 @@ end -class finite_po = finite + po -begin - -subclass chfin -apply default -apply (drule finite_range_imp_finch) -apply (rule finite) -apply (simp add: finite_chain_def) -done - -end - class flat = pcpo + assumes ax_flat: "x \ y \ x = \ \ x = y" begin @@ -299,7 +287,7 @@ end -text {* Discrete cpos *} +subsection {* Discrete cpos *} class discrete_cpo = below + assumes discrete_cpo [simp]: "x \ y \ x = y" @@ -320,14 +308,14 @@ thus "S i = S 0" by (rule sym) qed -subclass cpo +subclass chfin proof fix S :: "nat \ 'a" assume S: "chain S" - hence "\x. S = (\i. x)" - by (rule discrete_chain_const) - thus "\x. range S <<| x" - by (fast intro: lub_const) + hence "\x. S = (\i. x)" by (rule discrete_chain_const) + hence "max_in_chain 0 S" + unfolding max_in_chain_def by auto + thus "\i. max_in_chain i S" .. qed end diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/Pcpodef.thy --- a/src/HOLCF/Pcpodef.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/Pcpodef.thy Mon Oct 25 08:08:08 2010 +0200 @@ -46,15 +46,6 @@ by (simp only: type_definition.Abs_image [OF type]) qed -theorem typedef_finite_po: - fixes Abs :: "'a::finite_po \ 'b::po" - assumes type: "type_definition Rep Abs A" - shows "OFCLASS('b, finite_po_class)" - apply (intro_classes) - apply (rule typedef_finite_UNIV [OF type]) - apply (rule finite) -done - subsection {* Proving a subtype is chain-finite *} lemma ch2ch_Rep: diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/Product_Cpo.thy --- a/src/HOLCF/Product_Cpo.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/Product_Cpo.thy Mon Oct 25 08:08:08 2010 +0200 @@ -12,19 +12,16 @@ subsection {* Unit type is a pcpo *} -instantiation unit :: below +instantiation unit :: discrete_cpo begin definition below_unit_def [simp]: "x \ (y::unit) \ True" -instance .. -end +instance proof +qed simp -instance unit :: discrete_cpo -by intro_classes simp - -instance unit :: finite_po .. +end instance unit :: pcpo by intro_classes simp @@ -157,8 +154,6 @@ thus "\x. range S <<| x" .. qed -instance prod :: (finite_po, finite_po) finite_po .. - instance prod :: (discrete_cpo, discrete_cpo) discrete_cpo proof fix x y :: "'a \ 'b" diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/Sprod.thy --- a/src/HOLCF/Sprod.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/Sprod.thy Mon Oct 25 08:08:08 2010 +0200 @@ -12,37 +12,31 @@ subsection {* Definition of strict product type *} -pcpodef (Sprod) ('a, 'b) sprod (infixr "**" 20) = +pcpodef ('a, 'b) sprod (infixr "**" 20) = "{p::'a \ 'b. p = \ \ (fst p \ \ \ snd p \ \)}" by simp_all -instance sprod :: ("{finite_po,pcpo}", "{finite_po,pcpo}") finite_po -by (rule typedef_finite_po [OF type_definition_Sprod]) - instance sprod :: ("{chfin,pcpo}", "{chfin,pcpo}") chfin -by (rule typedef_chfin [OF type_definition_Sprod below_Sprod_def]) +by (rule typedef_chfin [OF type_definition_sprod below_sprod_def]) type_notation (xsymbols) sprod ("(_ \/ _)" [21,20] 20) type_notation (HTML output) sprod ("(_ \/ _)" [21,20] 20) -lemma spair_lemma: "(strict\b\a, strict\a\b) \ Sprod" -by (simp add: Sprod_def strict_conv_if) - subsection {* Definitions of constants *} definition sfst :: "('a ** 'b) \ 'a" where - "sfst = (\ p. fst (Rep_Sprod p))" + "sfst = (\ p. fst (Rep_sprod p))" definition ssnd :: "('a ** 'b) \ 'b" where - "ssnd = (\ p. snd (Rep_Sprod p))" + "ssnd = (\ p. snd (Rep_sprod p))" definition spair :: "'a \ 'b \ ('a ** 'b)" where - "spair = (\ a b. Abs_Sprod (strict\b\a, strict\a\b))" + "spair = (\ a b. Abs_sprod (strict\b\a, strict\a\b))" definition ssplit :: "('a \ 'b \ 'c) \ ('a ** 'b) \ 'c" where @@ -59,28 +53,20 @@ subsection {* Case analysis *} -lemma Rep_Sprod_spair: - "Rep_Sprod (:a, b:) = (strict\b\a, strict\a\b)" -unfolding spair_def -by (simp add: cont_Abs_Sprod Abs_Sprod_inverse spair_lemma) - -lemmas Rep_Sprod_simps = - Rep_Sprod_inject [symmetric] below_Sprod_def - Rep_Sprod_strict Rep_Sprod_spair +lemma spair_sprod: "(strict\b\a, strict\a\b) \ sprod" +by (simp add: sprod_def strict_conv_if) -lemma Exh_Sprod: - "z = \ \ (\a b. z = (:a, b:) \ a \ \ \ b \ \)" -apply (insert Rep_Sprod [of z]) -apply (simp add: Rep_Sprod_simps Pair_fst_snd_eq) -apply (simp add: Sprod_def) -apply (erule disjE, simp) -apply (simp add: strict_conv_if) -apply fast -done +lemma Rep_sprod_spair: "Rep_sprod (:a, b:) = (strict\b\a, strict\a\b)" +by (simp add: spair_def cont_Abs_sprod Abs_sprod_inverse spair_sprod) + +lemmas Rep_sprod_simps = + Rep_sprod_inject [symmetric] below_sprod_def + Pair_fst_snd_eq below_prod_def + Rep_sprod_strict Rep_sprod_spair lemma sprodE [case_names bottom spair, cases type: sprod]: - "\p = \ \ Q; \x y. \p = (:x, y:); x \ \; y \ \\ \ Q\ \ Q" -using Exh_Sprod [of p] by auto + obtains "p = \" | x y where "p = (:x, y:)" and "x \ \" and "y \ \" +using Rep_sprod [of p] by (auto simp add: sprod_def Rep_sprod_simps) lemma sprod_induct [case_names bottom spair, induct type: sprod]: "\P \; \x y. \x \ \; y \ \\ \ P (:x, y:)\ \ P x" @@ -89,22 +75,22 @@ subsection {* Properties of \emph{spair} *} lemma spair_strict1 [simp]: "(:\, y:) = \" -by (simp add: Rep_Sprod_simps strict_conv_if) +by (simp add: Rep_sprod_simps) lemma spair_strict2 [simp]: "(:x, \:) = \" -by (simp add: Rep_Sprod_simps strict_conv_if) +by (simp add: Rep_sprod_simps) lemma spair_strict_iff [simp]: "((:x, y:) = \) = (x = \ \ y = \)" -by (simp add: Rep_Sprod_simps strict_conv_if) +by (simp add: Rep_sprod_simps strict_conv_if) lemma spair_below_iff: "((:a, b:) \ (:c, d:)) = (a = \ \ b = \ \ (a \ c \ b \ d))" -by (simp add: Rep_Sprod_simps strict_conv_if) +by (simp add: Rep_sprod_simps strict_conv_if) lemma spair_eq_iff: "((:a, b:) = (:c, d:)) = (a = c \ b = d \ (a = \ \ b = \) \ (c = \ \ d = \))" -by (simp add: Rep_Sprod_simps strict_conv_if) +by (simp add: Rep_sprod_simps strict_conv_if) lemma spair_strict: "x = \ \ y = \ \ (:x, y:) = \" by simp @@ -118,6 +104,10 @@ lemma spair_defined_rev: "(:x, y:) = \ \ x = \ \ y = \" by simp +lemma spair_below: + "\x \ \; y \ \\ \ (:x, y:) \ (:a, b:) = (x \ a \ y \ b)" +by (simp add: spair_below_iff) + lemma spair_eq: "\x \ \; y \ \\ \ ((:x, y:) = (:a, b:)) = (x = a \ y = b)" by (simp add: spair_eq_iff) @@ -135,16 +125,16 @@ subsection {* Properties of \emph{sfst} and \emph{ssnd} *} lemma sfst_strict [simp]: "sfst\\ = \" -by (simp add: sfst_def cont_Rep_Sprod Rep_Sprod_strict) +by (simp add: sfst_def cont_Rep_sprod Rep_sprod_strict) lemma ssnd_strict [simp]: "ssnd\\ = \" -by (simp add: ssnd_def cont_Rep_Sprod Rep_Sprod_strict) +by (simp add: ssnd_def cont_Rep_sprod Rep_sprod_strict) lemma sfst_spair [simp]: "y \ \ \ sfst\(:x, y:) = x" -by (simp add: sfst_def cont_Rep_Sprod Rep_Sprod_spair) +by (simp add: sfst_def cont_Rep_sprod Rep_sprod_spair) lemma ssnd_spair [simp]: "x \ \ \ ssnd\(:x, y:) = y" -by (simp add: ssnd_def cont_Rep_Sprod Rep_Sprod_spair) +by (simp add: ssnd_def cont_Rep_sprod Rep_sprod_spair) lemma sfst_defined_iff [simp]: "(sfst\p = \) = (p = \)" by (cases p, simp_all) @@ -158,24 +148,15 @@ lemma ssnd_defined: "p \ \ \ ssnd\p \ \" by simp -lemma surjective_pairing_Sprod2: "(:sfst\p, ssnd\p:) = p" +lemma spair_sfst_ssnd: "(:sfst\p, ssnd\p:) = p" by (cases p, simp_all) lemma below_sprod: "x \ y = (sfst\x \ sfst\y \ ssnd\x \ ssnd\y)" -apply (simp add: below_Sprod_def sfst_def ssnd_def cont_Rep_Sprod) -apply (simp only: below_prod_def) -done +by (simp add: Rep_sprod_simps sfst_def ssnd_def cont_Rep_sprod) lemma eq_sprod: "(x = y) = (sfst\x = sfst\y \ ssnd\x = ssnd\y)" by (auto simp add: po_eq_conv below_sprod) -lemma spair_below: - "\x \ \; y \ \\ \ (:x, y:) \ (:a, b:) = (x \ a \ y \ b)" -apply (cases "a = \", simp) -apply (cases "b = \", simp) -apply (simp add: below_sprod) -done - lemma sfst_below_iff: "sfst\x \ y = x \ (:y, ssnd\x:)" apply (cases "x = \", simp, cases "y = \", simp) apply (simp add: below_sprod) @@ -195,7 +176,7 @@ by (rule compactI, simp add: ssnd_below_iff) lemma compact_spair: "\compact x; compact y\ \ compact (:x, y:)" -by (rule compact_Sprod, simp add: Rep_Sprod_spair strict_conv_if) +by (rule compact_sprod, simp add: Rep_sprod_spair strict_conv_if) lemma compact_spair_iff: "compact (:x, y:) = (x = \ \ y = \ \ (compact x \ compact y))" diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/Ssum.thy --- a/src/HOLCF/Ssum.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/Ssum.thy Mon Oct 25 08:08:08 2010 +0200 @@ -12,17 +12,14 @@ subsection {* Definition of strict sum type *} -pcpodef (Ssum) ('a, 'b) ssum (infixr "++" 10) = - "{p :: tr \ ('a \ 'b). - (fst p \ TT \ snd (snd p) = \) \ - (fst p \ FF \ fst (snd p) = \)}" +pcpodef ('a, 'b) ssum (infixr "++" 10) = + "{p :: tr \ ('a \ 'b). p = \ \ + (fst p = TT \ fst (snd p) \ \ \ snd (snd p) = \) \ + (fst p = FF \ fst (snd p) = \ \ snd (snd p) \ \) }" by simp_all -instance ssum :: ("{finite_po,pcpo}", "{finite_po,pcpo}") finite_po -by (rule typedef_finite_po [OF type_definition_Ssum]) - instance ssum :: ("{chfin,pcpo}", "{chfin,pcpo}") chfin -by (rule typedef_chfin [OF type_definition_Ssum below_Ssum_def]) +by (rule typedef_chfin [OF type_definition_ssum below_ssum_def]) type_notation (xsymbols) ssum ("(_ \/ _)" [21, 20] 20) @@ -34,45 +31,44 @@ definition sinl :: "'a \ ('a ++ 'b)" where - "sinl = (\ a. Abs_Ssum (strict\a\TT, a, \))" + "sinl = (\ a. Abs_ssum (strict\a\TT, a, \))" definition sinr :: "'b \ ('a ++ 'b)" where - "sinr = (\ b. Abs_Ssum (strict\b\FF, \, b))" + "sinr = (\ b. Abs_ssum (strict\b\FF, \, b))" -lemma sinl_Ssum: "(strict\a\TT, a, \) \ Ssum" -by (simp add: Ssum_def strict_conv_if) +lemma sinl_ssum: "(strict\a\TT, a, \) \ ssum" +by (simp add: ssum_def strict_conv_if) -lemma sinr_Ssum: "(strict\b\FF, \, b) \ Ssum" -by (simp add: Ssum_def strict_conv_if) +lemma sinr_ssum: "(strict\b\FF, \, b) \ ssum" +by (simp add: ssum_def strict_conv_if) -lemma sinl_Abs_Ssum: "sinl\a = Abs_Ssum (strict\a\TT, a, \)" -by (unfold sinl_def, simp add: cont_Abs_Ssum sinl_Ssum) - -lemma sinr_Abs_Ssum: "sinr\b = Abs_Ssum (strict\b\FF, \, b)" -by (unfold sinr_def, simp add: cont_Abs_Ssum sinr_Ssum) +lemma Rep_ssum_sinl: "Rep_ssum (sinl\a) = (strict\a\TT, a, \)" +by (simp add: sinl_def cont_Abs_ssum Abs_ssum_inverse sinl_ssum) -lemma Rep_Ssum_sinl: "Rep_Ssum (sinl\a) = (strict\a\TT, a, \)" -by (simp add: sinl_Abs_Ssum Abs_Ssum_inverse sinl_Ssum) +lemma Rep_ssum_sinr: "Rep_ssum (sinr\b) = (strict\b\FF, \, b)" +by (simp add: sinr_def cont_Abs_ssum Abs_ssum_inverse sinr_ssum) -lemma Rep_Ssum_sinr: "Rep_Ssum (sinr\b) = (strict\b\FF, \, b)" -by (simp add: sinr_Abs_Ssum Abs_Ssum_inverse sinr_Ssum) +lemmas Rep_ssum_simps = + Rep_ssum_inject [symmetric] below_ssum_def + Pair_fst_snd_eq below_prod_def + Rep_ssum_strict Rep_ssum_sinl Rep_ssum_sinr subsection {* Properties of \emph{sinl} and \emph{sinr} *} text {* Ordering *} lemma sinl_below [simp]: "(sinl\x \ sinl\y) = (x \ y)" -by (simp add: below_Ssum_def Rep_Ssum_sinl strict_conv_if) +by (simp add: Rep_ssum_simps strict_conv_if) lemma sinr_below [simp]: "(sinr\x \ sinr\y) = (x \ y)" -by (simp add: below_Ssum_def Rep_Ssum_sinr strict_conv_if) +by (simp add: Rep_ssum_simps strict_conv_if) lemma sinl_below_sinr [simp]: "(sinl\x \ sinr\y) = (x = \)" -by (simp add: below_Ssum_def Rep_Ssum_sinl Rep_Ssum_sinr strict_conv_if) +by (simp add: Rep_ssum_simps strict_conv_if) lemma sinr_below_sinl [simp]: "(sinr\x \ sinl\y) = (x = \)" -by (simp add: below_Ssum_def Rep_Ssum_sinl Rep_Ssum_sinr strict_conv_if) +by (simp add: Rep_ssum_simps strict_conv_if) text {* Equality *} @@ -97,30 +93,30 @@ text {* Strictness *} lemma sinl_strict [simp]: "sinl\\ = \" -by (simp add: sinl_Abs_Ssum Abs_Ssum_strict) +by (simp add: Rep_ssum_simps) lemma sinr_strict [simp]: "sinr\\ = \" -by (simp add: sinr_Abs_Ssum Abs_Ssum_strict) +by (simp add: Rep_ssum_simps) lemma sinl_defined_iff [simp]: "(sinl\x = \) = (x = \)" -by (cut_tac sinl_eq [of "x" "\"], simp) +using sinl_eq [of "x" "\"] by simp lemma sinr_defined_iff [simp]: "(sinr\x = \) = (x = \)" -by (cut_tac sinr_eq [of "x" "\"], simp) +using sinr_eq [of "x" "\"] by simp -lemma sinl_defined [intro!]: "x \ \ \ sinl\x \ \" +lemma sinl_defined: "x \ \ \ sinl\x \ \" by simp -lemma sinr_defined [intro!]: "x \ \ \ sinr\x \ \" +lemma sinr_defined: "x \ \ \ sinr\x \ \" by simp text {* Compactness *} lemma compact_sinl: "compact x \ compact (sinl\x)" -by (rule compact_Ssum, simp add: Rep_Ssum_sinl strict_conv_if) +by (rule compact_ssum, simp add: Rep_ssum_sinl) lemma compact_sinr: "compact x \ compact (sinr\x)" -by (rule compact_Ssum, simp add: Rep_Ssum_sinr strict_conv_if) +by (rule compact_ssum, simp add: Rep_ssum_sinr) lemma compact_sinlD: "compact (sinl\x) \ compact x" unfolding compact_def @@ -138,24 +134,11 @@ subsection {* Case analysis *} -lemma Exh_Ssum: - "z = \ \ (\a. z = sinl\a \ a \ \) \ (\b. z = sinr\b \ b \ \)" -apply (induct z rule: Abs_Ssum_induct) -apply (case_tac y, rename_tac t a b) -apply (case_tac t rule: trE) -apply (rule disjI1) -apply (simp add: Ssum_def Abs_Ssum_strict) -apply (rule disjI2, rule disjI1, rule_tac x=a in exI) -apply (simp add: sinl_Abs_Ssum Ssum_def) -apply (rule disjI2, rule disjI2, rule_tac x=b in exI) -apply (simp add: sinr_Abs_Ssum Ssum_def) -done - lemma ssumE [case_names bottom sinl sinr, cases type: ssum]: - "\p = \ \ Q; - \x. \p = sinl\x; x \ \\ \ Q; - \y. \p = sinr\y; y \ \\ \ Q\ \ Q" -using Exh_Ssum [of p] by auto + obtains "p = \" + | x where "p = sinl\x" and "x \ \" + | y where "p = sinr\y" and "y \ \" +using Rep_ssum [of p] by (auto simp add: ssum_def Rep_ssum_simps) lemma ssum_induct [case_names bottom sinl sinr, induct type: ssum]: "\P \; @@ -177,7 +160,7 @@ definition sscase :: "('a \ 'c) \ ('b \ 'c) \ ('a ++ 'b) \ 'c" where - "sscase = (\ f g s. (\(t, x, y). If t then f\x else g\y fi) (Rep_Ssum s))" + "sscase = (\ f g s. (\(t, x, y). If t then f\x else g\y fi) (Rep_ssum s))" translations "case s of XCONST sinl\x \ t1 | XCONST sinr\y \ t2" == "CONST sscase\(\ x. t1)\(\ y. t2)\s" @@ -187,17 +170,17 @@ "\(XCONST sinr\y). t" == "CONST sscase\\\(\ y. t)" lemma beta_sscase: - "sscase\f\g\s = (\(t, x, y). If t then f\x else g\y fi) (Rep_Ssum s)" -unfolding sscase_def by (simp add: cont_Rep_Ssum [THEN cont_compose]) + "sscase\f\g\s = (\(t, x, y). If t then f\x else g\y fi) (Rep_ssum s)" +unfolding sscase_def by (simp add: cont_Rep_ssum [THEN cont_compose]) lemma sscase1 [simp]: "sscase\f\g\\ = \" -unfolding beta_sscase by (simp add: Rep_Ssum_strict) +unfolding beta_sscase by (simp add: Rep_ssum_strict) lemma sscase2 [simp]: "x \ \ \ sscase\f\g\(sinl\x) = f\x" -unfolding beta_sscase by (simp add: Rep_Ssum_sinl) +unfolding beta_sscase by (simp add: Rep_ssum_sinl) lemma sscase3 [simp]: "y \ \ \ sscase\f\g\(sinr\y) = g\y" -unfolding beta_sscase by (simp add: Rep_Ssum_sinr) +unfolding beta_sscase by (simp add: Rep_ssum_sinr) lemma sscase4 [simp]: "sscase\sinl\sinr\z = z" by (cases z, simp_all) diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/Tools/Domain/domain.ML --- a/src/HOLCF/Tools/Domain/domain.ML Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/Tools/Domain/domain.ML Mon Oct 25 08:08:08 2010 +0200 @@ -8,25 +8,21 @@ signature DOMAIN = sig val add_domain_cmd: - binding -> ((string * string option) list * binding * mixfix * (binding * (bool * binding option * string) list * mixfix) list) list -> theory -> theory val add_domain: - binding -> ((string * string option) list * binding * mixfix * (binding * (bool * binding option * typ) list * mixfix) list) list -> theory -> theory val add_new_domain_cmd: - binding -> ((string * string option) list * binding * mixfix * (binding * (bool * binding option * string) list * mixfix) list) list -> theory -> theory val add_new_domain: - binding -> ((string * string option) list * binding * mixfix * (binding * (bool * binding option * typ) list * mixfix) list) list -> theory -> theory @@ -50,7 +46,6 @@ (prep_typ : theory -> (string * sort) list -> 'a -> typ) (add_isos : (binding * mixfix * (typ * typ)) list -> theory -> info * theory) (arg_sort : bool -> sort) - (comp_dbind : binding) (raw_specs : ((string * string option) list * binding * mixfix * (binding * (bool * binding option * 'a) list * mixfix) list) list) (thy : theory) = @@ -172,7 +167,7 @@ (dbinds ~~ rhss ~~ iso_infos); val (take_rews, thy) = - Domain_Induction.comp_theorems comp_dbind + Domain_Induction.comp_theorems dbinds take_info constr_infos thy; in thy @@ -240,27 +235,21 @@ (Parse.$$$ "=" |-- Parse.enum1 "|" cons_decl); val domains_decl = - Scan.option (Parse.$$$ "(" |-- Parse.binding --| Parse.$$$ ")") -- - Parse.and_list1 domain_decl; + Parse.and_list1 domain_decl; fun mk_domain (definitional : bool) - (opt_name : binding option, - doms : ((((string * string option) list * binding) * mixfix) * + (doms : ((((string * string option) list * binding) * mixfix) * ((binding * (bool * binding option * string) list) * mixfix) list) list ) = let - val names = map (fn (((_, t), _), _) => Binding.name_of t) doms; val specs : ((string * string option) list * binding * mixfix * (binding * (bool * binding option * string) list * mixfix) list) list = map (fn (((vs, t), mx), cons) => (vs, t, mx, map (fn ((c, ds), mx) => (c, ds, mx)) cons)) doms; - val comp_dbind = - case opt_name of NONE => Binding.name (space_implode "_" names) - | SOME s => s; in - if definitional - then add_new_domain_cmd comp_dbind specs - else add_domain_cmd comp_dbind specs + if definitional + then add_new_domain_cmd specs + else add_domain_cmd specs end; val _ = diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/Tools/Domain/domain_induction.ML --- a/src/HOLCF/Tools/Domain/domain_induction.ML Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/Tools/Domain/domain_induction.ML Mon Oct 25 08:08:08 2010 +0200 @@ -8,7 +8,7 @@ signature DOMAIN_INDUCTION = sig val comp_theorems : - binding -> binding list -> + binding list -> Domain_Take_Proofs.take_induct_info -> Domain_Constructors.constr_info list -> theory -> thm list * theory @@ -367,13 +367,14 @@ (******************************************************************************) fun comp_theorems - (comp_dbind : binding) (dbinds : binding list) (take_info : Domain_Take_Proofs.take_induct_info) (constr_infos : Domain_Constructors.constr_info list) (thy : theory) = let -val comp_dname = Binding.name_of comp_dbind; + +val comp_dname = space_implode "_" (map Binding.name_of dbinds); +val comp_dbind = Binding.name comp_dname; (* Test for emptiness *) (* FIXME: reimplement emptiness test diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/Tools/fixrec.ML --- a/src/HOLCF/Tools/fixrec.ML Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/Tools/fixrec.ML Mon Oct 25 08:08:08 2010 +0200 @@ -393,10 +393,16 @@ (******************************** Parsers ********************************) (*************************************************************************) +val opt_thm_name' : (bool * Attrib.binding) parser = + Parse.$$$ "(" -- Parse.$$$ "unchecked" -- Parse.$$$ ")" >> K (true, Attrib.empty_binding) + || Parse_Spec.opt_thm_name ":" >> pair false; + +val spec' : (bool * (Attrib.binding * string)) parser = + opt_thm_name' -- Parse.prop >> (fn ((a, b), c) => (a, (b, c))); + val alt_specs' : (bool * (Attrib.binding * string)) list parser = - Parse.enum1 "|" - (Parse.opt_keyword "unchecked" -- Parse_Spec.spec --| - Scan.option (Scan.ahead (Parse.name || Parse.$$$ "[") -- Parse.!!! (Parse.$$$ "|"))); + let val unexpected = Scan.ahead (Parse.name || Parse.$$$ "[" || Parse.$$$ "("); + in Parse.enum1 "|" (spec' --| Scan.option (unexpected -- Parse.!!! (Parse.$$$ "|"))) end; val _ = Outer_Syntax.local_theory "fixrec" "define recursive functions (HOLCF)" Keyword.thy_decl diff -r 82873a6f2b81 -r 0d579da1902a src/HOLCF/Up.thy --- a/src/HOLCF/Up.thy Fri Oct 22 18:38:59 2010 +0200 +++ b/src/HOLCF/Up.thy Mon Oct 25 08:08:08 2010 +0200 @@ -62,108 +62,67 @@ by (auto split: u.split_asm intro: below_trans) qed -lemma u_UNIV: "UNIV = insert Ibottom (range Iup)" -by (auto, case_tac x, auto) - -instance u :: (finite_po) finite_po -by (intro_classes, simp add: u_UNIV) - - subsection {* Lifted cpo is a cpo *} lemma is_lub_Iup: "range S <<| x \ range (\i. Iup (S i)) <<| Iup x" -apply (rule is_lubI) -apply (rule ub_rangeI) -apply (subst Iup_below) -apply (erule is_ub_lub) -apply (case_tac u) -apply (drule ub_rangeD) -apply simp -apply simp -apply (erule is_lub_lub) -apply (rule ub_rangeI) -apply (drule_tac i=i in ub_rangeD) -apply simp -done - -text {* Now some lemmas about chains of @{typ "'a u"} elements *} - -lemma up_lemma1: "z \ Ibottom \ Iup (THE a. Iup a = z) = z" -by (case_tac z, simp_all) - -lemma up_lemma2: - "\chain Y; Y j \ Ibottom\ \ Y (i + j) \ Ibottom" -apply (erule contrapos_nn) -apply (drule_tac i="j" and j="i + j" in chain_mono) -apply (rule le_add2) -apply (case_tac "Y j") -apply assumption -apply simp -done - -lemma up_lemma3: - "\chain Y; Y j \ Ibottom\ \ Iup (THE a. Iup a = Y (i + j)) = Y (i + j)" -by (rule up_lemma1 [OF up_lemma2]) - -lemma up_lemma4: - "\chain Y; Y j \ Ibottom\ \ chain (\i. THE a. Iup a = Y (i + j))" -apply (rule chainI) -apply (rule Iup_below [THEN iffD1]) -apply (subst up_lemma3, assumption+)+ -apply (simp add: chainE) -done - -lemma up_lemma5: - "\chain Y; Y j \ Ibottom\ \ - (\i. Y (i + j)) = (\i. Iup (THE a. Iup a = Y (i + j)))" -by (rule ext, rule up_lemma3 [symmetric]) - -lemma up_lemma6: - "\chain Y; Y j \ Ibottom\ - \ range Y <<| Iup (\i. THE a. Iup a = Y(i + j))" -apply (rule_tac j1 = j in is_lub_range_shift [THEN iffD1]) -apply assumption -apply (subst up_lemma5, assumption+) -apply (rule is_lub_Iup) -apply (rule cpo_lubI) -apply (erule (1) up_lemma4) -done +unfolding is_lub_def is_ub_def ball_simps +by (auto simp add: below_up_def split: u.split) lemma up_chain_lemma: - "chain Y \ - (\A. chain A \ (\i. Y i) = Iup (\i. A i) \ - (\j. \i. Y (i + j) = Iup (A i))) \ (Y = (\i. Ibottom))" -apply (rule disjCI) -apply (simp add: fun_eq_iff) -apply (erule exE, rename_tac j) -apply (rule_tac x="\i. THE a. Iup a = Y (i + j)" in exI) -apply (simp add: up_lemma4) -apply (simp add: up_lemma6 [THEN thelubI]) -apply (rule_tac x=j in exI) -apply (simp add: up_lemma3) -done - -lemma cpo_up: "chain (Y::nat \ 'a u) \ \x. range Y <<| x" -apply (frule up_chain_lemma, safe) -apply (rule_tac x="Iup (\i. A i)" in exI) -apply (erule_tac j="j" in is_lub_range_shift [THEN iffD1, standard]) -apply (simp add: is_lub_Iup cpo_lubI) -apply (rule exI, rule lub_const) -done + assumes Y: "chain Y" obtains "\i. Y i = Ibottom" + | A k where "\i. Iup (A i) = Y (i + k)" and "chain A" and "range Y <<| Iup (\i. A i)" +proof (cases "\k. Y k \ Ibottom") + case True + then obtain k where k: "Y k \ Ibottom" .. + def A \ "\i. THE a. Iup a = Y (i + k)" + have Iup_A: "\i. Iup (A i) = Y (i + k)" + proof + fix i :: nat + from Y le_add2 have "Y k \ Y (i + k)" by (rule chain_mono) + with k have "Y (i + k) \ Ibottom" by (cases "Y k", auto) + thus "Iup (A i) = Y (i + k)" + by (cases "Y (i + k)", simp_all add: A_def) + qed + from Y have chain_A: "chain A" + unfolding chain_def Iup_below [symmetric] + by (simp add: Iup_A) + hence "range A <<| (\i. A i)" + by (rule cpo_lubI) + hence "range (\i. Iup (A i)) <<| Iup (\i. A i)" + by (rule is_lub_Iup) + hence "range (\i. Y (i + k)) <<| Iup (\i. A i)" + by (simp only: Iup_A) + hence "range (\i. Y i) <<| Iup (\i. A i)" + by (simp only: is_lub_range_shift [OF Y]) + with Iup_A chain_A show ?thesis .. +next + case False + then have "\i. Y i = Ibottom" by simp + then show ?thesis .. +qed instance u :: (cpo) cpo -by intro_classes (rule cpo_up) +proof + fix S :: "nat \ 'a u" + assume S: "chain S" + thus "\x. range (\i. S i) <<| x" + proof (rule up_chain_lemma) + assume "\i. S i = Ibottom" + hence "range (\i. S i) <<| Ibottom" + by (simp add: lub_const) + thus ?thesis .. + next + fix A :: "nat \ 'a" + assume "range S <<| Iup (\i. A i)" + thus ?thesis .. + qed +qed subsection {* Lifted cpo is pointed *} -lemma least_up: "\x::'a u. \y. x \ y" -apply (rule_tac x = "Ibottom" in exI) -apply (rule minimal_up [THEN allI]) -done - instance u :: (cpo) pcpo -by intro_classes (rule least_up) +by intro_classes fast text {* for compatibility with old HOLCF-Version *} lemma inst_up_pcpo: "\ = Ibottom" @@ -192,13 +151,22 @@ done lemma cont_Ifup2: "cont (\x. Ifup f x)" -apply (rule contI) -apply (frule up_chain_lemma, safe) -apply (rule_tac j="j" in is_lub_range_shift [THEN iffD1, standard]) -apply (erule monofun_Ifup2 [THEN ch2ch_monofun]) -apply (simp add: cont_cfun_arg) -apply (simp add: lub_const) -done +proof (rule contI2) + fix Y assume Y: "chain Y" and Y': "chain (\i. Ifup f (Y i))" + from Y show "Ifup f (\i. Y i) \ (\i. Ifup f (Y i))" + proof (rule up_chain_lemma) + fix A and k + assume A: "\i. Iup (A i) = Y (i + k)" + assume "chain A" and "range Y <<| Iup (\i. A i)" + hence "Ifup f (\i. Y i) = (\i. Ifup f (Iup (A i)))" + by (simp add: thelubI contlub_cfun_arg) + also have "\ = (\i. Ifup f (Y (i + k)))" + by (simp add: A) + also have "\ = (\i. Ifup f (Y i))" + using Y' by (rule lub_range_shift) + finally show ?thesis by simp + qed simp +qed (rule monofun_Ifup2) subsection {* Continuous versions of constants *} @@ -251,18 +219,20 @@ text {* lifting preserves chain-finiteness *} lemma up_chain_cases: - "chain Y \ - (\A. chain A \ (\i. Y i) = up\(\i. A i) \ - (\j. \i. Y (i + j) = up\(A i))) \ Y = (\i. \)" -by (simp add: inst_up_pcpo up_def cont_Iup up_chain_lemma) + assumes Y: "chain Y" obtains "\i. Y i = \" + | A k where "\i. up\(A i) = Y (i + k)" and "chain A" and "(\i. Y i) = up\(\i. A i)" +apply (rule up_chain_lemma [OF Y]) +apply (simp_all add: inst_up_pcpo up_def cont_Iup thelubI) +done lemma compact_up: "compact x \ compact (up\x)" apply (rule compactI2) -apply (drule up_chain_cases, safe) +apply (erule up_chain_cases) +apply simp apply (drule (1) compactD2, simp) -apply (erule exE, rule_tac x="i + j" in exI) -apply simp -apply simp +apply (erule exE) +apply (drule_tac f="up" and x="x" in monofun_cfun_arg) +apply (simp, erule exI) done lemma compact_upD: "compact (up\x) \ compact x"