| author | haftmann | 
| Thu, 13 Dec 2007 07:09:00 +0100 | |
| changeset 25612 | 314d949c70b5 | 
| parent 23373 | ead82c82da9e | 
| child 25976 | 11c6811f232c | 
| permissions | -rw-r--r-- | 
| 13405 | 1 | (* Title: HOL/Extraction/Warshall.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Stefan Berghofer, TU Muenchen | |
| 4 | *) | |
| 5 | ||
| 6 | header {* Warshall's algorithm *}
 | |
| 7 | ||
| 16761 | 8 | theory Warshall | 
| 9 | imports Main | |
| 10 | begin | |
| 13405 | 11 | |
| 12 | text {*
 | |
| 13 | Derivation of Warshall's algorithm using program extraction, | |
| 14 |   based on Berger, Schwichtenberg and Seisenberger \cite{Berger-JAR-2001}.
 | |
| 15 | *} | |
| 16 | ||
| 17 | datatype b = T | F | |
| 18 | ||
| 19 | consts | |
| 20 |   is_path' :: "('a \<Rightarrow> 'a \<Rightarrow> b) \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a \<Rightarrow> bool"
 | |
| 21 | ||
| 22 | primrec | |
| 23 | "is_path' r x [] z = (r x z = T)" | |
| 24 | "is_path' r x (y # ys) z = (r x y = T \<and> is_path' r y ys z)" | |
| 25 | ||
| 26 | constdefs | |
| 27 | is_path :: "(nat \<Rightarrow> nat \<Rightarrow> b) \<Rightarrow> (nat * nat list * nat) \<Rightarrow> | |
| 28 | nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" | |
| 29 | "is_path r p i j k == fst p = j \<and> snd (snd p) = k \<and> | |
| 30 | list_all (\<lambda>x. x < i) (fst (snd p)) \<and> | |
| 31 | is_path' r (fst p) (fst (snd p)) (snd (snd p))" | |
| 32 | ||
| 33 |   conc :: "('a * 'a list * 'a) \<Rightarrow> ('a * 'a list * 'a) \<Rightarrow> ('a * 'a list * 'a)"
 | |
| 34 | "conc p q == (fst p, fst (snd p) @ fst q # fst (snd q), snd (snd q))" | |
| 35 | ||
| 36 | theorem is_path'_snoc [simp]: | |
| 37 | "\<And>x. is_path' r x (ys @ [y]) z = (is_path' r x ys y \<and> r y z = T)" | |
| 38 | by (induct ys) simp+ | |
| 39 | ||
| 40 | theorem list_all_scoc [simp]: "list_all P (xs @ [x]) = (P x \<and> list_all P xs)" | |
| 17604 | 41 | by (induct xs, simp+, iprover) | 
| 13405 | 42 | |
| 43 | theorem list_all_lemma: | |
| 44 | "list_all P xs \<Longrightarrow> (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> list_all Q xs" | |
| 45 | proof - | |
| 46 | assume PQ: "\<And>x. P x \<Longrightarrow> Q x" | |
| 47 | show "list_all P xs \<Longrightarrow> list_all Q xs" | |
| 48 | proof (induct xs) | |
| 49 | case Nil | |
| 50 | show ?case by simp | |
| 51 | next | |
| 52 | case (Cons y ys) | |
| 53 | hence Py: "P y" by simp | |
| 54 | from Cons have Pys: "list_all P ys" by simp | |
| 55 | show ?case | |
| 56 | by simp (rule conjI PQ Py Cons Pys)+ | |
| 57 | qed | |
| 58 | qed | |
| 59 | ||
| 60 | theorem lemma1: "\<And>p. is_path r p i j k \<Longrightarrow> is_path r p (Suc i) j k" | |
| 61 | apply (unfold is_path_def) | |
| 62 | apply (simp cong add: conj_cong add: split_paired_all) | |
| 63 | apply (erule conjE)+ | |
| 64 | apply (erule list_all_lemma) | |
| 65 | apply simp | |
| 66 | done | |
| 67 | ||
| 68 | theorem lemma2: "\<And>p. is_path r p 0 j k \<Longrightarrow> r j k = T" | |
| 69 | apply (unfold is_path_def) | |
| 70 | apply (simp cong add: conj_cong add: split_paired_all) | |
| 71 | apply (case_tac "aa") | |
| 72 | apply simp+ | |
| 73 | done | |
| 74 | ||
| 75 | theorem is_path'_conc: "is_path' r j xs i \<Longrightarrow> is_path' r i ys k \<Longrightarrow> | |
| 76 | is_path' r j (xs @ i # ys) k" | |
| 77 | proof - | |
| 78 | assume pys: "is_path' r i ys k" | |
| 79 | show "\<And>j. is_path' r j xs i \<Longrightarrow> is_path' r j (xs @ i # ys) k" | |
| 80 | proof (induct xs) | |
| 81 | case (Nil j) | |
| 82 | hence "r j i = T" by simp | |
| 83 | with pys show ?case by simp | |
| 84 | next | |
| 85 | case (Cons z zs j) | |
| 86 | hence jzr: "r j z = T" by simp | |
| 87 | from Cons have pzs: "is_path' r z zs i" by simp | |
| 88 | show ?case | |
| 89 | by simp (rule conjI jzr Cons pzs)+ | |
| 90 | qed | |
| 91 | qed | |
| 92 | ||
| 93 | theorem lemma3: | |
| 94 | "\<And>p q. is_path r p i j i \<Longrightarrow> is_path r q i i k \<Longrightarrow> | |
| 95 | is_path r (conc p q) (Suc i) j k" | |
| 96 | apply (unfold is_path_def conc_def) | |
| 97 | apply (simp cong add: conj_cong add: split_paired_all) | |
| 98 | apply (erule conjE)+ | |
| 99 | apply (rule conjI) | |
| 100 | apply (erule list_all_lemma) | |
| 101 | apply simp | |
| 102 | apply (rule conjI) | |
| 103 | apply (erule list_all_lemma) | |
| 104 | apply simp | |
| 105 | apply (rule is_path'_conc) | |
| 106 | apply assumption+ | |
| 107 | done | |
| 108 | ||
| 109 | theorem lemma5: | |
| 110 | "\<And>p. is_path r p (Suc i) j k \<Longrightarrow> ~ is_path r p i j k \<Longrightarrow> | |
| 111 | (\<exists>q. is_path r q i j i) \<and> (\<exists>q'. is_path r q' i i k)" | |
| 112 | proof (simp cong add: conj_cong add: split_paired_all is_path_def, (erule conjE)+) | |
| 113 | fix xs | |
| 23373 | 114 | assume asms: | 
| 115 | "list_all (\<lambda>x. x < Suc i) xs" | |
| 116 | "is_path' r j xs k" | |
| 117 | "\<not> list_all (\<lambda>x. x < i) xs" | |
| 13405 | 118 | show "(\<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r j ys i) \<and> | 
| 119 | (\<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r i ys k)" | |
| 120 | proof | |
| 121 | show "\<And>j. list_all (\<lambda>x. x < Suc i) xs \<Longrightarrow> is_path' r j xs k \<Longrightarrow> | |
| 122 | \<not> list_all (\<lambda>x. x < i) xs \<Longrightarrow> | |
| 123 | \<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r j ys i" (is "PROP ?ih xs") | |
| 124 | proof (induct xs) | |
| 125 | case Nil | |
| 126 | thus ?case by simp | |
| 127 | next | |
| 128 | case (Cons a as j) | |
| 129 | show ?case | |
| 130 | proof (cases "a=i") | |
| 131 | case True | |
| 132 | show ?thesis | |
| 133 | proof | |
| 134 | from True and Cons have "r j i = T" by simp | |
| 135 | thus "list_all (\<lambda>x. x < i) [] \<and> is_path' r j [] i" by simp | |
| 136 | qed | |
| 137 | next | |
| 138 | case False | |
| 139 | have "PROP ?ih as" by (rule Cons) | |
| 140 | then obtain ys where ys: "list_all (\<lambda>x. x < i) ys \<and> is_path' r a ys i" | |
| 141 | proof | |
| 142 | from Cons show "list_all (\<lambda>x. x < Suc i) as" by simp | |
| 143 | from Cons show "is_path' r a as k" by simp | |
| 16761 | 144 | from Cons and False show "\<not> list_all (\<lambda>x. x < i) as" by (simp) | 
| 13405 | 145 | qed | 
| 146 | show ?thesis | |
| 147 | proof | |
| 148 | from Cons False ys | |
| 16761 | 149 | show "list_all (\<lambda>x. x<i) (a#ys) \<and> is_path' r j (a#ys) i" by simp | 
| 13405 | 150 | qed | 
| 151 | qed | |
| 152 | qed | |
| 153 | show "\<And>k. list_all (\<lambda>x. x < Suc i) xs \<Longrightarrow> is_path' r j xs k \<Longrightarrow> | |
| 154 | \<not> list_all (\<lambda>x. x < i) xs \<Longrightarrow> | |
| 155 | \<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r i ys k" (is "PROP ?ih xs") | |
| 156 | proof (induct xs rule: rev_induct) | |
| 157 | case Nil | |
| 158 | thus ?case by simp | |
| 159 | next | |
| 160 | case (snoc a as k) | |
| 161 | show ?case | |
| 162 | proof (cases "a=i") | |
| 163 | case True | |
| 164 | show ?thesis | |
| 165 | proof | |
| 166 | from True and snoc have "r i k = T" by simp | |
| 167 | thus "list_all (\<lambda>x. x < i) [] \<and> is_path' r i [] k" by simp | |
| 168 | qed | |
| 169 | next | |
| 170 | case False | |
| 171 | have "PROP ?ih as" by (rule snoc) | |
| 172 | then obtain ys where ys: "list_all (\<lambda>x. x < i) ys \<and> is_path' r i ys a" | |
| 173 | proof | |
| 174 | from snoc show "list_all (\<lambda>x. x < Suc i) as" by simp | |
| 175 | from snoc show "is_path' r j as a" by simp | |
| 16761 | 176 | from snoc and False show "\<not> list_all (\<lambda>x. x < i) as" by simp | 
| 13405 | 177 | qed | 
| 178 | show ?thesis | |
| 179 | proof | |
| 180 | from snoc False ys | |
| 181 | show "list_all (\<lambda>x. x < i) (ys @ [a]) \<and> is_path' r i (ys @ [a]) k" | |
| 16761 | 182 | by simp | 
| 13405 | 183 | qed | 
| 184 | qed | |
| 185 | qed | |
| 23373 | 186 | qed (rule asms)+ | 
| 13405 | 187 | qed | 
| 188 | ||
| 189 | theorem lemma5': | |
| 190 | "\<And>p. is_path r p (Suc i) j k \<Longrightarrow> \<not> is_path r p i j k \<Longrightarrow> | |
| 191 | \<not> (\<forall>q. \<not> is_path r q i j i) \<and> \<not> (\<forall>q'. \<not> is_path r q' i i k)" | |
| 17604 | 192 | by (iprover dest: lemma5) | 
| 13405 | 193 | |
| 194 | theorem warshall: | |
| 195 | "\<And>j k. \<not> (\<exists>p. is_path r p i j k) \<or> (\<exists>p. is_path r p i j k)" | |
| 196 | proof (induct i) | |
| 197 | case (0 j k) | |
| 198 | show ?case | |
| 199 | proof (cases "r j k") | |
| 200 | assume "r j k = T" | |
| 201 | hence "is_path r (j, [], k) 0 j k" | |
| 202 | by (simp add: is_path_def) | |
| 203 | hence "\<exists>p. is_path r p 0 j k" .. | |
| 204 | thus ?thesis .. | |
| 205 | next | |
| 206 | assume "r j k = F" | |
| 207 | hence "r j k ~= T" by simp | |
| 208 | hence "\<not> (\<exists>p. is_path r p 0 j k)" | |
| 17604 | 209 | by (iprover dest: lemma2) | 
| 13405 | 210 | thus ?thesis .. | 
| 211 | qed | |
| 212 | next | |
| 213 | case (Suc i j k) | |
| 214 | thus ?case | |
| 215 | proof | |
| 216 | assume h1: "\<not> (\<exists>p. is_path r p i j k)" | |
| 217 | from Suc show ?case | |
| 218 | proof | |
| 219 | assume "\<not> (\<exists>p. is_path r p i j i)" | |
| 220 | with h1 have "\<not> (\<exists>p. is_path r p (Suc i) j k)" | |
| 17604 | 221 | by (iprover dest: lemma5') | 
| 13405 | 222 | thus ?case .. | 
| 223 | next | |
| 224 | assume "\<exists>p. is_path r p i j i" | |
| 225 | then obtain p where h2: "is_path r p i j i" .. | |
| 226 | from Suc show ?case | |
| 227 | proof | |
| 228 | assume "\<not> (\<exists>p. is_path r p i i k)" | |
| 229 | with h1 have "\<not> (\<exists>p. is_path r p (Suc i) j k)" | |
| 17604 | 230 | by (iprover dest: lemma5') | 
| 13405 | 231 | thus ?case .. | 
| 232 | next | |
| 233 | assume "\<exists>q. is_path r q i i k" | |
| 234 | then obtain q where "is_path r q i i k" .. | |
| 235 | with h2 have "is_path r (conc p q) (Suc i) j k" | |
| 236 | by (rule lemma3) | |
| 237 | hence "\<exists>pq. is_path r pq (Suc i) j k" .. | |
| 238 | thus ?case .. | |
| 239 | qed | |
| 240 | qed | |
| 241 | next | |
| 242 | assume "\<exists>p. is_path r p i j k" | |
| 243 | hence "\<exists>p. is_path r p (Suc i) j k" | |
| 17604 | 244 | by (iprover intro: lemma1) | 
| 13405 | 245 | thus ?case .. | 
| 246 | qed | |
| 247 | qed | |
| 248 | ||
| 249 | extract warshall | |
| 250 | ||
| 251 | text {*
 | |
| 252 | 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 | 253 |   @{thm [display, eta_contract=false] warshall_def [no_vars]}
 | 
| 13405 | 254 | The corresponding correctness theorem is | 
| 255 |   @{thm [display] warshall_correctness [no_vars]}
 | |
| 256 | *} | |
| 257 | ||
| 258 | end |