| author | wenzelm | 
| Tue, 18 Apr 2017 19:17:46 +0200 | |
| changeset 65507 | decdb95bd007 | 
| parent 63531 | 847eefdca90d | 
| permissions | -rw-r--r-- | 
| 61656 | 1 | (*:maxLineLen=78:*) | 
| 2 | ||
| 26779 | 3 | theory Quick_Reference | 
| 63531 | 4 | imports Main Base | 
| 26779 | 5 | begin | 
| 6 | ||
| 58618 | 7 | chapter \<open>Isabelle/Isar quick reference \label{ap:refcard}\<close>
 | 
| 26779 | 8 | |
| 58618 | 9 | section \<open>Proof commands\<close> | 
| 26779 | 10 | |
| 62278 | 11 | subsection \<open>Main grammar \label{ap:main-grammar}\<close>
 | 
| 62270 | 12 | |
| 13 | text \<open> | |
| 14 |   \begin{tabular}{rcl}
 | |
| 15 | \<open>main\<close> & = & \<^theory_text>\<open>notepad begin "statement\<^sup>*" end\<close> \\ | |
| 63292 | 16 | & \<open>|\<close> & \<^theory_text>\<open>theorem name: props if name: props for vars "proof"\<close> \\ | 
| 62270 | 17 | & \<open>|\<close> & \<^theory_text>\<open>theorem name:\<close> \\ | 
| 63285 | 18 | & & \quad\<^theory_text>\<open>fixes vars\<close> \\ | 
| 62270 | 19 | & & \quad\<^theory_text>\<open>assumes name: props\<close> \\ | 
| 20 | & & \quad\<^theory_text>\<open>shows name: props "proof"\<close> \\ | |
| 63286 | 21 | & \<open>|\<close> & \<^theory_text>\<open>theorem name:\<close> \\ | 
| 22 | & & \quad\<^theory_text>\<open>fixes vars\<close> \\ | |
| 23 | & & \quad\<^theory_text>\<open>assumes name: props\<close> \\ | |
| 24 | & & \quad\<^theory_text>\<open>obtains (name) vars where props | \<dots> "proof"\<close> \\ | |
| 63298 | 25 | \<open>proof\<close> & = & \<^theory_text>\<open>"refinement\<^sup>*" proper_proof\<close> \\ | 
| 62270 | 26 | \<open>refinement\<close> & = & \<^theory_text>\<open>apply method\<close> \\ | 
| 63292 | 27 | & \<open>|\<close> & \<^theory_text>\<open>supply name = thms\<close> \\ | 
| 63286 | 28 | & \<open>|\<close> & \<^theory_text>\<open>subgoal premises name for vars "proof"\<close> \\ | 
| 63292 | 29 | & \<open>|\<close> & \<^theory_text>\<open>using thms\<close> \\ | 
| 30 | & \<open>|\<close> & \<^theory_text>\<open>unfolding thms\<close> \\ | |
| 63298 | 31 | \<open>proper_proof\<close> & = & \<^theory_text>\<open>proof "method\<^sup>?" "statement\<^sup>*" qed "method\<^sup>?"\<close> \\ | 
| 63293 | 32 | & \<open>|\<close> & \<^theory_text>\<open>done\<close> \\ | 
| 62270 | 33 |     \<open>statement\<close> & = & \<^theory_text>\<open>{ "statement\<^sup>*" }\<close> \\
 | 
| 34 | & \<open>|\<close> & \<^theory_text>\<open>next\<close> \\ | |
| 63292 | 35 | & \<open>|\<close> & \<^theory_text>\<open>note name = thms\<close> \\ | 
| 62270 | 36 | & \<open>|\<close> & \<^theory_text>\<open>let "term" = "term"\<close> \\ | 
| 37 | & \<open>|\<close> & \<^theory_text>\<open>write name (mixfix)\<close> \\ | |
| 63285 | 38 | & \<open>|\<close> & \<^theory_text>\<open>fix vars\<close> \\ | 
| 39 | & \<open>|\<close> & \<^theory_text>\<open>assume name: props if props for vars\<close> \\ | |
| 62270 | 40 | & \<open>|\<close> & \<^theory_text>\<open>then"\<^sup>?" goal\<close> \\ | 
| 63286 | 41 | \<open>goal\<close> & = & \<^theory_text>\<open>have name: props if name: props for vars "proof"\<close> \\ | 
| 63285 | 42 | & \<open>|\<close> & \<^theory_text>\<open>show name: props if name: props for vars "proof"\<close> \\ | 
| 62270 | 43 |   \end{tabular}
 | 
