# HG changeset patch # User wenzelm # Date 1209813968 -7200 # Node ID 35809287ab231e75b44634c715701178c8b106e3 # Parent 378bdbce68e6e2efe068305eee4d0975adf01a29 converted refcard.tex to Thy/Quick_Reference.thy; diff -r 378bdbce68e6 -r 35809287ab23 doc-src/IsarRef/IsaMakefile --- a/doc-src/IsarRef/IsaMakefile Sat May 03 13:25:27 2008 +0200 +++ b/doc-src/IsarRef/IsaMakefile Sat May 03 13:26:08 2008 +0200 @@ -22,7 +22,7 @@ Thy: $(LOG)/HOL-Thy.gz $(LOG)/HOL-Thy.gz: Thy/ROOT.ML ../antiquote_setup.ML Thy/intro.thy \ - Thy/pure.thy Thy/syntax.thy + Thy/pure.thy Thy/syntax.thy Thy/Quick_Reference.thy @$(USEDIR) HOL Thy diff -r 378bdbce68e6 -r 35809287ab23 doc-src/IsarRef/Makefile --- a/doc-src/IsarRef/Makefile Sat May 03 13:25:27 2008 +0200 +++ b/doc-src/IsarRef/Makefile Sat May 03 13:26:08 2008 +0200 @@ -14,7 +14,8 @@ NAME = isar-ref FILES = isar-ref.tex Thy/document/intro.tex basics.tex Thy/document/syntax.tex \ - Thy/document/pure.tex generic.tex logics.tex refcard.tex conversion.tex \ + Thy/document/pure.tex generic.tex logics.tex Thy/document/Quick_Reference.tex \ + conversion.tex \ ../isar.sty ../rail.sty ../railsetup.sty ../proof.sty \ ../iman.sty ../extra.sty ../ttbox.sty ../manual.bib diff -r 378bdbce68e6 -r 35809287ab23 doc-src/IsarRef/Thy/Quick_Reference.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Thy/Quick_Reference.thy Sat May 03 13:26:08 2008 +0200 @@ -0,0 +1,219 @@ +(* $Id$ *) + +theory Quick_Reference +imports Main +begin + +chapter {* Isabelle/Isar quick reference \label{ap:refcard} *} + +section {* Proof commands *} + +subsection {* Primitives and basic syntax *} + +text {* + \begin{tabular}{ll} + @{command "fix"}~@{text x} & augment context by @{text "\x. \"} \\ + @{command "assume"}~@{text "a: \"} & augment context by @{text "\ \ \"} \\ + @{command "then"} & indicate forward chaining of facts \\ + @{command "have"}~@{text "a: \"} & prove local result \\ + @{command "show"}~@{text "a: \"} & prove local result, refining some goal \\ + @{command "using"}~@{text a} & indicate use of additional facts \\ + @{command "unfolding"}~@{text a} & unfold definitional equations \\ + @{command "proof"}~@{text "m\<^sub>1"}~\dots~@{command "qed"}~@{text "m\<^sub>2"} & indicate proof structure and refinements \\ + @{command "{"}~\dots~@{command "}"} & indicate explicit blocks \\ + @{command "next"} & switch blocks \\ + @{command "note"}~@{text "a = b"} & reconsider facts \\ + @{command "let"}~@{text "p = t"} & abbreviate terms by higher-order matching \\ + \end{tabular} + + \begin{matharray}{rcl} + @{text "theory\stmt"} & = & @{command "theorem"}~@{text "name: prop proof"} \Or @{command "definition"}~\dots \Or \dots \\[1ex] + @{text "proof"} & = & @{text "prfx\<^sup>*"}~@{command "proof"}~@{text "method stmt\<^sup>*"}~@{command "qed"}~@{text method} \\ + & \Or & @{text "prfx\<^sup>*"}~@{command "done"} \\[1ex] + @{text prfx} & = & @{command "apply"}~@{text method} \\ + & \Or & @{command "using"}~@{text "name\<^sup>+"} \\ + & \Or & @{command "unfolding"}~@{text "name\<^sup>+"} \\ + @{text stmt} & = & @{command "{"}~@{text "stmt\<^sup>*"}~@{command "}"} \\ + & \Or & @{command "next"} \\ + & \Or & @{command "note"}~@{text "name = name\<^sup>+"} \\ + & \Or & @{command "let"}~@{text "term = term"} \\ + & \Or & @{command "fix"}~@{text "var\<^sup>+"} \\ + & \Or & @{command "assume"}~@{text "name: prop\<^sup>+"} \\ + & \Or & @{command "then"}@{text "\<^sup>?"}~@{text goal} \\ + @{text goal} & = & @{command "have"}~@{text "name: prop\<^sup>+ proof"} \\ + & \Or & @{command "show"}~@{text "name: prop\<^sup>+ proof"} \\ + \end{matharray} +*} + + +subsection {* Abbreviations and synonyms *} + +text {* + \begin{matharray}{rcl} + @{command "by"}~@{text "m\<^sub>1 m\<^sub>2"} & \equiv & @{command "proof"}~@{text "m\<^sub>1"}~@{command "qed"}~@{text "m\<^sub>2"} \\ + @{command ".."} & \equiv & @{command "by"}~@{text rule} \\ + @{command "."} & \equiv & @{command "by"}~@{text this} \\ + @{command "hence"} & \equiv & @{command "then"}~@{command "have"} \\ + @{command "thus"} & \equiv & @{command "then"}~@{command "show"} \\ + @{command "from"}~@{text a} & \equiv & @{command "note"}~@{text a}~@{command "then"} \\ + @{command "with"}~@{text a} & \equiv & @{command "from"}~@{text "a \ this"} \\[1ex] + @{command "from"}~@{text this} & \equiv & @{command "then"} \\ + @{command "from"}~@{text this}~@{command "have"} & \equiv & @{command "hence"} \\ + @{command "from"}~@{text this}~@{command "show"} & \equiv & @{command "thus"} \\ + \end{matharray} +*} + + +subsection {* Derived elements *} + +text {* + \begin{matharray}{rcl} + @{command "also"}@{text "\<^sub>0"} & \approx & @{command "note"}~@{text "calculation = this"} \\ + @{command "also"}@{text "\<^sub>n\<^sub>+\<^sub>1"} & \approx & @{command "note"}~@{text "calculation = trans [OF calculation this]"} \\ + @{command "finally"} & \approx & @{command "also"}~@{command "from"}~@{text calculation} \\[0.5ex] + @{command "moreover"} & \approx & @{command "note"}~@{text "calculation = calculation this"} \\ + @{command "ultimately"} & \approx & @{command "moreover"}~@{command "from"}~@{text calculation} \\[0.5ex] + @{command "presume"}~@{text "a: \"} & \approx & @{command "assume"}~@{text "a: \"} \\ + @{command "def"}~@{text "a: x \ t"} & \approx & @{command "fix"}~@{text x}~@{command "assume"}~@{text "a: x \ t"} \\ + @{command "obtain"}~@{text "x \ a: \"} & \approx & \dots~@{command "fix"}~@{text x}~@{command "assume"}~@{text "a: \"} \\ + @{command "case"}~@{text c} & \approx & @{command "fix"}~@{text x}~@{command "assume"}~@{text "c: \"} \\ + @{command "sorry"} & \approx & @{command "by"}~@{text cheating} \\ + \end{matharray} +*} + + +subsection {* Diagnostic commands *} + +text {* + \begin{tabular}{ll} + @{command "pr"} & print current state \\ + @{command "thm"}~@{text a} & print fact \\ + @{command "term"}~@{text t} & print term \\ + @{command "prop"}~@{text \} & print meta-level proposition \\ + @{command "typ"}~@{text \} & print meta-level type \\ + \end{tabular} +*} + + +section {* Proof methods *} + +text {* + \begin{tabular}{ll} + \multicolumn{2}{l}{\textbf{Single steps (forward-chaining facts)}} \\[0.5ex] + @{method assumption} & apply some assumption \\ + @{method this} & apply current facts \\ + @{method rule}~@{text a} & apply some rule \\ + @{method rule} & apply standard rule (default for @{command "proof"}) \\ + @{method contradiction} & apply @{text "\"} elimination rule (any order) \\ + @{method cases}~@{text t} & case analysis (provides cases) \\ + @{method induct}~@{text x} & proof by induction (provides cases) \\[2ex] + + \multicolumn{2}{l}{\textbf{Repeated steps (inserting facts)}} \\[0.5ex] + @{method "-"} & no rules \\ + @{method intro}~@{text a} & introduction rules \\ + @{method intro_classes} & class introduction rules \\ + @{method elim}~@{text a} & elimination rules \\ + @{method unfold}~@{text a} & definitional rewrite rules \\[2ex] + + \multicolumn{2}{l}{\textbf{Automated proof tools (inserting facts)}} \\[0.5ex] + @{method iprover} & intuitionistic proof search \\ + @{method blast}, @{method fast} & Classical Reasoner \\ + @{method simp}, @{method simp_all} & Simplifier (+ Splitter) \\ + @{method auto}, @{method force} & Simplifier + Classical Reasoner \\ + @{method arith} & Arithmetic procedures \\ + \end{tabular} +*} + + +section {* Attributes *} + +text {* + \begin{tabular}{ll} + \multicolumn{2}{l}{\textbf{Operations}} \\[0.5ex] + @{attribute OF}~@{text a} & rule resolved with facts (skipping ``@{text _}'') \\ + @{attribute of}~@{text t} & rule instantiated with terms (skipping ``@{text _}'') \\ + @{attribute "where"}~@{text "x = t"} & rule instantiated with terms, by variable name \\ + @{attribute symmetric} & resolution with symmetry rule \\ + @{attribute THEN}~@{text b} & resolution with another rule \\ + @{attribute rule_format} & result put into standard rule format \\ + @{attribute elim_format} & destruct rule turned into elimination rule format \\[1ex] + + \multicolumn{2}{l}{\textbf{Declarations}} \\[0.5ex] + @{attribute simp} & Simplifier rule \\ + @{attribute intro}, @{attribute elim}, @{attribute dest} & Pure or Classical Reasoner rule \\ + @{attribute iff} & Simplifier + Classical Reasoner rule \\ + @{attribute split} & case split rule \\ + @{attribute trans} & transitivity rule \\ + @{attribute sym} & symmetry rule \\ + \end{tabular} +*} + + +section {* Rule declarations and methods *} + +text {* + \begin{tabular}{l|lllll} + & @{method rule} & @{method iprover} & @{method blast} & @{method simp} & @{method auto} \\ + & & & @{method fast} & @{method simp_all} & @{method force} \\ + \hline + @{attribute Pure.elim}@{text "!"} @{attribute Pure.intro}@{text "!"} + & @{text "\"} & @{text "\"} \\ + @{attribute Pure.elim} @{attribute Pure.intro} + & @{text "\"} & @{text "\"} \\ + @{attribute elim}@{text "!"} @{attribute intro}@{text "!"} + & @{text "\"} & & @{text "\"} & & @{text "\"} \\ + @{attribute elim} @{attribute intro} + & @{text "\"} & & @{text "\"} & & @{text "\"} \\ + @{attribute iff} + & @{text "\"} & & @{text "\"} & @{text "\"} & @{text "\"} \\ + @{attribute iff}@{text "?"} + & @{text "\"} \\ + @{attribute elim}@{text "?"} @{attribute intro}@{text "?"} + & @{text "\"} \\ + @{attribute simp} + & & & & @{text "\"} & @{text "\"} \\ + @{attribute cong} + & & & & @{text "\"} & @{text "\"} \\ + @{attribute split} + & & & & @{text "\"} & @{text "\"} \\ + \end{tabular} +*} + + +section {* Emulating tactic scripts *} + +subsection {* Commands *} + +text {* + \begin{tabular}{ll} + @{command "apply"}~@{text m} & apply proof method at initial position \\ + @{command "apply_end"}~@{text m} & apply proof method near terminal position \\ + @{command "done"} & complete proof \\ + @{command "defer"}~@{text n} & move subgoal to end \\ + @{command "prefer"}~@{text n} & move subgoal to beginning \\ + @{command "back"} & backtrack last command \\ + \end{tabular} +*} + + +subsection {* Methods *} + +text {* + \begin{tabular}{ll} + @{method rule_tac}~@{text insts} & resolution (with instantiation) \\ + @{method erule_tac}~@{text insts} & elim-resolution (with instantiation) \\ + @{method drule_tac}~@{text insts} & destruct-resolution (with instantiation) \\ + @{method frule_tac}~@{text insts} & forward-resolution (with instantiation) \\ + @{method cut_tac}~@{text insts} & insert facts (with instantiation) \\ + @{method thin_tac}~@{text \} & delete assumptions \\ + @{method subgoal_tac}~@{text \} & new claims \\ + @{method rename_tac}~@{text x} & rename innermost goal parameters \\ + @{method rotate_tac}~@{text n} & rotate assumptions of goal \\ + @{method tactic}~@{text "text"} & arbitrary ML tactic \\ + @{method case_tac}~@{text t} & exhaustion (datatypes) \\ + @{method induct_tac}~@{text x} & induction (datatypes) \\ + @{method ind_cases}~@{text t} & exhaustion + simplification (inductive predicates) \\ + \end{tabular} +*} + +end diff -r 378bdbce68e6 -r 35809287ab23 doc-src/IsarRef/Thy/ROOT.ML --- a/doc-src/IsarRef/Thy/ROOT.ML Sat May 03 13:25:27 2008 +0200 +++ b/doc-src/IsarRef/Thy/ROOT.ML Sat May 03 13:26:08 2008 +0200 @@ -5,3 +5,4 @@ use_thy "intro"; use_thy "syntax"; use_thy "pure"; +use_thy "Quick_Reference"; diff -r 378bdbce68e6 -r 35809287ab23 doc-src/IsarRef/Thy/document/Quick_Reference.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Thy/document/Quick_Reference.tex Sat May 03 13:26:08 2008 +0200 @@ -0,0 +1,278 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Quick{\isacharunderscore}Reference}% +% +\isadelimtheory +\isanewline +\isanewline +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Quick{\isacharunderscore}Reference\isanewline +\isakeyword{imports}\ Main\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupchapter{Isabelle/Isar quick reference \label{ap:refcard}% +} +\isamarkuptrue% +% +\isamarkupsection{Proof commands% +} +\isamarkuptrue% +% +\isamarkupsubsection{Primitives and basic syntax% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{tabular}{ll} + \mbox{\isa{\isacommand{fix}}}~\isa{x} & augment context by \isa{{\isasymAnd}x{\isachardot}\ {\isasymbox}} \\ + \mbox{\isa{\isacommand{assume}}}~\isa{a{\isacharcolon}\ {\isasymphi}} & augment context by \isa{{\isasymphi}\ {\isasymLongrightarrow}\ {\isasymbox}} \\ + \mbox{\isa{\isacommand{then}}} & indicate forward chaining of facts \\ + \mbox{\isa{\isacommand{have}}}~\isa{a{\isacharcolon}\ {\isasymphi}} & prove local result \\ + \mbox{\isa{\isacommand{show}}}~\isa{a{\isacharcolon}\ {\isasymphi}} & prove local result, refining some goal \\ + \mbox{\isa{\isacommand{using}}}~\isa{a} & indicate use of additional facts \\ + \mbox{\isa{\isacommand{unfolding}}}~\isa{a} & unfold definitional equations \\ + \mbox{\isa{\isacommand{proof}}}~\isa{m\isactrlsub {\isadigit{1}}}~\dots~\mbox{\isa{\isacommand{qed}}}~\isa{m\isactrlsub {\isadigit{2}}} & indicate proof structure and refinements \\ + \mbox{\isa{\isacommand{{\isacharbraceleft}}}}~\dots~\mbox{\isa{\isacommand{{\isacharbraceright}}}} & indicate explicit blocks \\ + \mbox{\isa{\isacommand{next}}} & switch blocks \\ + \mbox{\isa{\isacommand{note}}}~\isa{a\ {\isacharequal}\ b} & reconsider facts \\ + \mbox{\isa{\isacommand{let}}}~\isa{p\ {\isacharequal}\ t} & abbreviate terms by higher-order matching \\ + \end{tabular} + + \begin{matharray}{rcl} + \isa{theory{\isasymdash}stmt} & = & \mbox{\isa{\isacommand{theorem}}}~\isa{name{\isacharcolon}\ prop\ proof} \Or \mbox{\isa{\isacommand{definition}}}~\dots \Or \dots \\[1ex] + \isa{proof} & = & \isa{prfx\isactrlsup {\isacharasterisk}}~\mbox{\isa{\isacommand{proof}}}~\isa{method\ stmt\isactrlsup {\isacharasterisk}}~\mbox{\isa{\isacommand{qed}}}~\isa{method} \\ + & \Or & \isa{prfx\isactrlsup {\isacharasterisk}}~\mbox{\isa{\isacommand{done}}} \\[1ex] + \isa{prfx} & = & \mbox{\isa{\isacommand{apply}}}~\isa{method} \\ + & \Or & \mbox{\isa{\isacommand{using}}}~\isa{name\isactrlsup {\isacharplus}} \\ + & \Or & \mbox{\isa{\isacommand{unfolding}}}~\isa{name\isactrlsup {\isacharplus}} \\ + \isa{stmt} & = & \mbox{\isa{\isacommand{{\isacharbraceleft}}}}~\isa{stmt\isactrlsup {\isacharasterisk}}~\mbox{\isa{\isacommand{{\isacharbraceright}}}} \\ + & \Or & \mbox{\isa{\isacommand{next}}} \\ + & \Or & \mbox{\isa{\isacommand{note}}}~\isa{name\ {\isacharequal}\ name\isactrlsup {\isacharplus}} \\ + & \Or & \mbox{\isa{\isacommand{let}}}~\isa{term\ {\isacharequal}\ term} \\ + & \Or & \mbox{\isa{\isacommand{fix}}}~\isa{var\isactrlsup {\isacharplus}} \\ + & \Or & \mbox{\isa{\isacommand{assume}}}~\isa{name{\isacharcolon}\ prop\isactrlsup {\isacharplus}} \\ + & \Or & \mbox{\isa{\isacommand{then}}}\isa{\isactrlsup {\isacharquery}}~\isa{goal} \\ + \isa{goal} & = & \mbox{\isa{\isacommand{have}}}~\isa{name{\isacharcolon}\ prop\isactrlsup {\isacharplus}\ proof} \\ + & \Or & \mbox{\isa{\isacommand{show}}}~\isa{name{\isacharcolon}\ prop\isactrlsup {\isacharplus}\ proof} \\ + \end{matharray}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Abbreviations and synonyms% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{matharray}{rcl} + \mbox{\isa{\isacommand{by}}}~\isa{m\isactrlsub {\isadigit{1}}\ m\isactrlsub {\isadigit{2}}} & \equiv & \mbox{\isa{\isacommand{proof}}}~\isa{m\isactrlsub {\isadigit{1}}}~\mbox{\isa{\isacommand{qed}}}~\isa{m\isactrlsub {\isadigit{2}}} \\ + \mbox{\isa{\isacommand{{\isachardot}{\isachardot}}}} & \equiv & \mbox{\isa{\isacommand{by}}}~\isa{rule} \\ + \mbox{\isa{\isacommand{{\isachardot}}}} & \equiv & \mbox{\isa{\isacommand{by}}}~\isa{this} \\ + \mbox{\isa{\isacommand{hence}}} & \equiv & \mbox{\isa{\isacommand{then}}}~\mbox{\isa{\isacommand{have}}} \\ + \mbox{\isa{\isacommand{thus}}} & \equiv & \mbox{\isa{\isacommand{then}}}~\mbox{\isa{\isacommand{show}}} \\ + \mbox{\isa{\isacommand{from}}}~\isa{a} & \equiv & \mbox{\isa{\isacommand{note}}}~\isa{a}~\mbox{\isa{\isacommand{then}}} \\ + \mbox{\isa{\isacommand{with}}}~\isa{a} & \equiv & \mbox{\isa{\isacommand{from}}}~\isa{a\ {\isasymAND}\ this} \\[1ex] + \mbox{\isa{\isacommand{from}}}~\isa{this} & \equiv & \mbox{\isa{\isacommand{then}}} \\ + \mbox{\isa{\isacommand{from}}}~\isa{this}~\mbox{\isa{\isacommand{have}}} & \equiv & \mbox{\isa{\isacommand{hence}}} \\ + \mbox{\isa{\isacommand{from}}}~\isa{this}~\mbox{\isa{\isacommand{show}}} & \equiv & \mbox{\isa{\isacommand{thus}}} \\ + \end{matharray}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Derived elements% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{matharray}{rcl} + \mbox{\isa{\isacommand{also}}}\isa{\isactrlsub {\isadigit{0}}} & \approx & \mbox{\isa{\isacommand{note}}}~\isa{calculation\ {\isacharequal}\ this} \\ + \mbox{\isa{\isacommand{also}}}\isa{\isactrlsub n\isactrlsub {\isacharplus}\isactrlsub {\isadigit{1}}} & \approx & \mbox{\isa{\isacommand{note}}}~\isa{calculation\ {\isacharequal}\ trans\ {\isacharbrackleft}OF\ calculation\ this{\isacharbrackright}} \\ + \mbox{\isa{\isacommand{finally}}} & \approx & \mbox{\isa{\isacommand{also}}}~\mbox{\isa{\isacommand{from}}}~\isa{calculation} \\[0.5ex] + \mbox{\isa{\isacommand{moreover}}} & \approx & \mbox{\isa{\isacommand{note}}}~\isa{calculation\ {\isacharequal}\ calculation\ this} \\ + \mbox{\isa{\isacommand{ultimately}}} & \approx & \mbox{\isa{\isacommand{moreover}}}~\mbox{\isa{\isacommand{from}}}~\isa{calculation} \\[0.5ex] + \mbox{\isa{\isacommand{presume}}}~\isa{a{\isacharcolon}\ {\isasymphi}} & \approx & \mbox{\isa{\isacommand{assume}}}~\isa{a{\isacharcolon}\ {\isasymphi}} \\ + \mbox{\isa{\isacommand{def}}}~\isa{a{\isacharcolon}\ x\ {\isasymequiv}\ t} & \approx & \mbox{\isa{\isacommand{fix}}}~\isa{x}~\mbox{\isa{\isacommand{assume}}}~\isa{a{\isacharcolon}\ x\ {\isasymequiv}\ t} \\ + \mbox{\isa{\isacommand{obtain}}}~\isa{x\ {\isasymWHERE}\ a{\isacharcolon}\ {\isasymphi}} & \approx & \dots~\mbox{\isa{\isacommand{fix}}}~\isa{x}~\mbox{\isa{\isacommand{assume}}}~\isa{a{\isacharcolon}\ {\isasymphi}} \\ + \mbox{\isa{\isacommand{case}}}~\isa{c} & \approx & \mbox{\isa{\isacommand{fix}}}~\isa{x}~\mbox{\isa{\isacommand{assume}}}~\isa{c{\isacharcolon}\ {\isasymphi}} \\ + \mbox{\isa{\isacommand{sorry}}} & \approx & \mbox{\isa{\isacommand{by}}}~\isa{cheating} \\ + \end{matharray}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Diagnostic commands% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{tabular}{ll} + \mbox{\isa{\isacommand{pr}}} & print current state \\ + \mbox{\isa{\isacommand{thm}}}~\isa{a} & print fact \\ + \mbox{\isa{\isacommand{term}}}~\isa{t} & print term \\ + \mbox{\isa{\isacommand{prop}}}~\isa{{\isasymphi}} & print meta-level proposition \\ + \mbox{\isa{\isacommand{typ}}}~\isa{{\isasymtau}} & print meta-level type \\ + \end{tabular}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Proof methods% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{tabular}{ll} + \multicolumn{2}{l}{\textbf{Single steps (forward-chaining facts)}} \\[0.5ex] + \mbox{\isa{assumption}} & apply some assumption \\ + \mbox{\isa{this}} & apply current facts \\ + \mbox{\isa{rule}}~\isa{a} & apply some rule \\ + \mbox{\isa{rule}} & apply standard rule (default for \mbox{\isa{\isacommand{proof}}}) \\ + \mbox{\isa{contradiction}} & apply \isa{{\isasymnot}} elimination rule (any order) \\ + \mbox{\isa{cases}}~\isa{t} & case analysis (provides cases) \\ + \mbox{\isa{induct}}~\isa{x} & proof by induction (provides cases) \\[2ex] + + \multicolumn{2}{l}{\textbf{Repeated steps (inserting facts)}} \\[0.5ex] + \mbox{\isa{{\isacharminus}}} & no rules \\ + \mbox{\isa{intro}}~\isa{a} & introduction rules \\ + \mbox{\isa{intro{\isacharunderscore}classes}} & class introduction rules \\ + \mbox{\isa{elim}}~\isa{a} & elimination rules \\ + \mbox{\isa{unfold}}~\isa{a} & definitional rewrite rules \\[2ex] + + \multicolumn{2}{l}{\textbf{Automated proof tools (inserting facts)}} \\[0.5ex] + \mbox{\isa{iprover}} & intuitionistic proof search \\ + \mbox{\isa{blast}}, \mbox{\isa{fast}} & Classical Reasoner \\ + \mbox{\isa{simp}}, \mbox{\isa{simp{\isacharunderscore}all}} & Simplifier (+ Splitter) \\ + \mbox{\isa{auto}}, \mbox{\isa{force}} & Simplifier + Classical Reasoner \\ + \mbox{\isa{arith}} & Arithmetic procedures \\ + \end{tabular}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Attributes% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{tabular}{ll} + \multicolumn{2}{l}{\textbf{Operations}} \\[0.5ex] + \mbox{\isa{OF}}~\isa{a} & rule resolved with facts (skipping ``\isa{{\isacharunderscore}}'') \\ + \mbox{\isa{of}}~\isa{t} & rule instantiated with terms (skipping ``\isa{{\isacharunderscore}}'') \\ + \mbox{\isa{where}}~\isa{x\ {\isacharequal}\ t} & rule instantiated with terms, by variable name \\ + \mbox{\isa{symmetric}} & resolution with symmetry rule \\ + \mbox{\isa{THEN}}~\isa{b} & resolution with another rule \\ + \mbox{\isa{rule{\isacharunderscore}format}} & result put into standard rule format \\ + \mbox{\isa{elim{\isacharunderscore}format}} & destruct rule turned into elimination rule format \\[1ex] + + \multicolumn{2}{l}{\textbf{Declarations}} \\[0.5ex] + \mbox{\isa{simp}} & Simplifier rule \\ + \mbox{\isa{intro}}, \mbox{\isa{elim}}, \mbox{\isa{dest}} & Pure or Classical Reasoner rule \\ + \mbox{\isa{iff}} & Simplifier + Classical Reasoner rule \\ + \mbox{\isa{split}} & case split rule \\ + \mbox{\isa{trans}} & transitivity rule \\ + \mbox{\isa{sym}} & symmetry rule \\ + \end{tabular}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Rule declarations and methods% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{tabular}{l|lllll} + & \mbox{\isa{rule}} & \mbox{\isa{iprover}} & \mbox{\isa{blast}} & \mbox{\isa{simp}} & \mbox{\isa{auto}} \\ + & & & \mbox{\isa{fast}} & \mbox{\isa{simp{\isacharunderscore}all}} & \mbox{\isa{force}} \\ + \hline + \mbox{\isa{Pure{\isachardot}elim}}\isa{{\isacharbang}} \mbox{\isa{Pure{\isachardot}intro}}\isa{{\isacharbang}} + & \isa{{\isasymtimes}} & \isa{{\isasymtimes}} \\ + \mbox{\isa{Pure{\isachardot}elim}} \mbox{\isa{Pure{\isachardot}intro}} + & \isa{{\isasymtimes}} & \isa{{\isasymtimes}} \\ + \mbox{\isa{elim}}\isa{{\isacharbang}} \mbox{\isa{intro}}\isa{{\isacharbang}} + & \isa{{\isasymtimes}} & & \isa{{\isasymtimes}} & & \isa{{\isasymtimes}} \\ + \mbox{\isa{elim}} \mbox{\isa{intro}} + & \isa{{\isasymtimes}} & & \isa{{\isasymtimes}} & & \isa{{\isasymtimes}} \\ + \mbox{\isa{iff}} + & \isa{{\isasymtimes}} & & \isa{{\isasymtimes}} & \isa{{\isasymtimes}} & \isa{{\isasymtimes}} \\ + \mbox{\isa{iff}}\isa{{\isacharquery}} + & \isa{{\isasymtimes}} \\ + \mbox{\isa{elim}}\isa{{\isacharquery}} \mbox{\isa{intro}}\isa{{\isacharquery}} + & \isa{{\isasymtimes}} \\ + \mbox{\isa{simp}} + & & & & \isa{{\isasymtimes}} & \isa{{\isasymtimes}} \\ + \mbox{\isa{cong}} + & & & & \isa{{\isasymtimes}} & \isa{{\isasymtimes}} \\ + \mbox{\isa{split}} + & & & & \isa{{\isasymtimes}} & \isa{{\isasymtimes}} \\ + \end{tabular}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Emulating tactic scripts% +} +\isamarkuptrue% +% +\isamarkupsubsection{Commands% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{tabular}{ll} + \mbox{\isa{\isacommand{apply}}}~\isa{m} & apply proof method at initial position \\ + \mbox{\isa{\isacommand{apply{\isacharunderscore}end}}}~\isa{m} & apply proof method near terminal position \\ + \mbox{\isa{\isacommand{done}}} & complete proof \\ + \mbox{\isa{\isacommand{defer}}}~\isa{n} & move subgoal to end \\ + \mbox{\isa{\isacommand{prefer}}}~\isa{n} & move subgoal to beginning \\ + \mbox{\isa{\isacommand{back}}} & backtrack last command \\ + \end{tabular}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Methods% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{tabular}{ll} + \mbox{\isa{rule{\isacharunderscore}tac}}~\isa{insts} & resolution (with instantiation) \\ + \mbox{\isa{erule{\isacharunderscore}tac}}~\isa{insts} & elim-resolution (with instantiation) \\ + \mbox{\isa{drule{\isacharunderscore}tac}}~\isa{insts} & destruct-resolution (with instantiation) \\ + \mbox{\isa{frule{\isacharunderscore}tac}}~\isa{insts} & forward-resolution (with instantiation) \\ + \mbox{\isa{cut{\isacharunderscore}tac}}~\isa{insts} & insert facts (with instantiation) \\ + \mbox{\isa{thin{\isacharunderscore}tac}}~\isa{{\isasymphi}} & delete assumptions \\ + \mbox{\isa{subgoal{\isacharunderscore}tac}}~\isa{{\isasymphi}} & new claims \\ + \mbox{\isa{rename{\isacharunderscore}tac}}~\isa{x} & rename innermost goal parameters \\ + \mbox{\isa{rotate{\isacharunderscore}tac}}~\isa{n} & rotate assumptions of goal \\ + \mbox{\isa{tactic}}~\isa{text} & arbitrary ML tactic \\ + \mbox{\isa{case{\isacharunderscore}tac}}~\isa{t} & exhaustion (datatypes) \\ + \mbox{\isa{induct{\isacharunderscore}tac}}~\isa{x} & induction (datatypes) \\ + \mbox{\isa{ind{\isacharunderscore}cases}}~\isa{t} & exhaustion + simplification (inductive predicates) \\ + \end{tabular}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 378bdbce68e6 -r 35809287ab23 doc-src/IsarRef/Thy/document/session.tex --- a/doc-src/IsarRef/Thy/document/session.tex Sat May 03 13:25:27 2008 +0200 +++ b/doc-src/IsarRef/Thy/document/session.tex Sat May 03 13:26:08 2008 +0200 @@ -4,6 +4,8 @@ \input{pure.tex} +\input{Quick_Reference.tex} + %%% Local Variables: %%% mode: latex %%% TeX-master: "root" diff -r 378bdbce68e6 -r 35809287ab23 doc-src/IsarRef/isar-ref.tex --- a/doc-src/IsarRef/isar-ref.tex Sat May 03 13:25:27 2008 +0200 +++ b/doc-src/IsarRef/isar-ref.tex Sat May 03 13:26:08 2008 +0200 @@ -76,7 +76,7 @@ \input{logics.tex} \appendix -\input{refcard.tex} +\input{Thy/document/Quick_Reference.tex} \input{conversion.tex} \begingroup diff -r 378bdbce68e6 -r 35809287ab23 doc-src/IsarRef/refcard.tex --- a/doc-src/IsarRef/refcard.tex Sat May 03 13:25:27 2008 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,192 +0,0 @@ - -\chapter{Isabelle/Isar quick reference}\label{ap:refcard} - -\section{Proof commands} - -\subsection{Primitives and basic syntax} - -\begin{tabular}{ll} - $\FIX{\vec x}$ & augment context by $\All {\vec x} \Box$ \\ - $\ASSUME{a}{\vec\phi}$ & augment context by $\vec\phi \Imp \Box$ \\ - $\THEN$ & indicate forward chaining of facts \\ - $\HAVE{a}{\phi}$ & prove local result \\ - $\SHOW{a}{\phi}$ & prove local result, refining some goal \\ - $\USING{\vec a}$ & indicate use of additional facts \\ - $\UNFOLDING{\vec a}$ & unfold definitional equations \\ - $\PROOF{m@1}~\dots~\QED{m@2}$ & indicate proof structure and refinements \\ - $\BG~\dots~\EN$ & declare explicit blocks \\ - $\NEXT$ & switch blocks \\ - $\NOTE{a}{\vec b}$ & reconsider facts \\ - $\LET{p = t}$ & \Text{abbreviate terms by higher-order matching} \\ -\end{tabular} - -\begin{matharray}{rcl} - theory{\dsh}stmt & = & \THEOREM{name}{prop} ~proof \Or \isarkeyword{definition}~\dots \Or \dots \\[1ex] - proof & = & prfx^*~\PROOF{method}~stmt^*~\QED{method} \\ - & \Or & prfx^*~\DONE \\[1ex] - prfx & = & \APPLY{method} \\ - & \Or & \USING{name^+} \\ - & \Or & \UNFOLDING{name^+} \\ - stmt & = & \BG~stmt^*~\EN \\ - & \Or & \NEXT \\ - & \Or & \NOTE{name}{name^+} \\ - & \Or & \LET{term = term} \\ - & \Or & \FIX{var^+} \\ - & \Or & \ASSUME{name}{prop^+}\\ - & \Or & \THEN^?~goal \\ - goal & = & \HAVE{name}{prop}~proof \\ - & \Or & \SHOW{name}{prop}~proof \\ -\end{matharray} - - -\subsection{Abbreviations and synonyms} - -\begin{matharray}{rcl} - \BYY{m@1}{m@2} & \equiv & \PROOF{m@1}~\QED{m@2} \\ - \DDOT & \equiv & \BY{rule} \\ - \DOT & \equiv & \BY{this} \\ - \HENCENAME & \equiv & \THEN~\HAVENAME \\ - \THUSNAME & \equiv & \THEN~\SHOWNAME \\ - \FROM{\vec a} & \equiv & \NOTE{this}{\vec a}~\THEN \\ - \WITH{\vec a} & \equiv & \FROM{\vec a~\AND~this} \\[1ex] - \FROM{this} & \equiv & \THEN \\ - \FROM{this}~\HAVENAME & \equiv & \HENCENAME \\ - \FROM{this}~\SHOWNAME & \equiv & \THUSNAME \\ -\end{matharray} - - -\subsection{Derived elements} - -\begin{matharray}{rcl} - \ALSO@0 & \approx & \NOTE{calculation}{this} \\ - \ALSO@{n+1} & \approx & \NOTE{calculation}{trans~[OF~calculation~this]} \\ - \FINALLY & \approx & \ALSO~\FROM{calculation} \\[0.5ex] - \MOREOVER & \approx & \NOTE{calculation}{calculation~this} \\ - \ULTIMATELY & \approx & \MOREOVER~\FROM{calculation} \\[0.5ex] - \PRESUME{a}{\vec\phi} & \approx & \ASSUME{a}{\vec\phi} \\ -% & & \Text{(permissive assumption)} \\ - \DEF{a}{x \equiv t} & \approx & \FIX{x}~\ASSUME{a}{x \equiv t} \\ -% & & \Text{(definitional assumption)} \\ - \OBTAIN{\vec x}{a}{\vec\phi} & \approx & \dots~\FIX{\vec x}~\ASSUME{a}{\vec\phi} \\ -% & & \Text{(generalized existence)} \\ - \CASE{c} & \approx & \FIX{\vec x}~\ASSUME{c}{\vec\phi} \\ -% & & \Text{(named context)} \\[0.5ex] - \SORRY & \approx & \BY{cheating} \\ -\end{matharray} - - -\subsection{Diagnostic commands} - -\begin{matharray}{ll} - \isarkeyword{pr} & \Text{print current state} \\ - \isarkeyword{thm}~\vec a & \Text{print theorems} \\ - \isarkeyword{term}~t & \Text{print term} \\ - \isarkeyword{prop}~\phi & \Text{print meta-level proposition} \\ - \isarkeyword{typ}~\tau & \Text{print meta-level type} \\ -\end{matharray} - - -\section{Proof methods} - -\begin{tabular}{ll} - \multicolumn{2}{l}{\textbf{Single steps (forward-chaining facts)}} \\[0.5ex] - $assumption$ & apply some assumption \\ - $this$ & apply current facts \\ - $rule~\vec a$ & apply some rule \\ - $rule$ & apply standard rule (default for $\PROOFNAME$) \\ - $contradiction$ & apply $\neg{}$ elimination rule (any order) \\ - $cases~t$ & case analysis (provides cases) \\ - $induct~\vec x$ & proof by induction (provides cases) \\[2ex] - - \multicolumn{2}{l}{\textbf{Repeated steps (inserting facts)}} \\[0.5ex] - $-$ & \Text{no rules} \\ - $intro~\vec a$ & \Text{introduction rules} \\ - $intro_classes$ & \Text{class introduction rules} \\ - $elim~\vec a$ & \Text{elimination rules} \\ - $unfold~\vec a$ & \Text{definitions} \\[2ex] - - \multicolumn{2}{l}{\textbf{Automated proof tools (inserting facts, or even prems!)}} \\[0.5ex] - $rules$ & \Text{intuitionistic proof search} \\ - $blast$, $fast$ & Classical Reasoner \\ - $simp$, $simp_all$ & Simplifier (+ Splitter) \\ - $auto$, $force$ & Simplifier + Classical Reasoner \\ - $arith$ & Arithmetic procedure \\ -\end{tabular} - - -\section{Attributes} - -\begin{tabular}{ll} - \multicolumn{2}{l}{\textbf{Operations}} \\[0.5ex] - $OF~\vec a$ & rule resolved with facts (skipping ``$_$'') \\ - $of~\vec t$ & rule instantiated with terms (skipping ``$_$'') \\ - $where~\vec x = \vec t$ & rule instantiated with terms, by variable name \\ - $symmetric$ & resolution with symmetry rule \\ - $THEN~b$ & resolution with another rule \\ - $rule_format$ & result put into standard rule format \\ - $elim_format$ & destruct rule turned into elimination rule format \\[1ex] - - \multicolumn{2}{l}{\textbf{Declarations}} \\[0.5ex] - $simp$ & Simplifier rule \\ - $intro$, $elim$, $dest$ & Pure or Classical Reasoner rule \\ - $iff$ & Simplifier + Classical Reasoner rule \\ - $split$ & case split rule \\ - $trans$ & transitivity rule \\ - $sym$ & symmetry rule \\ -\end{tabular} - - -\section{Rule declarations and methods} - -\begin{tabular}{l|lllll} - & $rule$ & $rules$ & $blast$ etc. & $simp$ etc. & $auto$ etc. \\ - \hline - $elim!$ $intro!$ (Pure) & $\times$ & $\times$ \\ - $elim$ $intro$ (Pure) & $\times$ & $\times$ \\ - $elim!$ $intro!$ & $\times$ & & $\times$ & & $\times$ \\ - $elim$ $intro$ & $\times$ & & $\times$ & & $\times$ \\ - $iff$ & $\times$ & & $\times$ & $\times$ & $\times$ \\ - $iff?$ & $\times$ \\ - $elim?$ $intro?$ & $\times$ \\ - $simp$ & & & & $\times$ & $\times$ \\ - $cong$ & & & & $\times$ & $\times$ \\ - $split$ & & & & $\times$ & $\times$ \\ -\end{tabular} - - -\section{Emulating tactic scripts} - -\subsection{Commands} - -\begin{tabular}{ll} - $\APPLY{m}$ & apply proof method at initial position \\ - $\isarkeyword{apply_end}~(m)$ & apply proof method near terminal position \\ - $\isarkeyword{done}$ & complete proof \\ - $\isarkeyword{defer}~n$ & move subgoal to end \\ - $\isarkeyword{prefer}~n$ & move subgoal to beginning \\ - $\isarkeyword{back}$ & backtrack last command \\ -\end{tabular} - -\subsection{Methods} - -\begin{tabular}{ll} - $rule_tac~insts$ & resolution (with instantiation) \\ - $erule_tac~insts$ & elim-resolution (with instantiation) \\ - $drule_tac~insts$ & destruct-resolution (with instantiation) \\ - $frule_tac~insts$ & forward-resolution (with instantiation) \\ - $cut_tac~insts$ & insert facts (with instantiation) \\ - $thin_tac~\phi$ & delete assumptions \\ - $subgoal_tac~\phi$ & new claims \\ - $rename_tac~\vec x$ & rename suffix of goal parameters \\ - $rotate_tac~n$ & rotate assumptions of goal \\ - $tactic~text$ & arbitrary ML tactic \\ - $case_tac~t$ & exhaustion (datatypes) \\ - $induct_tac~\vec x$ & induction (datatypes) \\ - $ind_cases~t$ & exhaustion + simplification (inductive sets) \\ -\end{tabular} - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "isar-ref" -%%% End: