| author | wenzelm | 
| Mon, 28 Feb 2022 13:02:40 +0100 | |
| changeset 75166 | 1f6da5d18340 | 
| parent 69597 | ff784d5a5bfb | 
| permissions | -rw-r--r-- | 
| 60450 | 1 | (* Title: HOL/Isar_Examples/Structured_Statements.thy | 
| 2 | Author: Makarius | |
| 3 | *) | |
| 4 | ||
| 5 | section \<open>Structured statements within Isar proofs\<close> | |
| 6 | ||
| 7 | theory Structured_Statements | |
| 63583 | 8 | imports Main | 
| 60450 | 9 | begin | 
| 10 | ||
| 60470 | 11 | subsection \<open>Introduction steps\<close> | 
| 12 | ||
| 60450 | 13 | notepad | 
| 14 | begin | |
| 60470 | 15 | fix A B :: bool | 
| 16 | fix P :: "'a \<Rightarrow> bool" | |
| 60450 | 17 | |
| 60470 | 18 | have "A \<longrightarrow> B" | 
| 19 | proof | |
| 62314 | 20 | show B if A using that \<proof> | 
| 60470 | 21 | qed | 
| 22 | ||
| 23 | have "\<not> A" | |
| 24 | proof | |
| 62314 | 25 | show False if A using that \<proof> | 
| 60470 | 26 | qed | 
| 27 | ||
| 28 | have "\<forall>x. P x" | |
| 29 | proof | |
| 62314 | 30 | show "P x" for x \<proof> | 
| 60470 | 31 | qed | 
| 32 | end | |
| 33 | ||
| 34 | ||
| 35 | subsection \<open>If-and-only-if\<close> | |
| 36 | ||
| 37 | notepad | |
| 38 | begin | |
| 39 | fix A B :: bool | |
| 40 | ||
| 41 | have "A \<longleftrightarrow> B" | |
| 42 | proof | |
| 62314 | 43 | show B if A \<proof> | 
| 44 | show A if B \<proof> | |
| 60470 | 45 | qed | 
| 46 | next | |
| 60450 | 47 | fix A B :: bool | 
| 48 | ||
| 49 | have iff_comm: "(A \<and> B) \<longleftrightarrow> (B \<and> A)" | |
| 50 | proof | |
| 51 | show "B \<and> A" if "A \<and> B" | |
| 52 | proof | |
| 53 | show B using that .. | |
| 54 | show A using that .. | |
| 55 | qed | |
| 56 | show "A \<and> B" if "B \<and> A" | |
| 57 | proof | |
| 58 | show A using that .. | |
| 59 | show B using that .. | |
| 60 | qed | |
| 61 | qed | |
| 62 | ||
| 63 | text \<open>Alternative proof, avoiding redundant copy of symmetric argument.\<close> | |
| 64 | have iff_comm: "(A \<and> B) \<longleftrightarrow> (B \<and> A)" | |
| 65 | proof | |
| 66 | show "B \<and> A" if "A \<and> B" for A B | |
| 67 | proof | |
| 68 | show B using that .. | |
| 69 | show A using that .. | |
| 70 | qed | |
| 71 | then show "A \<and> B" if "B \<and> A" | |
| 72 | by this (rule that) | |
| 73 | qed | |
| 60470 | 74 | end | 
| 60450 | 75 | |
| 60470 | 76 | |
| 77 | subsection \<open>Elimination and cases\<close> | |
| 78 | ||
| 79 | notepad | |
| 80 | begin | |
| 81 | fix A B C D :: bool | |
| 82 | assume *: "A \<or> B \<or> C \<or> D" | |
| 83 | ||
| 84 | consider (a) A | (b) B | (c) C | (d) D | |
| 85 | using * by blast | |
| 86 | then have something | |
| 87 | proof cases | |
| 88 | case a thm \<open>A\<close> | |
| 62314 | 89 | then show ?thesis \<proof> | 
| 60470 | 90 | next | 
| 91 | case b thm \<open>B\<close> | |
| 62314 | 92 | then show ?thesis \<proof> | 
| 60470 | 93 | next | 
| 94 | case c thm \<open>C\<close> | |
| 62314 | 95 | then show ?thesis \<proof> | 
| 60470 | 96 | next | 
| 97 | case d thm \<open>D\<close> | |
| 62314 | 98 | then show ?thesis \<proof> | 
| 60470 | 99 | qed | 
| 100 | next | |
| 101 | fix A :: "'a \<Rightarrow> bool" | |
| 102 | fix B :: "'b \<Rightarrow> 'c \<Rightarrow> bool" | |
| 103 | assume *: "(\<exists>x. A x) \<or> (\<exists>y z. B y z)" | |
| 104 | ||
| 105 | consider (a) x where "A x" | (b) y z where "B y z" | |
| 106 | using * by blast | |
| 107 | then have something | |
| 108 | proof cases | |
| 109 | case a thm \<open>A x\<close> | |
| 62314 | 110 | then show ?thesis \<proof> | 
| 60470 | 111 | next | 
| 112 | case b thm \<open>B y z\<close> | |
| 62314 | 113 | then show ?thesis \<proof> | 
| 60470 | 114 | qed | 
| 115 | end | |
| 116 | ||
| 117 | ||
| 118 | subsection \<open>Induction\<close> | |
| 119 | ||
| 120 | notepad | |
| 121 | begin | |
| 122 | fix P :: "nat \<Rightarrow> bool" | |
| 123 | fix n :: nat | |
| 124 | ||
| 125 | have "P n" | |
| 126 | proof (induct n) | |
| 62314 | 127 | show "P 0" \<proof> | 
| 60470 | 128 | show "P (Suc n)" if "P n" for n thm \<open>P n\<close> | 
| 62314 | 129 | using that \<proof> | 
| 60470 | 130 | qed | 
| 60450 | 131 | end | 
| 132 | ||
| 60555 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 133 | |
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 134 | subsection \<open>Suffices-to-show\<close> | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 135 | |
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 136 | notepad | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 137 | begin | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 138 | fix A B C | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 139 | assume r: "A \<Longrightarrow> B \<Longrightarrow> C" | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 140 | |
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 141 | have C | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 142 | proof - | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 143 | show ?thesis when A (is ?A) and B (is ?B) | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 144 | using that by (rule r) | 
| 62314 | 145 | show ?A \<proof> | 
| 146 | show ?B \<proof> | |
| 60555 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 147 | qed | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 148 | next | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 149 | fix a :: 'a | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 150 | fix A :: "'a \<Rightarrow> bool" | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 151 | fix C | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 152 | |
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 153 | have C | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 154 | proof - | 
| 69597 | 155 | show ?thesis when "A x" (is ?A) for x :: 'a \<comment> \<open>abstract \<^term>\<open>x\<close>\<close> | 
| 62314 | 156 | using that \<proof> | 
| 69597 | 157 | show "?A a" \<comment> \<open>concrete \<^term>\<open>a\<close>\<close> | 
| 62314 | 158 | \<proof> | 
| 60555 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 159 | qed | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 160 | end | 
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 161 | |
| 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 wenzelm parents: 
60470diff
changeset | 162 | end |