| 44 | \<close> | |
| 45 | ||
| 46 | ||
| 47 | subsection \<open>Primitives\<close> | |
| 26779 | 48 | |
| 58618 | 49 | text \<open> | 
| 26779 | 50 |   \begin{tabular}{ll}
 | 
| 62270 | 51 | \<^theory_text>\<open>fix x\<close> & augment context by \<open>\<And>x. \<box>\<close> \\ | 
| 52 | \<^theory_text>\<open>assume a: A\<close> & augment context by \<open>A \<Longrightarrow> \<box>\<close> \\ | |
| 53 | \<^theory_text>\<open>then\<close> & indicate forward chaining of facts \\ | |
| 54 | \<^theory_text>\<open>have a: A\<close> & prove local result \\ | |
| 55 | \<^theory_text>\<open>show a: A\<close> & prove local result, refining some goal \\ | |
| 56 | \<^theory_text>\<open>using a\<close> & indicate use of additional facts \\ | |
| 57 | \<^theory_text>\<open>unfolding a\<close> & unfold definitional equations \\ | |
| 58 | \<^theory_text>\<open>proof m\<^sub>1 \<dots> qed m\<^sub>2\<close> & indicate proof structure and refinements \\ | |
| 59 |     \<^theory_text>\<open>{ \<dots> }\<close> & indicate explicit blocks \\
 | |
| 60 | \<^theory_text>\<open>next\<close> & switch proof blocks \\ | |
| 61 | \<^theory_text>\<open>note a = b\<close> & reconsider and declare facts \\ | |
| 62 | \<^theory_text>\<open>let p = t\<close> & abbreviate terms by higher-order matching \\ | |
| 63 | \<^theory_text>\<open>write c (mx)\<close> & declare local mixfix syntax \\ | |
| 26852 | 64 |   \end{tabular}
 | 
| 58618 | 65 | \<close> | 
| 26779 | 66 | |
| 67 | ||
| 58618 | 68 | subsection \<open>Abbreviations and synonyms\<close> | 
| 26779 | 69 | |
| 58618 | 70 | text \<open> | 
| 26852 | 71 |   \begin{tabular}{rcl}
 | 
| 62270 | 72 | \<^theory_text>\<open>by m\<^sub>1 m\<^sub>2\<close> & \<open>\<equiv>\<close> & \<^theory_text>\<open>proof m\<^sub>1 qed m\<^sub>2\<close> \\ | 
| 73 | \<^theory_text>\<open>..\<close> & \<open>\<equiv>\<close> & \<^theory_text>\<open>by standard\<close> \\ | |
| 74 | \<^theory_text>\<open>.\<close> & \<open>\<equiv>\<close> & \<^theory_text>\<open>by this\<close> \\ | |
| 75 | \<^theory_text>\<open>from a\<close> & \<open>\<equiv>\<close> & \<^theory_text>\<open>note a then\<close> \\ | |
| 76 | \<^theory_text>\<open>with a\<close> & \<open>\<equiv>\<close> & \<^theory_text>\<open>from a and this\<close> \\ | |
| 77 | \<^theory_text>\<open>from this\<close> & \<open>\<equiv>\<close> & \<^theory_text>\<open>then\<close> \\ | |
| 26852 | 78 |   \end{tabular}
 | 
| 58618 | 79 | \<close> | 
| 26779 | 80 | |
| 81 | ||
| 58618 | 82 | subsection \<open>Derived elements\<close> | 
| 26779 | 83 | |
| 58618 | 84 | text \<open> | 
| 26852 | 85 |   \begin{tabular}{rcl}
 | 
