# HG changeset patch # User wenzelm # Date 1286654643 -3600 # Node ID 1080dee73a5332aa9fbb008d10a78545fdee8b27 # Parent 350857040d09c11c7607e77831209caf8cb7bded various concrete ML antiquotations; diff -r 350857040d09 -r 1080dee73a53 doc-src/IsarImplementation/Thy/Logic.thy --- a/doc-src/IsarImplementation/Thy/Logic.thy Sat Oct 09 19:49:19 2010 +0100 +++ b/doc-src/IsarImplementation/Thy/Logic.thy Sat Oct 09 21:04:03 2010 +0100 @@ -185,6 +185,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} *} @@ -381,6 +426,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} *} @@ -633,6 +721,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 {* diff -r 350857040d09 -r 1080dee73a53 doc-src/IsarImplementation/Thy/Prelim.thy --- a/doc-src/IsarImplementation/Thy/Prelim.thy Sat Oct 09 19:49:19 2010 +0100 +++ b/doc-src/IsarImplementation/Thy/Prelim.thy Sat Oct 09 21:04:03 2010 +0100 @@ -211,6 +211,33 @@ \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} *} @@ -270,6 +297,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} *} @@ -916,7 +960,9 @@ \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 @@ -990,4 +1036,24 @@ \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} +*} + end