|
1 (* $Id$ *) |
|
2 |
|
3 theory Quick_Reference |
|
4 imports Main |
|
5 begin |
|
6 |
|
7 chapter {* Isabelle/Isar quick reference \label{ap:refcard} *} |
|
8 |
|
9 section {* Proof commands *} |
|
10 |
|
11 subsection {* Primitives and basic syntax *} |
|
12 |
|
13 text {* |
|
14 \begin{tabular}{ll} |
|
15 @{command "fix"}~@{text x} & augment context by @{text "\<And>x. \<box>"} \\ |
|
16 @{command "assume"}~@{text "a: \<phi>"} & augment context by @{text "\<phi> \<Longrightarrow> \<box>"} \\ |
|
17 @{command "then"} & indicate forward chaining of facts \\ |
|
18 @{command "have"}~@{text "a: \<phi>"} & prove local result \\ |
|
19 @{command "show"}~@{text "a: \<phi>"} & prove local result, refining some goal \\ |
|
20 @{command "using"}~@{text a} & indicate use of additional facts \\ |
|
21 @{command "unfolding"}~@{text a} & unfold definitional equations \\ |
|
22 @{command "proof"}~@{text "m\<^sub>1"}~\dots~@{command "qed"}~@{text "m\<^sub>2"} & indicate proof structure and refinements \\ |
|
23 @{command "{"}~\dots~@{command "}"} & indicate explicit blocks \\ |
|
24 @{command "next"} & switch blocks \\ |
|
25 @{command "note"}~@{text "a = b"} & reconsider facts \\ |
|
26 @{command "let"}~@{text "p = t"} & abbreviate terms by higher-order matching \\ |
|
27 \end{tabular} |
|
28 |
|
29 \begin{matharray}{rcl} |
|
30 @{text "theory\<dash>stmt"} & = & @{command "theorem"}~@{text "name: prop proof"} \Or @{command "definition"}~\dots \Or \dots \\[1ex] |
|
31 @{text "proof"} & = & @{text "prfx\<^sup>*"}~@{command "proof"}~@{text "method stmt\<^sup>*"}~@{command "qed"}~@{text method} \\ |
|
32 & \Or & @{text "prfx\<^sup>*"}~@{command "done"} \\[1ex] |
|
33 @{text prfx} & = & @{command "apply"}~@{text method} \\ |
|
34 & \Or & @{command "using"}~@{text "name\<^sup>+"} \\ |
|
35 & \Or & @{command "unfolding"}~@{text "name\<^sup>+"} \\ |
|
36 @{text stmt} & = & @{command "{"}~@{text "stmt\<^sup>*"}~@{command "}"} \\ |
|
37 & \Or & @{command "next"} \\ |
|
38 & \Or & @{command "note"}~@{text "name = name\<^sup>+"} \\ |
|
39 & \Or & @{command "let"}~@{text "term = term"} \\ |
|
40 & \Or & @{command "fix"}~@{text "var\<^sup>+"} \\ |
|
41 & \Or & @{command "assume"}~@{text "name: prop\<^sup>+"} \\ |
|
42 & \Or & @{command "then"}@{text "\<^sup>?"}~@{text goal} \\ |
|
43 @{text goal} & = & @{command "have"}~@{text "name: prop\<^sup>+ proof"} \\ |
|
44 & \Or & @{command "show"}~@{text "name: prop\<^sup>+ proof"} \\ |
|
45 \end{matharray} |
|
46 *} |
|
47 |
|
48 |
|
49 subsection {* Abbreviations and synonyms *} |
|
50 |
|
51 text {* |
|
52 \begin{matharray}{rcl} |
|
53 @{command "by"}~@{text "m\<^sub>1 m\<^sub>2"} & \equiv & @{command "proof"}~@{text "m\<^sub>1"}~@{command "qed"}~@{text "m\<^sub>2"} \\ |
|
54 @{command ".."} & \equiv & @{command "by"}~@{text rule} \\ |
|
55 @{command "."} & \equiv & @{command "by"}~@{text this} \\ |
|
56 @{command "hence"} & \equiv & @{command "then"}~@{command "have"} \\ |
|
57 @{command "thus"} & \equiv & @{command "then"}~@{command "show"} \\ |
|
58 @{command "from"}~@{text a} & \equiv & @{command "note"}~@{text a}~@{command "then"} \\ |
|
59 @{command "with"}~@{text a} & \equiv & @{command "from"}~@{text "a \<AND> this"} \\[1ex] |
|
60 @{command "from"}~@{text this} & \equiv & @{command "then"} \\ |
|
61 @{command "from"}~@{text this}~@{command "have"} & \equiv & @{command "hence"} \\ |
|
62 @{command "from"}~@{text this}~@{command "show"} & \equiv & @{command "thus"} \\ |
|
63 \end{matharray} |
|
64 *} |
|
65 |
|
66 |
|
67 subsection {* Derived elements *} |
|
68 |
|
69 text {* |
|
70 \begin{matharray}{rcl} |
|
71 @{command "also"}@{text "\<^sub>0"} & \approx & @{command "note"}~@{text "calculation = this"} \\ |
|
72 @{command "also"}@{text "\<^sub>n\<^sub>+\<^sub>1"} & \approx & @{command "note"}~@{text "calculation = trans [OF calculation this]"} \\ |
|
73 @{command "finally"} & \approx & @{command "also"}~@{command "from"}~@{text calculation} \\[0.5ex] |
|
74 @{command "moreover"} & \approx & @{command "note"}~@{text "calculation = calculation this"} \\ |
|
75 @{command "ultimately"} & \approx & @{command "moreover"}~@{command "from"}~@{text calculation} \\[0.5ex] |
|
76 @{command "presume"}~@{text "a: \<phi>"} & \approx & @{command "assume"}~@{text "a: \<phi>"} \\ |
|
77 @{command "def"}~@{text "a: x \<equiv> t"} & \approx & @{command "fix"}~@{text x}~@{command "assume"}~@{text "a: x \<equiv> t"} \\ |
|
78 @{command "obtain"}~@{text "x \<WHERE> a: \<phi>"} & \approx & \dots~@{command "fix"}~@{text x}~@{command "assume"}~@{text "a: \<phi>"} \\ |
|
79 @{command "case"}~@{text c} & \approx & @{command "fix"}~@{text x}~@{command "assume"}~@{text "c: \<phi>"} \\ |
|
80 @{command "sorry"} & \approx & @{command "by"}~@{text cheating} \\ |
|
81 \end{matharray} |
|
82 *} |
|
83 |
|
84 |
|
85 subsection {* Diagnostic commands *} |
|
86 |
|
87 text {* |
|
88 \begin{tabular}{ll} |
|
89 @{command "pr"} & print current state \\ |
|
90 @{command "thm"}~@{text a} & print fact \\ |
|
91 @{command "term"}~@{text t} & print term \\ |
|
92 @{command "prop"}~@{text \<phi>} & print meta-level proposition \\ |
|
93 @{command "typ"}~@{text \<tau>} & print meta-level type \\ |
|
94 \end{tabular} |
|
95 *} |
|
96 |
|
97 |
|
98 section {* Proof methods *} |
|
99 |
|
100 text {* |
|
101 \begin{tabular}{ll} |
|
102 \multicolumn{2}{l}{\textbf{Single steps (forward-chaining facts)}} \\[0.5ex] |
|
103 @{method assumption} & apply some assumption \\ |
|
104 @{method this} & apply current facts \\ |
|
105 @{method rule}~@{text a} & apply some rule \\ |
|
106 @{method rule} & apply standard rule (default for @{command "proof"}) \\ |
|
107 @{method contradiction} & apply @{text "\<not>"} elimination rule (any order) \\ |
|
108 @{method cases}~@{text t} & case analysis (provides cases) \\ |
|
109 @{method induct}~@{text x} & proof by induction (provides cases) \\[2ex] |
|
110 |
|
111 \multicolumn{2}{l}{\textbf{Repeated steps (inserting facts)}} \\[0.5ex] |
|
112 @{method "-"} & no rules \\ |
|
113 @{method intro}~@{text a} & introduction rules \\ |
|
114 @{method intro_classes} & class introduction rules \\ |
|
115 @{method elim}~@{text a} & elimination rules \\ |
|
116 @{method unfold}~@{text a} & definitional rewrite rules \\[2ex] |
|
117 |
|
118 \multicolumn{2}{l}{\textbf{Automated proof tools (inserting facts)}} \\[0.5ex] |
|
119 @{method iprover} & intuitionistic proof search \\ |
|
120 @{method blast}, @{method fast} & Classical Reasoner \\ |
|
121 @{method simp}, @{method simp_all} & Simplifier (+ Splitter) \\ |
|
122 @{method auto}, @{method force} & Simplifier + Classical Reasoner \\ |
|
123 @{method arith} & Arithmetic procedures \\ |
|
124 \end{tabular} |
|
125 *} |
|
126 |
|
127 |
|
128 section {* Attributes *} |
|
129 |
|
130 text {* |
|
131 \begin{tabular}{ll} |
|
132 \multicolumn{2}{l}{\textbf{Operations}} \\[0.5ex] |
|
133 @{attribute OF}~@{text a} & rule resolved with facts (skipping ``@{text _}'') \\ |
|
134 @{attribute of}~@{text t} & rule instantiated with terms (skipping ``@{text _}'') \\ |
|
135 @{attribute "where"}~@{text "x = t"} & rule instantiated with terms, by variable name \\ |
|
136 @{attribute symmetric} & resolution with symmetry rule \\ |
|
137 @{attribute THEN}~@{text b} & resolution with another rule \\ |
|
138 @{attribute rule_format} & result put into standard rule format \\ |
|
139 @{attribute elim_format} & destruct rule turned into elimination rule format \\[1ex] |
|
140 |
|
141 \multicolumn{2}{l}{\textbf{Declarations}} \\[0.5ex] |
|
142 @{attribute simp} & Simplifier rule \\ |
|
143 @{attribute intro}, @{attribute elim}, @{attribute dest} & Pure or Classical Reasoner rule \\ |
|
144 @{attribute iff} & Simplifier + Classical Reasoner rule \\ |
|
145 @{attribute split} & case split rule \\ |
|
146 @{attribute trans} & transitivity rule \\ |
|
147 @{attribute sym} & symmetry rule \\ |
|
148 \end{tabular} |
|
149 *} |
|
150 |
|
151 |
|
152 section {* Rule declarations and methods *} |
|
153 |
|
154 text {* |
|
155 \begin{tabular}{l|lllll} |
|
156 & @{method rule} & @{method iprover} & @{method blast} & @{method simp} & @{method auto} \\ |
|
157 & & & @{method fast} & @{method simp_all} & @{method force} \\ |
|
158 \hline |
|
159 @{attribute Pure.elim}@{text "!"} @{attribute Pure.intro}@{text "!"} |
|
160 & @{text "\<times>"} & @{text "\<times>"} \\ |
|
161 @{attribute Pure.elim} @{attribute Pure.intro} |
|
162 & @{text "\<times>"} & @{text "\<times>"} \\ |
|
163 @{attribute elim}@{text "!"} @{attribute intro}@{text "!"} |
|
164 & @{text "\<times>"} & & @{text "\<times>"} & & @{text "\<times>"} \\ |
|
165 @{attribute elim} @{attribute intro} |
|
166 & @{text "\<times>"} & & @{text "\<times>"} & & @{text "\<times>"} \\ |
|
167 @{attribute iff} |
|
168 & @{text "\<times>"} & & @{text "\<times>"} & @{text "\<times>"} & @{text "\<times>"} \\ |
|
169 @{attribute iff}@{text "?"} |
|
170 & @{text "\<times>"} \\ |
|
171 @{attribute elim}@{text "?"} @{attribute intro}@{text "?"} |
|
172 & @{text "\<times>"} \\ |
|
173 @{attribute simp} |
|
174 & & & & @{text "\<times>"} & @{text "\<times>"} \\ |
|
175 @{attribute cong} |
|
176 & & & & @{text "\<times>"} & @{text "\<times>"} \\ |
|
177 @{attribute split} |
|
178 & & & & @{text "\<times>"} & @{text "\<times>"} \\ |
|
179 \end{tabular} |
|
180 *} |
|
181 |
|
182 |
|
183 section {* Emulating tactic scripts *} |
|
184 |
|
185 subsection {* Commands *} |
|
186 |
|
187 text {* |
|
188 \begin{tabular}{ll} |
|
189 @{command "apply"}~@{text m} & apply proof method at initial position \\ |
|
190 @{command "apply_end"}~@{text m} & apply proof method near terminal position \\ |
|
191 @{command "done"} & complete proof \\ |
|
192 @{command "defer"}~@{text n} & move subgoal to end \\ |
|
193 @{command "prefer"}~@{text n} & move subgoal to beginning \\ |
|
194 @{command "back"} & backtrack last command \\ |
|
195 \end{tabular} |
|
196 *} |
|
197 |
|
198 |
|
199 subsection {* Methods *} |
|
200 |
|
201 text {* |
|
202 \begin{tabular}{ll} |
|
203 @{method rule_tac}~@{text insts} & resolution (with instantiation) \\ |
|
204 @{method erule_tac}~@{text insts} & elim-resolution (with instantiation) \\ |
|
205 @{method drule_tac}~@{text insts} & destruct-resolution (with instantiation) \\ |
|
206 @{method frule_tac}~@{text insts} & forward-resolution (with instantiation) \\ |
|
207 @{method cut_tac}~@{text insts} & insert facts (with instantiation) \\ |
|
208 @{method thin_tac}~@{text \<phi>} & delete assumptions \\ |
|
209 @{method subgoal_tac}~@{text \<phi>} & new claims \\ |
|
210 @{method rename_tac}~@{text x} & rename innermost goal parameters \\ |
|
211 @{method rotate_tac}~@{text n} & rotate assumptions of goal \\ |
|
212 @{method tactic}~@{text "text"} & arbitrary ML tactic \\ |
|
213 @{method case_tac}~@{text t} & exhaustion (datatypes) \\ |
|
214 @{method induct_tac}~@{text x} & induction (datatypes) \\ |
|
215 @{method ind_cases}~@{text t} & exhaustion + simplification (inductive predicates) \\ |
|
216 \end{tabular} |
|
217 *} |
|
218 |
|
219 end |