| 62270 | 86 | \<^theory_text>\<open>also"\<^sub>0"\<close> & \<open>\<approx>\<close> & \<^theory_text>\<open>note calculation = this\<close> \\ | 
| 87 | \<^theory_text>\<open>also"\<^sub>n\<^sub>+\<^sub>1"\<close> & \<open>\<approx>\<close> & \<^theory_text>\<open>note calculation = trans [OF calculation this]\<close> \\ | |
| 88 | \<^theory_text>\<open>finally\<close> & \<open>\<approx>\<close> & \<^theory_text>\<open>also from calculation\<close> \\[0.5ex] | |
| 89 | \<^theory_text>\<open>moreover\<close> & \<open>\<approx>\<close> & \<^theory_text>\<open>note calculation = calculation this\<close> \\ | |
| 90 | \<^theory_text>\<open>ultimately\<close> & \<open>\<approx>\<close> & \<^theory_text>\<open>moreover from calculation\<close> \\[0.5ex] | |
| 91 | \<^theory_text>\<open>presume a: A\<close> & \<open>\<approx>\<close> & \<^theory_text>\<open>assume a: A\<close> \\ | |
| 63039 | 92 | \<^theory_text>\<open>define x where "x = t"\<close> & \<open>\<approx>\<close> & \<^theory_text>\<open>fix x assume x_def: "x = t"\<close> \\ | 
| 62270 | 93 | \<^theory_text>\<open>consider x where A | \<dots>\<close> & \<open>\<approx>\<close> & \<^theory_text>\<open>have thesis\<close> \\ | 
| 94 | & & \quad \<^theory_text>\<open>if "\<And>x. A \<Longrightarrow> thesis" and \<dots> for thesis\<close> \\ | |
| 95 | \<^theory_text>\<open>obtain x where a: A \<proof>\<close> & \<open>\<approx>\<close> & \<^theory_text>\<open>consider x where A \<proof>\<close> \\ | |
| 96 | & & \<^theory_text>\<open>fix x assume a: A\<close> \\ | |
| 97 | \<^theory_text>\<open>case c\<close> & \<open>\<approx>\<close> & \<^theory_text>\<open>fix x assume c: A\<close> \\ | |
| 98 | \<^theory_text>\<open>sorry\<close> & \<open>\<approx>\<close> & \<^theory_text>\<open>by cheating\<close> \\ | |
| 26852 | 99 |   \end{tabular}
 | 
| 58618 | 100 | \<close> | 
| 26779 | 101 | |
| 102 | ||
| 58618 | 103 | subsection \<open>Diagnostic commands\<close> | 
| 26779 | 104 | |
| 58618 | 105 | text \<open> | 
| 26779 | 106 |   \begin{tabular}{ll}
 | 
| 62270 | 107 | \<^theory_text>\<open>typ \<tau>\<close> & print type \\ | 
| 108 | \<^theory_text>\<open>term t\<close> & print term \\ | |
| 109 | \<^theory_text>\<open>prop \<phi>\<close> & print proposition \\ | |
| 110 | \<^theory_text>\<open>thm a\<close> & print fact \\ | |
| 111 | \<^theory_text>\<open>print_statement a\<close> & print fact in long statement form \\ | |
| 26779 | 112 |   \end{tabular}
 | 
| 58618 | 113 | \<close> | 
| 26779 | 114 | |
| 115 | ||
| 58618 | 116 | section \<open>Proof methods\<close> | 
| 26779 | 117 | |
| 58618 | 118 | text \<open> | 
| 26779 | 119 |   \begin{tabular}{ll}
 | 
| 61477 | 120 |     \multicolumn{2}{l}{\<^bold>\<open>Single steps (forward-chaining facts)\<close>} \\[0.5ex]
 | 
| 62270 | 121 |     @{method assumption} & apply some goal assumption \\
 | 
| 26779 | 122 |     @{method this} & apply current facts \\
 | 
| 61493 | 123 |     @{method rule}~\<open>a\<close> & apply some rule  \\
 | 
| 60618 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 wenzelm parents: 
58618diff
changeset | 124 |     @{method standard} & apply standard rule (default for @{command "proof"}) \\
 | 
| 61493 | 125 |     @{method contradiction} & apply \<open>\<not>\<close> elimination rule (any order) \\
 | 
| 126 |     @{method cases}~\<open>t\<close> & case analysis (provides cases) \\
 | |
| 127 |     @{method induct}~\<open>x\<close> & proof by induction (provides cases) \\[2ex]
 | |
| 26779 | 128 | |
| 61477 | 129 |     \multicolumn{2}{l}{\<^bold>\<open>Repeated steps (inserting facts)\<close>} \\[0.5ex]
 | 
| 26779 | 130 |     @{method "-"} & no rules \\
 | 
| 61493 | 131 |     @{method intro}~\<open>a\<close> & introduction rules \\
 | 
| 26779 | 132 |     @{method intro_classes} & class introduction rules \\
 | 
| 62270 | 133 |     @{method intro_locales} & locale introduction rules (without body) \\
 | 
