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