diff -r c56cff1c0e73 -r 77e3ffb5aeb3 src/Doc/Isar_Ref/Quick_Reference.thy --- a/src/Doc/Isar_Ref/Quick_Reference.thy Sun Feb 07 19:32:35 2016 +0100 +++ b/src/Doc/Isar_Ref/Quick_Reference.thy Sun Feb 07 19:33:42 2016 +0100 @@ -8,43 +8,58 @@ section \Proof commands\ -subsection \Primitives and basic syntax\ +subsection \Main grammar\ + +text \ + \begin{tabular}{rcl} + \main\ & = & \<^theory_text>\notepad begin "statement\<^sup>*" end\ \\ + & \|\ & \<^theory_text>\theorem name: props "proof"\ \\ + & \|\ & \<^theory_text>\theorem name:\ \\ + & & \quad\<^theory_text>\fixes "var\<^sup>+"\ \\ + & & \quad\<^theory_text>\assumes name: props\ \\ + & & \quad\<^theory_text>\shows name: props "proof"\ \\ + \proof\ & = & \<^theory_text>\"refinement\<^sup>*" proof "method\<^sup>?" "statement\<^sup>*" qed "method\<^sup>?"\ \\ + & \|\ & \<^theory_text>\"refinement\<^sup>*" done\ \\ + \refinement\ & = & \<^theory_text>\apply method\ \\ + & \|\ & \<^theory_text>\supply facts\ \\ + & \|\ & \<^theory_text>\subgoal "proof"\ \\ + & \|\ & \<^theory_text>\subgoal for "var\<^sup>+" "proof"\ \\ + & \|\ & \<^theory_text>\using facts\ \\ + & \|\ & \<^theory_text>\unfolding facts\ \\ + \statement\ & = & \<^theory_text>\{ "statement\<^sup>*" }\ \\ + & \|\ & \<^theory_text>\next\ \\ + & \|\ & \<^theory_text>\note name = facts\ \\ + & \|\ & \<^theory_text>\let "term" = "term"\ \\ + & \|\ & \<^theory_text>\write name (mixfix)\ \\ + & \|\ & \<^theory_text>\fix "var\<^sup>+"\ \\ + & \|\ & \<^theory_text>\assume name: props\ \\ + & \|\ & \<^theory_text>\assume name: props if name: props for "var\<^sup>+"\ \\ + & \|\ & \<^theory_text>\then"\<^sup>?" goal\ \\ + \goal\ & = & \<^theory_text>\have name: props "proof"\ \\ + & \|\ & \<^theory_text>\have name: props if name: props for "var\<^sup>+" "proof"\ \\ + & \|\ & \<^theory_text>\show name: props "proof"\ \\ + & \|\ & \<^theory_text>\show name: props if name: props for "var\<^sup>+" "proof"\ \\ + \end{tabular} +\ + + +subsection \Primitives\ text \ \begin{tabular}{ll} - @{command "fix"}~\x\ & augment context by \\x. \\ \\ - @{command "assume"}~\a: \\ & augment context by \\ \ \\ \\ - @{command "then"} & indicate forward chaining of facts \\ - @{command "have"}~\a: \\ & prove local result \\ - @{command "show"}~\a: \\ & prove local result, refining some goal \\ - @{command "using"}~\a\ & indicate use of additional facts \\ - @{command "unfolding"}~\a\ & unfold definitional equations \\ - @{command "proof"}~\m\<^sub>1\~\dots~@{command "qed"}~\m\<^sub>2\ & indicate proof structure and refinements \\ - @{command "{"}~\\\~@{command "}"} & indicate explicit blocks \\ - @{command "next"} & switch blocks \\ - @{command "note"}~\a = b\ & reconsider facts \\ - @{command "let"}~\p = t\ & abbreviate terms by higher-order matching \\ - @{command "write"}~\c (mx)\ & declare local mixfix syntax \\ - \end{tabular} - - \<^medskip> - - \begin{tabular}{rcl} - \proof\ & = & \prfx\<^sup>*\~@{command "proof"}~\method\<^sup>? stmt\<^sup>*\~@{command "qed"}~\method\<^sup>?\ \\ - & \|\ & \prfx\<^sup>*\~@{command "done"} \\ - \prfx\ & = & @{command "apply"}~\method\ \\ - & \|\ & @{command "using"}~\facts\ \\ - & \|\ & @{command "unfolding"}~\facts\ \\ - \stmt\ & = & @{command "{"}~\stmt\<^sup>*\~@{command "}"} \\ - & \|\ & @{command "next"} \\ - & \|\ & @{command "note"}~\name = facts\ \\ - & \|\ & @{command "let"}~\term = term\ \\ - & \|\ & @{command "write"}~\name (mixfix)\ \\ - & \|\ & @{command "fix"}~\var\<^sup>+\ \\ - & \|\ & @{command "assume"}~\name: props\ \\ - & \|\ & @{command "then"}\\<^sup>?\~\goal\ \\ - \goal\ & = & @{command "have"}~\name: props proof\ \\ - & \|\ & @{command "show"}~\name: props proof\ \\ + \<^theory_text>\fix x\ & augment context by \\x. \\ \\ + \<^theory_text>\assume a: A\ & augment context by \A \ \\ \\ + \<^theory_text>\then\ & indicate forward chaining of facts \\ + \<^theory_text>\have a: A\ & prove local result \\ + \<^theory_text>\show a: A\ & prove local result, refining some goal \\ + \<^theory_text>\using a\ & indicate use of additional facts \\ + \<^theory_text>\unfolding a\ & unfold definitional equations \\ + \<^theory_text>\proof m\<^sub>1 \ qed m\<^sub>2\ & indicate proof structure and refinements \\ + \<^theory_text>\{ \ }\ & indicate explicit blocks \\ + \<^theory_text>\next\ & switch proof blocks \\ + \<^theory_text>\note a = b\ & reconsider and declare facts \\ + \<^theory_text>\let p = t\ & abbreviate terms by higher-order matching \\ + \<^theory_text>\write c (mx)\ & declare local mixfix syntax \\ \end{tabular} \ @@ -53,13 +68,12 @@ text \ \begin{tabular}{rcl} - @{command "by"}~\m\<^sub>1 m\<^sub>2\ & \\\ & - @{command "proof"}~\m\<^sub>1\~@{command "qed"}~\m\<^sub>2\ \\ - @{command ".."} & \\\ & @{command "by"}~\standard\ \\ - @{command "."} & \\\ & @{command "by"}~\this\ \\ - @{command "from"}~\a\ & \\\ & @{command "note"}~\a\~@{command "then"} \\ - @{command "with"}~\a\ & \\\ & @{command "from"}~\a \ this\ \\ - @{command "from"}~\this\ & \\\ & @{command "then"} \\ + \<^theory_text>\by m\<^sub>1 m\<^sub>2\ & \\\ & \<^theory_text>\proof m\<^sub>1 qed m\<^sub>2\ \\ + \<^theory_text>\..\ & \\\ & \<^theory_text>\by standard\ \\ + \<^theory_text>\.\ & \\\ & \<^theory_text>\by this\ \\ + \<^theory_text>\from a\ & \\\ & \<^theory_text>\note a then\ \\ + \<^theory_text>\with a\ & \\\ & \<^theory_text>\from a and this\ \\ + \<^theory_text>\from this\ & \\\ & \<^theory_text>\then\ \\ \end{tabular} \ @@ -68,26 +82,19 @@ text \ \begin{tabular}{rcl} - @{command "also"}\\<^sub>0\ & \\\ & - @{command "note"}~\calculation = this\ \\ - @{command "also"}\\<^sub>n\<^sub>+\<^sub>1\ & \\\ & - @{command "note"}~\calculation = trans [OF calculation this]\ \\ - @{command "finally"} & \\\ & - @{command "also"}~@{command "from"}~\calculation\ \\[0.5ex] - @{command "moreover"} & \\\ & - @{command "note"}~\calculation = calculation this\ \\ - @{command "ultimately"} & \\\ & - @{command "moreover"}~@{command "from"}~\calculation\ \\[0.5ex] - @{command "presume"}~\a: \\ & \\\ & - @{command "assume"}~\a: \\ \\ - @{command "def"}~\a: x \ t\ & \\\ & - @{command "fix"}~\x\~@{command "assume"}~\a: x \ t\ \\ - @{command "obtain"}~\x \ a: \\ & \\\ & - \\\~@{command "fix"}~\x\~@{command "assume"}~\a: \\ \\ - @{command "case"}~\c\ & \\\ & - @{command "fix"}~\x\~@{command "assume"}~\c: \\ \\ - @{command "sorry"} & \\\ & - @{command "by"}~\cheating\ \\ + \<^theory_text>\also"\<^sub>0"\ & \\\ & \<^theory_text>\note calculation = this\ \\ + \<^theory_text>\also"\<^sub>n\<^sub>+\<^sub>1"\ & \\\ & \<^theory_text>\note calculation = trans [OF calculation this]\ \\ + \<^theory_text>\finally\ & \\\ & \<^theory_text>\also from calculation\ \\[0.5ex] + \<^theory_text>\moreover\ & \\\ & \<^theory_text>\note calculation = calculation this\ \\ + \<^theory_text>\ultimately\ & \\\ & \<^theory_text>\moreover from calculation\ \\[0.5ex] + \<^theory_text>\presume a: A\ & \\\ & \<^theory_text>\assume a: A\ \\ + \<^theory_text>\def "x \ t"\ & \\\ & \<^theory_text>\fix x assume x_def: "x \ t"\ \\ + \<^theory_text>\consider x where A | \\ & \\\ & \<^theory_text>\have thesis\ \\ + & & \quad \<^theory_text>\if "\x. A \ thesis" and \ for thesis\ \\ + \<^theory_text>\obtain x where a: A \\ & \\\ & \<^theory_text>\consider x where A \\ \\ + & & \<^theory_text>\fix x assume a: A\ \\ + \<^theory_text>\case c\ & \\\ & \<^theory_text>\fix x assume c: A\ \\ + \<^theory_text>\sorry\ & \\\ & \<^theory_text>\by cheating\ \\ \end{tabular} \ @@ -96,12 +103,11 @@ text \ \begin{tabular}{ll} - @{command "print_state"} & print proof state \\ - @{command "print_statement"} & print fact in long statement form \\ - @{command "thm"}~\a\ & print fact \\ - @{command "prop"}~\\\ & print proposition \\ - @{command "term"}~\t\ & print term \\ - @{command "typ"}~\\\ & print type \\ + \<^theory_text>\typ \\ & print type \\ + \<^theory_text>\term t\ & print term \\ + \<^theory_text>\prop \\ & print proposition \\ + \<^theory_text>\thm a\ & print fact \\ + \<^theory_text>\print_statement a\ & print fact in long statement form \\ \end{tabular} \ @@ -111,7 +117,7 @@ text \ \begin{tabular}{ll} \multicolumn{2}{l}{\<^bold>\Single steps (forward-chaining facts)\} \\[0.5ex] - @{method assumption} & apply some assumption \\ + @{method assumption} & apply some goal assumption \\ @{method this} & apply current facts \\ @{method rule}~\a\ & apply some rule \\ @{method standard} & apply standard rule (default for @{command "proof"}) \\ @@ -123,6 +129,8 @@ @{method "-"} & no rules \\ @{method intro}~\a\ & introduction rules \\ @{method intro_classes} & class introduction rules \\ + @{method intro_locales} & locale introduction rules (without body) \\ + @{method unfold_locales} & locale introduction rules (with body) \\ @{method elim}~\a\ & elimination rules \\ @{method unfold}~\a\ & definitional rewrite rules \\[2ex] @@ -191,18 +199,20 @@ \ -section \Emulating tactic scripts\ +section \Proof scripts\ subsection \Commands\ text \ \begin{tabular}{ll} - @{command "apply"}~\m\ & apply proof method at initial position \\ - @{command "apply_end"}~\m\ & apply proof method near terminal position \\ - @{command "done"} & complete proof \\ - @{command "defer"}~\n\ & move subgoal to end \\ - @{command "prefer"}~\n\ & move subgoal to beginning \\ - @{command "back"} & backtrack last command \\ + \<^theory_text>\apply m\ & apply proof method during backwards refinement \\ + \<^theory_text>\apply_end m\ & apply proof method (as if in terminal position) \\ + \<^theory_text>\supply a\ & supply facts during backwards refinement \\ + \<^theory_text>\subgoal\ & nested proof during backwards refinement \\ + \<^theory_text>\defer n\ & move subgoal to end \\ + \<^theory_text>\prefer n\ & move subgoal to start \\ + \<^theory_text>\back\ & backtrack last command \\ + \<^theory_text>\done\ & complete proof \\ \end{tabular} \