| 134 |     @{method unfold_locales} & locale introduction rules (with body) \\
 | |
| 61493 | 135 |     @{method elim}~\<open>a\<close> & elimination rules \\
 | 
| 136 |     @{method unfold}~\<open>a\<close> & definitional rewrite rules \\[2ex]
 | |
| 26779 | 137 | |
| 61477 | 138 |     \multicolumn{2}{l}{\<^bold>\<open>Automated proof tools (inserting facts)\<close>} \\[0.5ex]
 | 
| 26779 | 139 |     @{method iprover} & intuitionistic proof search \\
 | 
| 140 |     @{method blast}, @{method fast} & Classical Reasoner \\
 | |
| 141 |     @{method simp}, @{method simp_all} & Simplifier (+ Splitter) \\
 | |
| 142 |     @{method auto}, @{method force} & Simplifier + Classical Reasoner \\
 | |
| 143 |     @{method arith} & Arithmetic procedures \\
 | |
| 144 |   \end{tabular}
 | |
| 58618 | 145 | \<close> | 
| 26779 | 146 | |
| 147 | ||
| 58618 | 148 | section \<open>Attributes\<close> | 
| 26779 | 149 | |
| 58618 | 150 | text \<open> | 
| 26779 | 151 |   \begin{tabular}{ll}
 | 
| 61477 | 152 |     \multicolumn{2}{l}{\<^bold>\<open>Rules\<close>} \\[0.5ex]
 | 
| 61493 | 153 |     @{attribute OF}~\<open>a\<close> & rule resolved with facts (skipping ``\<open>_\<close>'') \\
 | 
| 154 |     @{attribute of}~\<open>t\<close> & rule instantiated with terms (skipping ``\<open>_\<close>'') \\
 | |
| 155 |     @{attribute "where"}~\<open>x = t\<close> & rule instantiated with terms, by variable name \\
 | |
| 26779 | 156 |     @{attribute symmetric} & resolution with symmetry rule \\
 | 
| 61493 | 157 |     @{attribute THEN}~\<open>b\<close> & resolution with another rule \\
 | 
| 26779 | 158 |     @{attribute rule_format} & result put into standard rule format \\
 | 
| 159 |     @{attribute elim_format} & destruct rule turned into elimination rule format \\[1ex]
 | |
| 160 | ||
| 61477 | 161 |     \multicolumn{2}{l}{\<^bold>\<open>Declarations\<close>} \\[0.5ex]
 | 
| 26779 | 162 |     @{attribute simp} & Simplifier rule \\
 | 
| 163 |     @{attribute intro}, @{attribute elim}, @{attribute dest} & Pure or Classical Reasoner rule \\
 | |
| 164 |     @{attribute iff} & Simplifier + Classical Reasoner rule \\
 | |
| 165 |     @{attribute split} & case split rule \\
 | |
| 166 |     @{attribute trans} & transitivity rule \\
 | |
| 167 |     @{attribute sym} & symmetry rule \\
 | |
| 168 |   \end{tabular}
 | |
| 58618 | 169 | \<close> | 
| 26779 | 170 | |
| 171 | ||
| 58618 | 172 | section \<open>Rule declarations and methods\<close> | 
| 26779 | 173 | |
| 58618 | 174 | text \<open> | 
| 26779 | 175 |   \begin{tabular}{l|lllll}
 | 
| 176 |       & @{method rule} & @{method iprover} & @{method blast} & @{method simp} & @{method auto} \\
 | |
| 177 |       &                &                   & @{method fast} & @{method simp_all} & @{method force} \\
 | |
| 178 | \hline | |
| 61493 | 179 |     @{attribute Pure.elim}\<open>!\<close> @{attribute Pure.intro}\<open>!\<close>
 | 
| 180 | & \<open>\<times>\<close> & \<open>\<times>\<close> \\ | |
| 26779 | 181 |     @{attribute Pure.elim} @{attribute Pure.intro}
 | 
| 61493 | 182 | & \<open>\<times>\<close> & \<open>\<times>\<close> \\ | 
| 183 |     @{attribute elim}\<open>!\<close> @{attribute intro}\<open>!\<close>
 | |
| 184 | & \<open>\<times>\<close> & & \<open>\<times>\<close> & & \<open>\<times>\<close> \\ | |
| 26779 | 185 |     @{attribute elim} @{attribute intro}
 | 
