| author | wenzelm | 
| Fri, 11 Jul 2025 14:55:49 +0200 | |
| changeset 82845 | d4da7d857bb7 | 
| parent 63233 | e53830948c4f | 
| permissions | -rw-r--r-- | 
| 24608 | 1 | (* Title: HOL/Induct/Common_Patterns.thy | 
| 2 | Author: Makarius | |
| 3 | *) | |
| 4 | ||
| 60530 | 5 | section \<open>Common patterns of induction\<close> | 
| 24608 | 6 | |
| 7 | theory Common_Patterns | |
| 8 | imports Main | |
| 9 | begin | |
| 10 | ||
| 60530 | 11 | text \<open> | 
| 24608 | 12 | The subsequent Isar proof schemes illustrate common proof patterns | 
| 63167 | 13 | supported by the generic \<open>induct\<close> method. | 
| 24608 | 14 | |
| 15 | To demonstrate variations on statement (goal) structure we refer to | |
| 16 |   the induction rule of Peano natural numbers: @{thm nat.induct
 | |
| 17 | [no_vars]}, which is the simplest case of datatype induction. We | |
| 18 | shall also see more complex (mutual) datatype inductions involving | |
| 19 | several rules. Working with inductive predicates is similar, but | |
| 20 | involves explicit facts about membership, instead of implicit | |
| 21 | syntactic typing. | |
| 60530 | 22 | \<close> | 
| 24608 | 23 | |
| 24 | ||
| 60530 | 25 | subsection \<open>Variations on statement structure\<close> | 
| 24608 | 26 | |
| 60530 | 27 | subsubsection \<open>Local facts and parameters\<close> | 
| 24608 | 28 | |
| 60530 | 29 | text \<open> | 
| 24608 | 30 | Augmenting a problem by additional facts and locally fixed variables | 
| 31 | is a bread-and-butter method in many applications. This is where | |
| 63167 | 32 | unwieldy object-level \<open>\<forall>\<close> and \<open>\<longrightarrow>\<close> used to occur in | 
| 33 | the past. The \<open>induct\<close> method works with primary means of | |
| 24608 | 34 | the proof language instead. | 
| 60530 | 35 | \<close> | 
| 24608 | 36 | |
| 37 | lemma | |
| 38 | fixes n :: nat | |
| 39 | and x :: 'a | |
| 40 | assumes "A n x" | |
| 60530 | 41 | shows "P n x" using \<open>A n x\<close> | 
| 24608 | 42 | proof (induct n arbitrary: x) | 
| 43 | case 0 | |
| 60530 | 44 | note prem = \<open>A 0 x\<close> | 
| 62315 | 45 | show "P 0 x" \<proof> | 
| 24608 | 46 | next | 
| 47 | case (Suc n) | |
| 60530 | 48 | note hyp = \<open>\<And>x. A n x \<Longrightarrow> P n x\<close> | 
| 49 | and prem = \<open>A (Suc n) x\<close> | |
| 62315 | 50 | show "P (Suc n) x" \<proof> | 
| 24608 | 51 | qed | 
| 52 | ||
| 53 | ||
| 60530 | 54 | subsubsection \<open>Local definitions\<close> | 
| 24608 | 55 | |
| 60530 | 56 | text \<open> | 
| 24608 | 57 | Here the idea is to turn sub-expressions of the problem into a | 
| 58 | defined induction variable. This is often accompanied with fixing | |
| 59 | of auxiliary parameters in the original expression, otherwise the | |
| 60 | induction step would refer invariably to particular entities. This | |
| 61 | combination essentially expresses a partially abstracted | |
| 62 | representation of inductive expressions. | |
| 60530 | 63 | \<close> | 
| 24608 | 64 | |
| 65 | lemma | |
| 66 | fixes a :: "'a \<Rightarrow> nat" | |
| 67 | assumes "A (a x)" | |
| 60530 | 68 | shows "P (a x)" using \<open>A (a x)\<close> | 
| 24608 | 69 | proof (induct n \<equiv> "a x" arbitrary: x) | 
| 70 | case 0 | |
| 60530 | 71 | note prem = \<open>A (a x)\<close> | 
| 72 | and defn = \<open>0 = a x\<close> | |
| 62315 | 73 | show "P (a x)" \<proof> | 
| 24608 | 74 | next | 
| 75 | case (Suc n) | |
| 60530 | 76 | note hyp = \<open>\<And>x. n = a x \<Longrightarrow> A (a x) \<Longrightarrow> P (a x)\<close> | 
| 77 | and prem = \<open>A (a x)\<close> | |
| 78 | and defn = \<open>Suc n = a x\<close> | |
| 62315 | 79 | show "P (a x)" \<proof> | 
| 24608 | 80 | qed | 
| 81 | ||
| 60530 | 82 | text \<open> | 
| 63167 | 83 | Observe how the local definition \<open>n = a x\<close> recurs in the | 
| 84 | inductive cases as \<open>0 = a x\<close> and \<open>Suc n = a x\<close>, | |
| 24608 | 85 | according to underlying induction rule. | 
| 60530 | 86 | \<close> | 
| 24608 | 87 | |
| 88 | ||
| 60530 | 89 | subsubsection \<open>Simple simultaneous goals\<close> | 
| 24608 | 90 | |
| 60530 | 91 | text \<open> | 
| 24608 | 92 | The most basic simultaneous induction operates on several goals | 
| 93 | one-by-one, where each case refers to induction hypotheses that are | |
| 94 | duplicated according to the number of conclusions. | |
| 60530 | 95 | \<close> | 
| 24608 | 96 | |
| 97 | lemma | |
| 98 | fixes n :: nat | |
| 99 | shows "P n" and "Q n" | |
| 100 | proof (induct n) | |
| 101 | case 0 case 1 | |
| 62315 | 102 | show "P 0" \<proof> | 
| 24608 | 103 | next | 
| 104 | case 0 case 2 | |
| 62315 | 105 | show "Q 0" \<proof> | 
| 24608 | 106 | next | 
| 107 | case (Suc n) case 1 | |
| 60530 | 108 | note hyps = \<open>P n\<close> \<open>Q n\<close> | 
| 62315 | 109 | show "P (Suc n)" \<proof> | 
| 24608 | 110 | next | 
| 111 | case (Suc n) case 2 | |
| 60530 | 112 | note hyps = \<open>P n\<close> \<open>Q n\<close> | 
| 62315 | 113 | show "Q (Suc n)" \<proof> | 
| 24608 | 114 | qed | 
| 115 | ||
| 60530 | 116 | text \<open> | 
| 24608 | 117 | The split into subcases may be deferred as follows -- this is | 
| 118 | particularly relevant for goal statements with local premises. | |
| 60530 | 119 | \<close> | 
| 24608 | 120 | |
| 121 | lemma | |
| 122 | fixes n :: nat | |
| 123 | shows "A n \<Longrightarrow> P n" | |
| 124 | and "B n \<Longrightarrow> Q n" | |
| 125 | proof (induct n) | |
| 126 | case 0 | |
| 127 |   {
 | |
| 128 | case 1 | |
| 60530 | 129 | note \<open>A 0\<close> | 
| 62315 | 130 | show "P 0" \<proof> | 
| 24608 | 131 | next | 
| 132 | case 2 | |
| 60530 | 133 | note \<open>B 0\<close> | 
| 62315 | 134 | show "Q 0" \<proof> | 
| 24608 | 135 | } | 
| 136 | next | |
| 137 | case (Suc n) | |
| 60530 | 138 | note \<open>A n \<Longrightarrow> P n\<close> | 
| 139 | and \<open>B n \<Longrightarrow> Q n\<close> | |
| 24608 | 140 |   {
 | 
| 141 | case 1 | |
| 60530 | 142 | note \<open>A (Suc n)\<close> | 
| 62315 | 143 | show "P (Suc n)" \<proof> | 
| 24608 | 144 | next | 
| 145 | case 2 | |
| 60530 | 146 | note \<open>B (Suc n)\<close> | 
| 62315 | 147 | show "Q (Suc n)" \<proof> | 
| 24608 | 148 | } | 
| 149 | qed | |
| 150 | ||
| 151 | ||
| 60530 | 152 | subsubsection \<open>Compound simultaneous goals\<close> | 
| 24608 | 153 | |
| 60530 | 154 | text \<open> | 
| 24608 | 155 | The following pattern illustrates the slightly more complex | 
| 156 | situation of simultaneous goals with individual local assumptions. | |
| 157 | In compound simultaneous statements like this, local assumptions | |
| 63167 | 158 | need to be included into each goal, using \<open>\<Longrightarrow>\<close> of the Pure | 
| 24608 | 159 | framework. In contrast, local parameters do not require separate | 
| 63167 | 160 | \<open>\<And>\<close> prefixes here, but may be moved into the common context | 
| 24608 | 161 | of the whole statement. | 
| 60530 | 162 | \<close> | 
| 24608 | 163 | |
| 164 | lemma | |
| 165 | fixes n :: nat | |
| 166 | and x :: 'a | |
| 167 | and y :: 'b | |
| 168 | shows "A n x \<Longrightarrow> P n x" | |
| 169 | and "B n y \<Longrightarrow> Q n y" | |
| 170 | proof (induct n arbitrary: x y) | |
| 171 | case 0 | |
| 172 |   {
 | |
| 173 | case 1 | |
| 60530 | 174 | note prem = \<open>A 0 x\<close> | 
| 62315 | 175 | show "P 0 x" \<proof> | 
| 24608 | 176 | } | 
| 177 |   {
 | |
| 178 | case 2 | |
| 60530 | 179 | note prem = \<open>B 0 y\<close> | 
| 62315 | 180 | show "Q 0 y" \<proof> | 
| 24608 | 181 | } | 
| 182 | next | |
| 183 | case (Suc n) | |
| 60530 | 184 | note hyps = \<open>\<And>x. A n x \<Longrightarrow> P n x\<close> \<open>\<And>y. B n y \<Longrightarrow> Q n y\<close> | 
| 62315 | 185 | then have some_intermediate_result \<proof> | 
| 24608 | 186 |   {
 | 
| 187 | case 1 | |
| 60530 | 188 | note prem = \<open>A (Suc n) x\<close> | 
| 62315 | 189 | show "P (Suc n) x" \<proof> | 
| 24608 | 190 | } | 
| 191 |   {
 | |
| 192 | case 2 | |
| 60530 | 193 | note prem = \<open>B (Suc n) y\<close> | 
| 62315 | 194 | show "Q (Suc n) y" \<proof> | 
| 24608 | 195 | } | 
| 196 | qed | |
| 197 | ||
| 60530 | 198 | text \<open> | 
| 63167 | 199 | Here \<open>induct\<close> provides again nested cases with numbered | 
| 24608 | 200 | sub-cases, which allows to share common parts of the body context. | 
| 201 | In typical applications, there could be a long intermediate proof of | |
| 202 | general consequences of the induction hypotheses, before finishing | |
| 203 | each conclusion separately. | |
| 60530 | 204 | \<close> | 
| 24608 | 205 | |
| 206 | ||
| 60530 | 207 | subsection \<open>Multiple rules\<close> | 
| 24608 | 208 | |
| 60530 | 209 | text \<open> | 
| 24608 | 210 | Multiple induction rules emerge from mutual definitions of | 
| 63167 | 211 | datatypes, inductive predicates, functions etc. The \<open>induct\<close> method accepts replicated arguments (with \<open>and\<close> | 
| 24608 | 212 | separator), corresponding to each projection of the induction | 
| 213 | principle. | |
| 214 | ||
| 215 | The goal statement essentially follows the same arrangement, | |
| 216 | although it might be subdivided into simultaneous sub-problems as | |
| 217 | before! | |
| 60530 | 218 | \<close> | 
| 24608 | 219 | |
| 58310 | 220 | datatype foo = Foo1 nat | Foo2 bar | 
| 24608 | 221 | and bar = Bar1 bool | Bar2 bazar | 
| 222 | and bazar = Bazar foo | |
| 223 | ||
| 60530 | 224 | text \<open> | 
| 24608 | 225 |   The pack of induction rules for this datatype is: @{thm [display]
 | 
| 58270 | 226 | foo.induct [no_vars] bar.induct [no_vars] bazar.induct [no_vars]} | 
| 24608 | 227 | |
| 228 | This corresponds to the following basic proof pattern: | |
| 60530 | 229 | \<close> | 
| 24608 | 230 | |
| 231 | lemma | |
| 232 | fixes foo :: foo | |
| 233 | and bar :: bar | |
| 234 | and bazar :: bazar | |
| 235 | shows "P foo" | |
| 236 | and "Q bar" | |
| 237 | and "R bazar" | |
| 238 | proof (induct foo and bar and bazar) | |
| 239 | case (Foo1 n) | |
| 62315 | 240 | show "P (Foo1 n)" \<proof> | 
| 24608 | 241 | next | 
| 242 | case (Foo2 bar) | |
| 60530 | 243 | note \<open>Q bar\<close> | 
| 62315 | 244 | show "P (Foo2 bar)" \<proof> | 
| 24608 | 245 | next | 
| 246 | case (Bar1 b) | |
| 62315 | 247 | show "Q (Bar1 b)" \<proof> | 
| 24608 | 248 | next | 
| 249 | case (Bar2 bazar) | |
| 60530 | 250 | note \<open>R bazar\<close> | 
| 62315 | 251 | show "Q (Bar2 bazar)" \<proof> | 
| 24608 | 252 | next | 
| 253 | case (Bazar foo) | |
| 60530 | 254 | note \<open>P foo\<close> | 
| 62315 | 255 | show "R (Bazar foo)" \<proof> | 
| 24608 | 256 | qed | 
| 257 | ||
| 60530 | 258 | text \<open> | 
| 24608 | 259 | This can be combined with the previous techniques for compound | 
| 260 | statements, e.g.\ like this. | |
| 60530 | 261 | \<close> | 
| 24608 | 262 | |
| 263 | lemma | |
| 264 | fixes x :: 'a and y :: 'b and z :: 'c | |
| 265 | and foo :: foo | |
| 266 | and bar :: bar | |
| 267 | and bazar :: bazar | |
| 268 | shows | |
| 269 | "A x foo \<Longrightarrow> P x foo" | |
| 270 | and | |
| 271 | "B1 y bar \<Longrightarrow> Q1 y bar" | |
| 272 | "B2 y bar \<Longrightarrow> Q2 y bar" | |
| 273 | and | |
| 274 | "C1 z bazar \<Longrightarrow> R1 z bazar" | |
| 275 | "C2 z bazar \<Longrightarrow> R2 z bazar" | |
| 276 | "C3 z bazar \<Longrightarrow> R3 z bazar" | |
| 277 | proof (induct foo and bar and bazar arbitrary: x and y and z) | |
| 278 | oops | |
| 279 | ||
| 280 | ||
| 60530 | 281 | subsection \<open>Inductive predicates\<close> | 
| 24608 | 282 | |
| 60530 | 283 | text \<open> | 
| 24608 | 284 | The most basic form of induction involving predicates (or sets) | 
| 285 | essentially eliminates a given membership fact. | |
| 60530 | 286 | \<close> | 
| 24608 | 287 | |
| 288 | inductive Even :: "nat \<Rightarrow> bool" where | |
| 289 | zero: "Even 0" | |
| 63233 | 290 | | double: "Even (2 * n)" if "Even n" for n | 
| 24608 | 291 | |
| 292 | lemma | |
| 293 | assumes "Even n" | |
| 294 | shows "P n" | |
| 295 | using assms | |
| 296 | proof induct | |
| 297 | case zero | |
| 62315 | 298 | show "P 0" \<proof> | 
| 24608 | 299 | next | 
| 300 | case (double n) | |
| 60530 | 301 | note \<open>Even n\<close> and \<open>P n\<close> | 
| 62315 | 302 | show "P (2 * n)" \<proof> | 
| 24608 | 303 | qed | 
| 304 | ||
| 60530 | 305 | text \<open> | 
| 24608 | 306 | Alternatively, an initial rule statement may be proven as follows, | 
| 307 | performing ``in-situ'' elimination with explicit rule specification. | |
| 60530 | 308 | \<close> | 
| 24608 | 309 | |
| 310 | lemma "Even n \<Longrightarrow> P n" | |
| 311 | proof (induct rule: Even.induct) | |
| 312 | oops | |
| 313 | ||
| 60530 | 314 | text \<open> | 
| 24608 | 315 | Simultaneous goals do not introduce anything new. | 
| 60530 | 316 | \<close> | 
| 24608 | 317 | |
| 318 | lemma | |
| 319 | assumes "Even n" | |
| 320 | shows "P1 n" and "P2 n" | |
| 321 | using assms | |
| 322 | proof induct | |
| 323 | case zero | |
| 324 |   {
 | |
| 325 | case 1 | |
| 62315 | 326 | show "P1 0" \<proof> | 
| 24608 | 327 | next | 
| 328 | case 2 | |
| 62315 | 329 | show "P2 0" \<proof> | 
| 24608 | 330 | } | 
| 331 | next | |
| 332 | case (double n) | |
| 60530 | 333 | note \<open>Even n\<close> and \<open>P1 n\<close> and \<open>P2 n\<close> | 
| 24608 | 334 |   {
 | 
| 335 | case 1 | |
| 62315 | 336 | show "P1 (2 * n)" \<proof> | 
| 24608 | 337 | next | 
| 338 | case 2 | |
| 62315 | 339 | show "P2 (2 * n)" \<proof> | 
| 24608 | 340 | } | 
| 341 | qed | |
| 342 | ||
| 60530 | 343 | text \<open> | 
| 24608 | 344 | Working with mutual rules requires special care in composing the | 
| 345 | statement as a two-level conjunction, using lists of propositions | |
| 63167 | 346 | separated by \<open>and\<close>. For example: | 
| 60530 | 347 | \<close> | 
| 24608 | 348 | |
| 349 | inductive Evn :: "nat \<Rightarrow> bool" and Odd :: "nat \<Rightarrow> bool" | |
| 350 | where | |
| 351 | zero: "Evn 0" | |
| 63233 | 352 | | succ_Evn: "Odd (Suc n)" if "Evn n" for n | 
| 353 | | succ_Odd: "Evn (Suc n)" if "Odd n" for n | |
| 24608 | 354 | |
| 24609 | 355 | lemma | 
| 356 | "Evn n \<Longrightarrow> P1 n" | |
| 357 | "Evn n \<Longrightarrow> P2 n" | |
| 358 | "Evn n \<Longrightarrow> P3 n" | |
| 359 | and | |
| 360 | "Odd n \<Longrightarrow> Q1 n" | |
| 361 | "Odd n \<Longrightarrow> Q2 n" | |
| 24608 | 362 | proof (induct rule: Evn_Odd.inducts) | 
| 363 | case zero | |
| 62315 | 364 |   { case 1 show "P1 0" \<proof> }
 | 
