| author | desharna | 
| Thu, 28 Mar 2024 08:30:42 +0100 | |
| changeset 80047 | 19cc354ba625 | 
| parent 76987 | 4c275405faae | 
| permissions | -rw-r--r-- | 
| 39157 
b98909faaea8
more explicit HOL-Proofs sessions, including former ex/Hilbert_Classical.thy which works in parallel mode without the antiquotation option "margin" (which is still critical);
 wenzelm parents: 
37599diff
changeset | 1 | (* Title: HOL/Proofs/Extraction/Warshall.thy | 
| 13405 | 2 | Author: Stefan Berghofer, TU Muenchen | 
| 3 | *) | |
| 4 | ||
| 61986 | 5 | section \<open>Warshall's algorithm\<close> | 
| 13405 | 6 | |
| 16761 | 7 | theory Warshall | 
| 67320 | 8 | imports "HOL-Library.Realizers" | 
| 16761 | 9 | begin | 
| 13405 | 10 | |
| 61986 | 11 | text \<open> | 
| 13405 | 12 | Derivation of Warshall's algorithm using program extraction, | 
| 76987 | 13 | based on Berger, Schwichtenberg and Seisenberger \<^cite>\<open>"Berger-JAR-2001"\<close>. | 
| 61986 | 14 | \<close> | 
| 13405 | 15 | |
| 58310 | 16 | datatype b = T | F | 
| 58272 | 17 | |
| 63361 | 18 | primrec is_path' :: "('a \<Rightarrow> 'a \<Rightarrow> b) \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 25976 | 19 | where | 
| 63361 | 20 | "is_path' r x [] z \<longleftrightarrow> r x z = T" | 
| 21 | | "is_path' r x (y # ys) z \<longleftrightarrow> r x y = T \<and> is_path' r y ys z" | |
| 13405 | 22 | |
| 63361 | 23 | definition is_path :: "(nat \<Rightarrow> nat \<Rightarrow> b) \<Rightarrow> (nat * nat list * nat) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" | 
| 24 | where "is_path r p i j k \<longleftrightarrow> | |
| 25 | fst p = j \<and> snd (snd p) = k \<and> | |
| 26 | list_all (\<lambda>x. x < i) (fst (snd p)) \<and> | |
| 27 | is_path' r (fst p) (fst (snd p)) (snd (snd p))" | |
| 13405 | 28 | |
| 63361 | 29 | definition conc :: "'a \<times> 'a list \<times> 'a \<Rightarrow> 'a \<times> 'a list \<times> 'a \<Rightarrow> 'a \<times> 'a list * 'a" | 
| 30 | where "conc p q = (fst p, fst (snd p) @ fst q # fst (snd q), snd (snd q))" | |
| 31 | ||
| 32 | theorem is_path'_snoc [simp]: "\<And>x. is_path' r x (ys @ [y]) z = (is_path' r x ys y \<and> r y z = T)" | |
| 13405 | 33 | by (induct ys) simp+ | 
| 34 | ||
| 37599 | 35 | theorem list_all_scoc [simp]: "list_all P (xs @ [x]) \<longleftrightarrow> P x \<and> list_all P xs" | 
| 63361 | 36 | by (induct xs) (simp+, iprover) | 
| 13405 | 37 | |
| 63361 | 38 | theorem list_all_lemma: "list_all P xs \<Longrightarrow> (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> list_all Q xs" | 
| 13405 | 39 | proof - | 
| 40 | assume PQ: "\<And>x. P x \<Longrightarrow> Q x" | |
| 41 | show "list_all P xs \<Longrightarrow> list_all Q xs" | |
| 42 | proof (induct xs) | |
| 43 | case Nil | |
| 44 | show ?case by simp | |
| 45 | next | |
| 46 | case (Cons y ys) | |
| 63361 | 47 | then have Py: "P y" by simp | 
| 13405 | 48 | from Cons have Pys: "list_all P ys" by simp | 
| 49 | show ?case | |
| 50 | by simp (rule conjI PQ Py Cons Pys)+ | |
| 51 | qed | |
| 52 | qed | |
| 53 | ||
| 54 | theorem lemma1: "\<And>p. is_path r p i j k \<Longrightarrow> is_path r p (Suc i) j k" | |
| 63361 | 55 | unfolding is_path_def | 
| 13405 | 56 | apply (simp cong add: conj_cong add: split_paired_all) | 
| 57 | apply (erule conjE)+ | |
| 58 | apply (erule list_all_lemma) | |
| 59 | apply simp | |
| 60 | done | |
| 61 | ||
| 62 | theorem lemma2: "\<And>p. is_path r p 0 j k \<Longrightarrow> r j k = T" | |
| 63361 | 63 | unfolding is_path_def | 
| 13405 | 64 | apply (simp cong add: conj_cong add: split_paired_all) | 
| 71989 
bad75618fb82
extraction of equations x = t from premises beneath meta-all
 haftmann parents: 
67320diff
changeset | 65 | apply (case_tac a) | 
| 
bad75618fb82
extraction of equations x = t from premises beneath meta-all
 haftmann parents: 
67320diff
changeset | 66 | apply simp_all | 
| 13405 | 67 | done | 
| 68 | ||
| 69 | theorem is_path'_conc: "is_path' r j xs i \<Longrightarrow> is_path' r i ys k \<Longrightarrow> | |
| 70 | is_path' r j (xs @ i # ys) k" | |
| 71 | proof - | |
| 72 | assume pys: "is_path' r i ys k" | |
| 73 | show "\<And>j. is_path' r j xs i \<Longrightarrow> is_path' r j (xs @ i # ys) k" | |
| 74 | proof (induct xs) | |
| 75 | case (Nil j) | |
| 63361 | 76 | then have "r j i = T" by simp | 
| 13405 | 77 | with pys show ?case by simp | 
| 78 | next | |
| 79 | case (Cons z zs j) | |
| 63361 | 80 | then have jzr: "r j z = T" by simp | 
| 13405 | 81 | from Cons have pzs: "is_path' r z zs i" by simp | 
| 82 | show ?case | |
| 83 | by simp (rule conjI jzr Cons pzs)+ | |
| 84 | qed | |
| 85 | qed | |
| 86 | ||
| 87 | theorem lemma3: | |
| 88 | "\<And>p q. is_path r p i j i \<Longrightarrow> is_path r q i i k \<Longrightarrow> | |
| 63361 | 89 | is_path r (conc p q) (Suc i) j k" | 
| 13405 | 90 | apply (unfold is_path_def conc_def) | 
| 91 | apply (simp cong add: conj_cong add: split_paired_all) | |
| 92 | apply (erule conjE)+ | |
| 93 | apply (rule conjI) | |
| 94 | apply (erule list_all_lemma) | |
| 95 | apply simp | |
| 96 | apply (rule conjI) | |
| 97 | apply (erule list_all_lemma) | |
| 98 | apply simp | |
| 99 | apply (rule is_path'_conc) | |
| 100 | apply assumption+ | |
| 101 | done | |
| 102 | ||
| 103 | theorem lemma5: | |
| 63361 | 104 | "\<And>p. is_path r p (Suc i) j k \<Longrightarrow> \<not> is_path r p i j k \<Longrightarrow> | 
| 105 | (\<exists>q. is_path r q i j i) \<and> (\<exists>q'. is_path r q' i i k)" | |
| 13405 | 106 | proof (simp cong add: conj_cong add: split_paired_all is_path_def, (erule conjE)+) | 
| 107 | fix xs | |
| 23373 | 108 | assume asms: | 
| 109 | "list_all (\<lambda>x. x < Suc i) xs" | |
| 110 | "is_path' r j xs k" | |
| 111 | "\<not> list_all (\<lambda>x. x < i) xs" | |
| 13405 | 112 | show "(\<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r j ys i) \<and> | 
| 113 | (\<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r i ys k)" | |
| 114 | proof | |
| 60595 | 115 | have "\<And>j. list_all (\<lambda>x. x < Suc i) xs \<Longrightarrow> is_path' r j xs k \<Longrightarrow> | 
| 13405 | 116 | \<not> list_all (\<lambda>x. x < i) xs \<Longrightarrow> | 
| 117 | \<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r j ys i" (is "PROP ?ih xs") | |
| 118 | proof (induct xs) | |
| 119 | case Nil | |
| 63361 | 120 | then show ?case by simp | 
| 13405 | 121 | next | 
| 122 | case (Cons a as j) | |
| 123 | show ?case | |
| 124 | proof (cases "a=i") | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 125 | case True | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 126 | show ?thesis | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 127 | proof | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 128 | from True and Cons have "r j i = T" by simp | 
| 63361 | 129 | then show "list_all (\<lambda>x. x < i) [] \<and> is_path' r j [] i" by simp | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 130 | qed | 
| 13405 | 131 | next | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 132 | case False | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 133 | have "PROP ?ih as" by (rule Cons) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 134 | then obtain ys where ys: "list_all (\<lambda>x. x < i) ys \<and> is_path' r a ys i" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 135 | proof | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 136 | from Cons show "list_all (\<lambda>x. x < Suc i) as" by simp | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 137 | from Cons show "is_path' r a as k" by simp | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 138 | from Cons and False show "\<not> list_all (\<lambda>x. x < i) as" by (simp) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 139 | qed | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 140 | show ?thesis | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 141 | proof | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 142 | from Cons False ys | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 143 | show "list_all (\<lambda>x. x<i) (a#ys) \<and> is_path' r j (a#ys) i" by simp | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 144 | qed | 
| 13405 | 145 | qed | 
| 146 | qed | |
| 60595 | 147 | from this asms show "\<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r j ys i" . | 
| 148 | have "\<And>k. list_all (\<lambda>x. x < Suc i) xs \<Longrightarrow> is_path' r j xs k \<Longrightarrow> | |
| 13405 | 149 | \<not> list_all (\<lambda>x. x < i) xs \<Longrightarrow> | 
| 150 | \<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r i ys k" (is "PROP ?ih xs") | |
| 151 | proof (induct xs rule: rev_induct) | |
| 152 | case Nil | |
| 63361 | 153 | then show ?case by simp | 
| 13405 | 154 | next | 
| 155 | case (snoc a as k) | |
| 156 | show ?case | |
| 157 | proof (cases "a=i") | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 158 | case True | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 159 | show ?thesis | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 160 | proof | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 161 | from True and snoc have "r i k = T" by simp | 
| 63361 | 162 | then show "list_all (\<lambda>x. x < i) [] \<and> is_path' r i [] k" by simp | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 163 | qed | 
| 13405 | 164 | next | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 165 | case False | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 166 | have "PROP ?ih as" by (rule snoc) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 167 | then obtain ys where ys: "list_all (\<lambda>x. x < i) ys \<and> is_path' r i ys a" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 168 | proof | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 169 | from snoc show "list_all (\<lambda>x. x < Suc i) as" by simp | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 170 | from snoc show "is_path' r j as a" by simp | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 171 | from snoc and False show "\<not> list_all (\<lambda>x. x < i) as" by simp | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 172 | qed | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 173 | show ?thesis | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 174 | proof | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 175 | from snoc False ys | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 176 | show "list_all (\<lambda>x. x < i) (ys @ [a]) \<and> is_path' r i (ys @ [a]) k" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 177 | by simp | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 178 | qed | 
| 13405 | 179 | qed | 
| 180 | qed | |
| 60595 | 181 | from this asms show "\<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r i ys k" . | 
| 182 | qed | |
| 13405 | 183 | qed | 
| 184 | ||
| 185 | theorem lemma5': | |
| 186 | "\<And>p. is_path r p (Suc i) j k \<Longrightarrow> \<not> is_path r p i j k \<Longrightarrow> | |
| 63361 | 187 | \<not> (\<forall>q. \<not> is_path r q i j i) \<and> \<not> (\<forall>q'. \<not> is_path r q' i i k)" | 
| 17604 | 188 | by (iprover dest: lemma5) | 
| 13405 | 189 | |
| 63361 | 190 | theorem warshall: "\<And>j k. \<not> (\<exists>p. is_path r p i j k) \<or> (\<exists>p. is_path r p i j k)" | 
| 13405 | 191 | proof (induct i) | 
| 192 | case (0 j k) | |
| 193 | show ?case | |
| 194 | proof (cases "r j k") | |
| 195 | assume "r j k = T" | |
| 63361 | 196 | then have "is_path r (j, [], k) 0 j k" | 
| 13405 | 197 | by (simp add: is_path_def) | 
| 63361 | 198 | then have "\<exists>p. is_path r p 0 j k" .. | 
| 199 | then show ?thesis .. | |
| 13405 | 200 | next | 
| 201 | assume "r j k = F" | |
| 63361 | 202 | then have "r j k \<noteq> T" by simp | 
| 203 | then have "\<not> (\<exists>p. is_path r p 0 j k)" | |
| 17604 | 204 | by (iprover dest: lemma2) | 
| 63361 | 205 | then show ?thesis .. | 
| 13405 | 206 | qed | 
| 207 | next | |
| 208 | case (Suc i j k) | |
| 63361 | 209 | then show ?case | 
| 13405 | 210 | proof | 
| 211 | assume h1: "\<not> (\<exists>p. is_path r p i j k)" | |
| 212 | from Suc show ?case | |
| 213 | proof | |
| 214 | assume "\<not> (\<exists>p. is_path r p i j i)" | |
| 215 | with h1 have "\<not> (\<exists>p. is_path r p (Suc i) j k)" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 216 | by (iprover dest: lemma5') | 
| 63361 | 217 | then show ?case .. | 
| 13405 | 218 | next | 
| 219 | assume "\<exists>p. is_path r p i j i" | |
| 220 | then obtain p where h2: "is_path r p i j i" .. | |
| 221 | from Suc show ?case | |
| 222 | proof | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 223 | assume "\<not> (\<exists>p. is_path r p i i k)" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 224 | with h1 have "\<not> (\<exists>p. is_path r p (Suc i) j k)" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 225 | by (iprover dest: lemma5') | 
| 63361 | 226 | then show ?case .. | 
| 13405 | 227 | next | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 228 | assume "\<exists>q. is_path r q i i k" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 229 | then obtain q where "is_path r q i i k" .. | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 230 | with h2 have "is_path r (conc p q) (Suc i) j k" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
29823diff
changeset | 231 | by (rule lemma3) | 
| 63361 | 232 | then have "\<exists>pq. is_path r pq (Suc i) j k" .. | 
| 233 | then show ?case .. | |
| 13405 | 234 | qed | 
| 235 | qed | |
| 236 | next | |
| 237 | assume "\<exists>p. is_path r p i j k" | |
| 63361 | 238 | then have "\<exists>p. is_path r p (Suc i) j k" | 
| 17604 | 239 | by (iprover intro: lemma1) | 
| 63361 | 240 | then show ?case .. | 
| 13405 | 241 | qed | 
| 242 | qed | |
| 243 | ||
| 244 | extract warshall | |
| 245 | ||
| 61986 | 246 | text \<open> | 
| 13405 | 247 | The program extracted from the above proof looks as follows | 
| 13671 
eec2582923f6
Eta contraction is now switched off when printing extracted program.
 berghofe parents: 
13471diff
changeset | 248 |   @{thm [display, eta_contract=false] warshall_def [no_vars]}
 | 
| 13405 | 249 | The corresponding correctness theorem is | 
| 250 |   @{thm [display] warshall_correctness [no_vars]}
 | |
| 61986 | 251 | \<close> | 
| 13405 | 252 | |
| 51272 | 253 | ML_val "@{code warshall}"
 | 
| 27982 | 254 | |
| 13405 | 255 | end |