| 61493 | 186 | & \<open>\<times>\<close> & & \<open>\<times>\<close> & & \<open>\<times>\<close> \\ | 
| 26779 | 187 |     @{attribute iff}
 | 
| 61493 | 188 | & \<open>\<times>\<close> & & \<open>\<times>\<close> & \<open>\<times>\<close> & \<open>\<times>\<close> \\ | 
| 189 |     @{attribute iff}\<open>?\<close>
 | |
| 190 | & \<open>\<times>\<close> \\ | |
| 191 |     @{attribute elim}\<open>?\<close> @{attribute intro}\<open>?\<close>
 | |
| 192 | & \<open>\<times>\<close> \\ | |
| 26779 | 193 |     @{attribute simp}
 | 
| 61493 | 194 | & & & & \<open>\<times>\<close> & \<open>\<times>\<close> \\ | 
| 26779 | 195 |     @{attribute cong}
 | 
| 61493 | 196 | & & & & \<open>\<times>\<close> & \<open>\<times>\<close> \\ | 
| 26779 | 197 |     @{attribute split}
 | 
| 61493 | 198 | & & & & \<open>\<times>\<close> & \<open>\<times>\<close> \\ | 
| 26779 | 199 |   \end{tabular}
 | 
| 58618 | 200 | \<close> | 
| 26779 | 201 | |
| 202 | ||
| 62270 | 203 | section \<open>Proof scripts\<close> | 
| 26779 | 204 | |
| 58618 | 205 | subsection \<open>Commands\<close> | 
| 26779 | 206 | |
| 58618 | 207 | text \<open> | 
| 26779 | 208 |   \begin{tabular}{ll}
 | 
| 62270 | 209 | \<^theory_text>\<open>apply m\<close> & apply proof method during backwards refinement \\ | 
| 210 | \<^theory_text>\<open>apply_end m\<close> & apply proof method (as if in terminal position) \\ | |
| 211 | \<^theory_text>\<open>supply a\<close> & supply facts during backwards refinement \\ | |
| 212 | \<^theory_text>\<open>subgoal\<close> & nested proof during backwards refinement \\ | |
| 213 | \<^theory_text>\<open>defer n\<close> & move subgoal to end \\ | |
| 214 | \<^theory_text>\<open>prefer n\<close> & move subgoal to start \\ | |
| 215 | \<^theory_text>\<open>back\<close> & backtrack last command \\ | |
| 216 | \<^theory_text>\<open>done\<close> & complete proof \\ | |
| 26779 | 217 |   \end{tabular}
 | 
| 58618 | 218 | \<close> | 
| 26779 | 219 | |
| 220 | ||
| 58618 | 221 | subsection \<open>Methods\<close> | 
| 26779 | 222 | |
| 58618 | 223 | text \<open> | 
| 26779 | 224 |   \begin{tabular}{ll}
 | 
| 61493 | 225 |     @{method rule_tac}~\<open>insts\<close> & resolution (with instantiation) \\
 | 
| 226 |     @{method erule_tac}~\<open>insts\<close> & elim-resolution (with instantiation) \\
 | |
| 227 |     @{method drule_tac}~\<open>insts\<close> & destruct-resolution (with instantiation) \\
 | |
| 228 |     @{method frule_tac}~\<open>insts\<close> & forward-resolution (with instantiation) \\
 | |
| 229 |     @{method cut_tac}~\<open>insts\<close> & insert facts (with instantiation) \\
 | |
| 230 |     @{method thin_tac}~\<open>\<phi>\<close> & delete assumptions \\
 | |
| 231 |     @{method subgoal_tac}~\<open>\<phi>\<close> & new claims \\
 | |
| 232 |     @{method rename_tac}~\<open>x\<close> & rename innermost goal parameters \\
 | |
| 233 |     @{method rotate_tac}~\<open>n\<close> & rotate assumptions of goal \\
 | |
| 234 |     @{method tactic}~\<open>text\<close> & arbitrary ML tactic \\
 | |
| 235 |     @{method case_tac}~\<open>t\<close> & exhaustion (datatypes) \\
 | |
| 236 |     @{method induct_tac}~\<open>x\<close> & induction (datatypes) \\
 | |
| 237 |     @{method ind_cases}~\<open>t\<close> & exhaustion + simplification (inductive predicates) \\
 | |
| 26779 | 238 |   \end{tabular}
 | 
| 58618 | 239 | \<close> | 
| 26779 | 240 | |
| 241 | end |