| 365 |   { case 2 show "P2 0" \<proof> }
 | |
| 366 |   { case 3 show "P3 0" \<proof> }
 | |
| 24608 | 367 | next | 
| 368 | case (succ_Evn n) | |
| 60530 | 369 | note \<open>Evn n\<close> and \<open>P1 n\<close> \<open>P2 n\<close> \<open>P3 n\<close> | 
| 62315 | 370 |   { case 1 show "Q1 (Suc n)" \<proof> }
 | 
| 371 |   { case 2 show "Q2 (Suc n)" \<proof> }
 | |
| 24608 | 372 | next | 
| 373 | case (succ_Odd n) | |
| 60530 | 374 | note \<open>Odd n\<close> and \<open>Q1 n\<close> \<open>Q2 n\<close> | 
| 62315 | 375 |   { case 1 show "P1 (Suc n)" \<proof> }
 | 
| 376 |   { case 2 show "P2 (Suc n)" \<proof> }
 | |
| 377 |   { case 3 show "P3 (Suc n)" \<proof> }
 | |
| 24608 | 378 | qed | 
| 379 | ||
| 44164 
238c5ea1e2ce
documented extended version of case_names attribute
 nipkow parents: 
34915diff
changeset | 380 | |
| 60530 | 381 | text \<open> | 
| 44164 
238c5ea1e2ce
documented extended version of case_names attribute
 nipkow parents: 
34915diff
changeset | 382 | Cases and hypotheses in each case can be named explicitly. | 
| 60530 | 383 | \<close> | 
| 44164 
238c5ea1e2ce
documented extended version of case_names attribute
 nipkow parents: 
34915diff
changeset | 384 | |
| 
238c5ea1e2ce
documented extended version of case_names attribute
 nipkow parents: 
34915diff
changeset | 385 | inductive star :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" for r
 | 
| 
238c5ea1e2ce
documented extended version of case_names attribute
 nipkow parents: 
34915diff
changeset | 386 | where | 
| 63233 | 387 | refl: "star r x x" for x | 
| 388 | | step: "star r x z" if "r x y" and "star r y z" for x y z | |
| 44164 
238c5ea1e2ce
documented extended version of case_names attribute
 nipkow parents: 
34915diff
changeset | 389 | |
| 60530 | 390 | text \<open>Underscores are replaced by the default name hyps:\<close> | 
| 44164 
238c5ea1e2ce
documented extended version of case_names attribute
 nipkow parents: 
34915diff
changeset | 391 | |
| 53379 | 392 | lemmas star_induct = star.induct [case_names base step[r _ IH]] | 
| 44164 
238c5ea1e2ce
documented extended version of case_names attribute
 nipkow parents: 
34915diff
changeset | 393 | |
| 
238c5ea1e2ce
documented extended version of case_names attribute
 nipkow parents: 
34915diff
changeset | 394 | lemma "star r x y \<Longrightarrow> star r y z \<Longrightarrow> star r x z" | 
| 53379 | 395 | proof (induct rule: star_induct) print_cases | 
| 396 | case base | |
| 397 | then show ?case . | |
| 44164 
238c5ea1e2ce
documented extended version of case_names attribute
 nipkow parents: 
34915diff
changeset | 398 | next | 
| 53379 | 399 | case (step a b c) print_facts | 
| 400 | from step.prems have "star r b z" by (rule step.IH) | |
| 401 | with step.r show ?case by (rule star.step) | |
| 44164 
238c5ea1e2ce
documented extended version of case_names attribute
 nipkow parents: 
34915diff
changeset | 402 | qed | 
| 
238c5ea1e2ce
documented extended version of case_names attribute
 nipkow parents: 
34915diff
changeset | 403 | |
| 62390 | 404 | end |