| author | wenzelm | 
| Wed, 03 Apr 2024 11:02:09 +0200 | |
| changeset 80077 | ee07b7738a24 | 
| parent 74563 | 042041c0ebeb | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 17456 | 1 | (* Title: CCL/Wfd.thy | 
| 1474 | 2 | Author: Martin Coen, Cambridge University Computer Laboratory | 
| 0 | 3 | Copyright 1993 University of Cambridge | 
| 4 | *) | |
| 5 | ||
| 60770 | 6 | section \<open>Well-founded relations in CCL\<close> | 
| 17456 | 7 | |
| 8 | theory Wfd | |
| 9 | imports Trancl Type Hered | |
| 10 | begin | |
| 0 | 11 | |
| 58977 | 12 | definition Wfd :: "[i set] \<Rightarrow> o" | 
| 13 | where "Wfd(R) == ALL P.(ALL x.(ALL y.<y,x> : R \<longrightarrow> y:P) \<longrightarrow> x:P) \<longrightarrow> (ALL a. a:P)" | |
| 0 | 14 | |
| 58977 | 15 | definition wf :: "[i set] \<Rightarrow> i set" | 
| 16 |   where "wf(R) == {x. x:R \<and> Wfd(R)}"
 | |
| 57521 | 17 | |
| 58977 | 18 | definition wmap :: "[i\<Rightarrow>i,i set] \<Rightarrow> i set" | 
| 19 |   where "wmap(f,R) == {p. EX x y. p=<x,y> \<and> <f(x),f(y)> : R}"
 | |
| 0 | 20 | |
| 57521 | 21 | definition lex :: "[i set,i set] => i set" (infixl "**" 70) | 
| 58977 | 22 |   where "ra**rb == {p. EX a a' b b'. p = <<a,b>,<a',b'>> \<and> (<a,a'> : ra | (a=a' \<and> <b,b'> : rb))}"
 | 
| 0 | 23 | |
| 57521 | 24 | definition NatPR :: "i set" | 
| 25 |   where "NatPR == {p. EX x:Nat. p=<x,succ(x)>}"
 | |
| 0 | 26 | |
| 58977 | 27 | definition ListPR :: "i set \<Rightarrow> i set" | 
| 57521 | 28 |   where "ListPR(A) == {p. EX h:A. EX t:List(A). p=<t,h$t>}"
 | 
| 17456 | 29 | |
| 20140 | 30 | |
| 31 | lemma wfd_induct: | |
| 32 | assumes 1: "Wfd(R)" | |
| 58977 | 33 | and 2: "\<And>x. ALL y. <y,x>: R \<longrightarrow> P(y) \<Longrightarrow> P(x)" | 
| 20140 | 34 | shows "P(a)" | 
| 35 | apply (rule 1 [unfolded Wfd_def, rule_format, THEN CollectD]) | |
| 36 | using 2 apply blast | |
| 37 | done | |
| 38 | ||
| 39 | lemma wfd_strengthen_lemma: | |
| 58977 | 40 | assumes 1: "\<And>x y.<x,y> : R \<Longrightarrow> Q(x)" | 
| 41 | and 2: "ALL x. (ALL y. <y,x> : R \<longrightarrow> y : P) \<longrightarrow> x : P" | |
| 42 | and 3: "\<And>x. Q(x) \<Longrightarrow> x:P" | |
| 20140 | 43 | shows "a:P" | 
| 44 | apply (rule 2 [rule_format]) | |
| 45 | using 1 3 | |
| 46 | apply blast | |
| 47 | done | |
| 48 | ||
| 60770 | 49 | method_setup wfd_strengthen = \<open> | 
| 74563 | 50 | Scan.lift Parse.embedded_inner_syntax >> (fn s => fn ctxt => | 
| 58971 | 51 | SIMPLE_METHOD' (fn i => | 
| 59780 | 52 |       Rule_Insts.res_inst_tac ctxt [((("Q", 0), Position.none), s)] [] @{thm wfd_strengthen_lemma} i
 | 
| 59755 | 53 | THEN assume_tac ctxt (i + 1))) | 
| 60770 | 54 | \<close> | 
| 20140 | 55 | |
| 58977 | 56 | lemma wf_anti_sym: "\<lbrakk>Wfd(r); <a,x>:r; <x,a>:r\<rbrakk> \<Longrightarrow> P" | 
| 57 | apply (subgoal_tac "ALL x. <a,x>:r \<longrightarrow> <x,a>:r \<longrightarrow> P") | |
| 20140 | 58 | apply blast | 
| 59 | apply (erule wfd_induct) | |
| 60 | apply blast | |
| 61 | done | |
| 62 | ||
| 58977 | 63 | lemma wf_anti_refl: "\<lbrakk>Wfd(r); <a,a>: r\<rbrakk> \<Longrightarrow> P" | 
| 20140 | 64 | apply (rule wf_anti_sym) | 
| 65 | apply assumption+ | |
| 66 | done | |
| 67 | ||
| 68 | ||
| 60770 | 69 | subsection \<open>Irreflexive transitive closure\<close> | 
| 20140 | 70 | |
| 71 | lemma trancl_wf: | |
| 72 | assumes 1: "Wfd(R)" | |
| 73 | shows "Wfd(R^+)" | |
| 74 | apply (unfold Wfd_def) | |
| 75 | apply (rule allI ballI impI)+ | |
| 76 | (*must retain the universal formula for later use!*) | |
| 77 | apply (rule allE, assumption) | |
| 78 | apply (erule mp) | |
| 79 | apply (rule 1 [THEN wfd_induct]) | |
| 80 | apply (rule impI [THEN allI]) | |
| 81 | apply (erule tranclE) | |
| 82 | apply blast | |
| 83 | apply (erule spec [THEN mp, THEN spec, THEN mp]) | |
| 84 | apply assumption+ | |
| 85 | done | |
| 86 | ||
| 87 | ||
| 60770 | 88 | subsection \<open>Lexicographic Ordering\<close> | 
| 20140 | 89 | |
| 90 | lemma lexXH: | |
| 58977 | 91 | "p : ra**rb \<longleftrightarrow> (EX a a' b b'. p = <<a,b>,<a',b'>> \<and> (<a,a'> : ra | a=a' \<and> <b,b'> : rb))" | 
| 20140 | 92 | unfolding lex_def by blast | 
| 93 | ||
| 58977 | 94 | lemma lexI1: "<a,a'> : ra \<Longrightarrow> <<a,b>,<a',b'>> : ra**rb" | 
| 20140 | 95 | by (blast intro!: lexXH [THEN iffD2]) | 
| 96 | ||
| 58977 | 97 | lemma lexI2: "<b,b'> : rb \<Longrightarrow> <<a,b>,<a,b'>> : ra**rb" | 
| 20140 | 98 | by (blast intro!: lexXH [THEN iffD2]) | 
| 99 | ||
| 100 | lemma lexE: | |
| 101 | assumes 1: "p : ra**rb" | |
| 58977 | 102 | and 2: "\<And>a a' b b'. \<lbrakk><a,a'> : ra; p=<<a,b>,<a',b'>>\<rbrakk> \<Longrightarrow> R" | 
| 103 | and 3: "\<And>a b b'. \<lbrakk><b,b'> : rb; p = <<a,b>,<a,b'>>\<rbrakk> \<Longrightarrow> R" | |
| 20140 | 104 | shows R | 
| 105 | apply (rule 1 [THEN lexXH [THEN iffD1], THEN exE]) | |
| 106 | using 2 3 | |
| 107 | apply blast | |
| 108 | done | |
| 109 | ||
| 58977 | 110 | lemma lex_pair: "\<lbrakk>p : r**s; \<And>a a' b b'. p = <<a,b>,<a',b'>> \<Longrightarrow> P\<rbrakk> \<Longrightarrow>P" | 
| 20140 | 111 | apply (erule lexE) | 
| 112 | apply blast+ | |
| 113 | done | |
| 114 | ||
| 115 | lemma lex_wf: | |
| 116 | assumes 1: "Wfd(R)" | |
| 117 | and 2: "Wfd(S)" | |
| 118 | shows "Wfd(R**S)" | |
| 119 | apply (unfold Wfd_def) | |
| 120 | apply safe | |
| 58977 | 121 | apply (wfd_strengthen "\<lambda>x. EX a b. x=<a,b>") | 
| 20140 | 122 | apply (blast elim!: lex_pair) | 
| 123 | apply (subgoal_tac "ALL a b.<a,b>:P") | |
| 124 | apply blast | |
| 125 | apply (rule 1 [THEN wfd_induct, THEN allI]) | |
| 126 | apply (rule 2 [THEN wfd_induct, THEN allI]) back | |
| 127 | apply (fast elim!: lexE) | |
| 128 | done | |
| 129 | ||
| 130 | ||
| 60770 | 131 | subsection \<open>Mapping\<close> | 
| 20140 | 132 | |
| 58977 | 133 | lemma wmapXH: "p : wmap(f,r) \<longleftrightarrow> (EX x y. p=<x,y> \<and> <f(x),f(y)> : r)" | 
| 20140 | 134 | unfolding wmap_def by blast | 
| 135 | ||
| 58977 | 136 | lemma wmapI: "<f(a),f(b)> : r \<Longrightarrow> <a,b> : wmap(f,r)" | 
| 20140 | 137 | by (blast intro!: wmapXH [THEN iffD2]) | 
| 138 | ||
| 58977 | 139 | lemma wmapE: "\<lbrakk>p : wmap(f,r); \<And>a b. \<lbrakk><f(a),f(b)> : r; p=<a,b>\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" | 
| 20140 | 140 | by (blast dest!: wmapXH [THEN iffD1]) | 
| 141 | ||
| 142 | lemma wmap_wf: | |
| 143 | assumes 1: "Wfd(r)" | |
| 144 | shows "Wfd(wmap(f,r))" | |
| 145 | apply (unfold Wfd_def) | |
| 146 | apply clarify | |
| 58977 | 147 | apply (subgoal_tac "ALL b. ALL a. f (a) = b \<longrightarrow> a:P") | 
| 20140 | 148 | apply blast | 
| 149 | apply (rule 1 [THEN wfd_induct, THEN allI]) | |
| 150 | apply clarify | |
| 151 | apply (erule spec [THEN mp]) | |
| 152 | apply (safe elim!: wmapE) | |
| 153 | apply (erule spec [THEN mp, THEN spec, THEN mp]) | |
| 154 | apply assumption | |
| 155 | apply (rule refl) | |
| 156 | done | |
| 157 | ||
| 158 | ||
| 60770 | 159 | subsection \<open>Projections\<close> | 
| 20140 | 160 | |
| 58977 | 161 | lemma wfstI: "<xa,ya> : r \<Longrightarrow> <<xa,xb>,<ya,yb>> : wmap(fst,r)" | 
| 20140 | 162 | apply (rule wmapI) | 
| 163 | apply simp | |
| 164 | done | |
| 165 | ||
| 58977 | 166 | lemma wsndI: "<xb,yb> : r \<Longrightarrow> <<xa,xb>,<ya,yb>> : wmap(snd,r)" | 
| 20140 | 167 | apply (rule wmapI) | 
| 168 | apply simp | |
| 169 | done | |
| 170 | ||
| 58977 | 171 | lemma wthdI: "<xc,yc> : r \<Longrightarrow> <<xa,<xb,xc>>,<ya,<yb,yc>>> : wmap(thd,r)" | 
| 20140 | 172 | apply (rule wmapI) | 
| 173 | apply simp | |
| 174 | done | |
| 175 | ||
| 176 | ||
| 60770 | 177 | subsection \<open>Ground well-founded relations\<close> | 
| 20140 | 178 | |
| 58977 | 179 | lemma wfI: "\<lbrakk>Wfd(r); a : r\<rbrakk> \<Longrightarrow> a : wf(r)" | 
| 20140 | 180 | unfolding wf_def by blast | 
| 181 | ||
| 182 | lemma Empty_wf: "Wfd({})"
 | |
| 183 | unfolding Wfd_def by (blast elim: EmptyXH [THEN iffD1, THEN FalseE]) | |
| 184 | ||
| 185 | lemma wf_wf: "Wfd(wf(R))" | |
| 186 | unfolding wf_def | |
| 187 | apply (rule_tac Q = "Wfd(R)" in excluded_middle [THEN disjE]) | |
| 188 | apply simp_all | |
| 189 | apply (rule Empty_wf) | |
| 190 | done | |
| 191 | ||
| 58977 | 192 | lemma NatPRXH: "p : NatPR \<longleftrightarrow> (EX x:Nat. p=<x,succ(x)>)" | 
| 20140 | 193 | unfolding NatPR_def by blast | 
| 194 | ||
| 58977 | 195 | lemma ListPRXH: "p : ListPR(A) \<longleftrightarrow> (EX h:A. EX t:List(A).p=<t,h$t>)" | 
| 20140 | 196 | unfolding ListPR_def by blast | 
| 197 | ||
| 58977 | 198 | lemma NatPRI: "x : Nat \<Longrightarrow> <x,succ(x)> : NatPR" | 
| 20140 | 199 | by (auto simp: NatPRXH) | 
| 200 | ||
| 58977 | 201 | lemma ListPRI: "\<lbrakk>t : List(A); h : A\<rbrakk> \<Longrightarrow> <t,h $ t> : ListPR(A)" | 
| 20140 | 202 | by (auto simp: ListPRXH) | 
| 203 | ||
| 204 | lemma NatPR_wf: "Wfd(NatPR)" | |
| 205 | apply (unfold Wfd_def) | |
| 206 | apply clarify | |
| 58977 | 207 | apply (wfd_strengthen "\<lambda>x. x:Nat") | 
| 47966 | 208 | apply (fastforce iff: NatPRXH) | 
| 20140 | 209 | apply (erule Nat_ind) | 
| 47966 | 210 | apply (fastforce iff: NatPRXH)+ | 
| 20140 | 211 | done | 
| 212 | ||
| 213 | lemma ListPR_wf: "Wfd(ListPR(A))" | |
| 214 | apply (unfold Wfd_def) | |
| 215 | apply clarify | |
| 58977 | 216 | apply (wfd_strengthen "\<lambda>x. x:List (A)") | 
| 47966 | 217 | apply (fastforce iff: ListPRXH) | 
| 20140 | 218 | apply (erule List_ind) | 
| 47966 | 219 | apply (fastforce iff: ListPRXH)+ | 
| 20140 | 220 | done | 
| 221 | ||
| 222 | ||
| 60770 | 223 | subsection \<open>General Recursive Functions\<close> | 
| 20140 | 224 | |
| 225 | lemma letrecT: | |
| 226 | assumes 1: "a : A" | |
| 58977 | 227 |     and 2: "\<And>p g. \<lbrakk>p:A; ALL x:{x: A. <x,p>:wf(R)}. g(x) : D(x)\<rbrakk> \<Longrightarrow> h(p,g) : D(p)"
 | 
| 20140 | 228 | shows "letrec g x be h(x,g) in g(a) : D(a)" | 
| 229 | apply (rule 1 [THEN rev_mp]) | |
| 230 | apply (rule wf_wf [THEN wfd_induct]) | |
| 231 | apply (subst letrecB) | |
| 232 | apply (rule impI) | |
| 233 | apply (erule 2) | |
| 234 | apply blast | |
| 235 | done | |
| 236 | ||
| 237 | lemma SPLITB: "SPLIT(<a,b>,B) = B(a,b)" | |
| 238 | unfolding SPLIT_def | |
| 239 | apply (rule set_ext) | |
| 240 | apply blast | |
| 241 | done | |
| 17456 | 242 | |
| 20140 | 243 | lemma letrec2T: | 
| 244 | assumes "a : A" | |
| 245 | and "b : B" | |
| 58977 | 246 | and "\<And>p q g. \<lbrakk>p:A; q:B; | 
| 247 |               ALL x:A. ALL y:{y: B. <<x,y>,<p,q>>:wf(R)}. g(x,y) : D(x,y)\<rbrakk> \<Longrightarrow> 
 | |
| 20140 | 248 | h(p,q,g) : D(p,q)" | 
| 249 | shows "letrec g x y be h(x,y,g) in g(a,b) : D(a,b)" | |
| 250 | apply (unfold letrec2_def) | |
| 251 | apply (rule SPLITB [THEN subst]) | |
| 41526 | 252 | apply (assumption | rule letrecT pairT splitT assms)+ | 
| 20140 | 253 | apply (subst SPLITB) | 
| 41526 | 254 | apply (assumption | rule ballI SubtypeI assms)+ | 
| 20140 | 255 | apply (rule SPLITB [THEN subst]) | 
| 41526 | 256 | apply (assumption | rule letrecT SubtypeI pairT splitT assms | | 
| 20140 | 257 | erule bspec SubtypeE sym [THEN subst])+ | 
| 258 | done | |
| 259 | ||
| 58977 | 260 | lemma lem: "SPLIT(<a,<b,c>>,\<lambda>x xs. SPLIT(xs,\<lambda>y z. B(x,y,z))) = B(a,b,c)" | 
| 20140 | 261 | by (simp add: SPLITB) | 
| 262 | ||
| 263 | lemma letrec3T: | |
| 264 | assumes "a : A" | |
| 265 | and "b : B" | |
| 266 | and "c : C" | |
| 58977 | 267 | and "\<And>p q r g. \<lbrakk>p:A; q:B; r:C; | 
| 268 |        ALL x:A. ALL y:B. ALL z:{z:C. <<x,<y,z>>,<p,<q,r>>> : wf(R)}.
 | |
| 269 | g(x,y,z) : D(x,y,z) \<rbrakk> \<Longrightarrow> | |
| 20140 | 270 | h(p,q,r,g) : D(p,q,r)" | 
| 271 | shows "letrec g x y z be h(x,y,z,g) in g(a,b,c) : D(a,b,c)" | |
| 272 | apply (unfold letrec3_def) | |
| 273 | apply (rule lem [THEN subst]) | |
| 41526 | 274 | apply (assumption | rule letrecT pairT splitT assms)+ | 
| 20140 | 275 | apply (simp add: SPLITB) | 
| 41526 | 276 | apply (assumption | rule ballI SubtypeI assms)+ | 
| 20140 | 277 | apply (rule lem [THEN subst]) | 
| 41526 | 278 | apply (assumption | rule letrecT SubtypeI pairT splitT assms | | 
| 20140 | 279 | erule bspec SubtypeE sym [THEN subst])+ | 
| 280 | done | |
| 281 | ||
| 282 | lemmas letrecTs = letrecT letrec2T letrec3T | |
| 283 | ||
| 284 | ||
| 60770 | 285 | subsection \<open>Type Checking for Recursive Calls\<close> | 
| 20140 | 286 | |
| 287 | lemma rcallT: | |
| 58977 | 288 |   "\<lbrakk>ALL x:{x:A.<x,p>:wf(R)}.g(x):D(x);  
 | 
| 289 | g(a) : D(a) \<Longrightarrow> g(a) : E; a:A; <a,p>:wf(R)\<rbrakk> \<Longrightarrow> g(a) : E" | |
| 20140 | 290 | by blast | 
| 291 | ||
| 292 | lemma rcall2T: | |
| 58977 | 293 |   "\<lbrakk>ALL x:A. ALL y:{y:B.<<x,y>,<p,q>>:wf(R)}.g(x,y):D(x,y);
 | 
| 294 | g(a,b) : D(a,b) \<Longrightarrow> g(a,b) : E; a:A; b:B; <<a,b>,<p,q>>:wf(R)\<rbrakk> \<Longrightarrow> g(a,b) : E" | |
| 20140 | 295 | by blast | 
| 296 | ||
| 297 | lemma rcall3T: | |
| 58977 | 298 |   "\<lbrakk>ALL x:A. ALL y:B. ALL z:{z:C.<<x,<y,z>>,<p,<q,r>>>:wf(R)}. g(x,y,z):D(x,y,z);
 | 
| 299 | g(a,b,c) : D(a,b,c) \<Longrightarrow> g(a,b,c) : E; | |
| 300 | a:A; b:B; c:C; <<a,<b,c>>,<p,<q,r>>> : wf(R)\<rbrakk> \<Longrightarrow> g(a,b,c) : E" | |
| 20140 | 301 | by blast | 
| 302 | ||
| 303 | lemmas rcallTs = rcallT rcall2T rcall3T | |
| 304 | ||
| 305 | ||
| 60770 | 306 | subsection \<open>Instantiating an induction hypothesis with an equality assumption\<close> | 
| 20140 | 307 | |
| 308 | lemma hyprcallT: | |
| 309 | assumes 1: "g(a) = b" | |
| 310 |     and 2: "ALL x:{x:A.<x,p>:wf(R)}.g(x):D(x)"
 | |
| 58977 | 311 |     and 3: "ALL x:{x:A.<x,p>:wf(R)}.g(x):D(x) \<Longrightarrow> b=g(a) \<Longrightarrow> g(a) : D(a) \<Longrightarrow> P"
 | 
| 312 |     and 4: "ALL x:{x:A.<x,p>:wf(R)}.g(x):D(x) \<Longrightarrow> a:A"
 | |
| 313 |     and 5: "ALL x:{x:A.<x,p>:wf(R)}.g(x):D(x) \<Longrightarrow> <a,p>:wf(R)"
 | |
| 20140 | 314 | shows P | 
| 315 | apply (rule 3 [OF 2, OF 1 [symmetric]]) | |
| 316 | apply (rule rcallT [OF 2]) | |
| 317 | apply assumption | |
| 318 | apply (rule 4 [OF 2]) | |
| 319 | apply (rule 5 [OF 2]) | |
| 320 | done | |
| 321 | ||
| 322 | lemma hyprcall2T: | |
| 323 | assumes 1: "g(a,b) = c" | |
| 324 |     and 2: "ALL x:A. ALL y:{y:B.<<x,y>,<p,q>>:wf(R)}.g(x,y):D(x,y)"
 | |
| 58977 | 325 | and 3: "\<lbrakk>c = g(a,b); g(a,b) : D(a,b)\<rbrakk> \<Longrightarrow> P" | 
| 20140 | 326 | and 4: "a:A" | 
| 327 | and 5: "b:B" | |
| 328 | and 6: "<<a,b>,<p,q>>:wf(R)" | |
| 329 | shows P | |
| 330 | apply (rule 3) | |
| 331 | apply (rule 1 [symmetric]) | |
| 332 | apply (rule rcall2T) | |
| 23467 | 333 | apply (rule 2) | 
| 334 | apply assumption | |
| 335 | apply (rule 4) | |
| 336 | apply (rule 5) | |
| 337 | apply (rule 6) | |
| 20140 | 338 | done | 
| 339 | ||
| 340 | lemma hyprcall3T: | |
| 341 | assumes 1: "g(a,b,c) = d" | |
| 342 |     and 2: "ALL x:A. ALL y:B. ALL z:{z:C.<<x,<y,z>>,<p,<q,r>>>:wf(R)}.g(x,y,z):D(x,y,z)"
 | |
| 58977 | 343 | and 3: "\<lbrakk>d = g(a,b,c); g(a,b,c) : D(a,b,c)\<rbrakk> \<Longrightarrow> P" | 
| 20140 | 344 | and 4: "a:A" | 
| 345 | and 5: "b:B" | |
| 346 | and 6: "c:C" | |
| 347 | and 7: "<<a,<b,c>>,<p,<q,r>>> : wf(R)" | |
| 348 | shows P | |
| 349 | apply (rule 3) | |
| 350 | apply (rule 1 [symmetric]) | |
| 351 | apply (rule rcall3T) | |
| 23467 | 352 | apply (rule 2) | 
| 353 | apply assumption | |
| 354 | apply (rule 4) | |
| 355 | apply (rule 5) | |
| 356 | apply (rule 6) | |
| 357 | apply (rule 7) | |
| 20140 | 358 | done | 
| 359 | ||
| 360 | lemmas hyprcallTs = hyprcallT hyprcall2T hyprcall3T | |
| 361 | ||
| 362 | ||
| 60770 | 363 | subsection \<open>Rules to Remove Induction Hypotheses after Type Checking\<close> | 
| 20140 | 364 | |
| 58977 | 365 | lemma rmIH1: "\<lbrakk>ALL x:{x:A.<x,p>:wf(R)}.g(x):D(x); P\<rbrakk> \<Longrightarrow> P" .
 | 
| 20140 | 366 | |
| 58977 | 367 | lemma rmIH2: "\<lbrakk>ALL x:A. ALL y:{y:B.<<x,y>,<p,q>>:wf(R)}.g(x,y):D(x,y); P\<rbrakk> \<Longrightarrow> P" .
 | 
| 20140 | 368 | |
| 369 | lemma rmIH3: | |
| 58977 | 370 |  "\<lbrakk>ALL x:A. ALL y:B. ALL z:{z:C.<<x,<y,z>>,<p,<q,r>>>:wf(R)}.g(x,y,z):D(x,y,z); P\<rbrakk> \<Longrightarrow> P" .
 | 
| 20140 | 371 | |
| 372 | lemmas rmIHs = rmIH1 rmIH2 rmIH3 | |
| 373 | ||
| 374 | ||
| 60770 | 375 | subsection \<open>Lemmas for constructors and subtypes\<close> | 
| 20140 | 376 | |
| 377 | (* 0-ary constructors do not need additional rules as they are handled *) | |
| 378 | (* correctly by applying SubtypeI *) | |
| 379 | ||
| 380 | lemma Subtype_canTs: | |
| 58977 | 381 |   "\<And>a b A B P. a : {x:A. b:{y:B(a).P(<x,y>)}} \<Longrightarrow> <a,b> : {x:Sigma(A,B).P(x)}"
 | 
| 382 |   "\<And>a A B P. a : {x:A. P(inl(x))} \<Longrightarrow> inl(a) : {x:A+B. P(x)}"
 | |
| 383 |   "\<And>b A B P. b : {x:B. P(inr(x))} \<Longrightarrow> inr(b) : {x:A+B. P(x)}"
 | |
| 384 |   "\<And>a P. a : {x:Nat. P(succ(x))} \<Longrightarrow> succ(a) : {x:Nat. P(x)}"
 | |
| 385 |   "\<And>h t A P. h : {x:A. t : {y:List(A).P(x$y)}} \<Longrightarrow> h$t : {x:List(A).P(x)}"
 | |
| 20140 | 386 | by (assumption | rule SubtypeI canTs icanTs | erule SubtypeE)+ | 
| 387 | ||
| 58977 | 388 | lemma letT: "\<lbrakk>f(t):B; \<not>t=bot\<rbrakk> \<Longrightarrow> let x be t in f(x) : B" | 
| 20140 | 389 | apply (erule letB [THEN ssubst]) | 
| 390 | apply assumption | |
| 391 | done | |
| 392 | ||
| 58977 | 393 | lemma applyT2: "\<lbrakk>a:A; f : Pi(A,B)\<rbrakk> \<Longrightarrow> f ` a : B(a)" | 
| 20140 | 394 | apply (erule applyT) | 
| 395 | apply assumption | |
| 396 | done | |
| 397 | ||
| 58977 | 398 | lemma rcall_lemma1: "\<lbrakk>a:A; a:A \<Longrightarrow> P(a)\<rbrakk> \<Longrightarrow> a : {x:A. P(x)}"
 | 
| 20140 | 399 | by blast | 
| 400 | ||
| 58977 | 401 | lemma rcall_lemma2: "\<lbrakk>a:{x:A. Q(x)}; \<lbrakk>a:A; Q(a)\<rbrakk> \<Longrightarrow> P(a)\<rbrakk> \<Longrightarrow> a : {x:A. P(x)}"
 | 
| 20140 | 402 | by blast | 
| 403 | ||
| 404 | lemmas rcall_lemmas = asm_rl rcall_lemma1 SubtypeD1 rcall_lemma2 | |
| 405 | ||
| 406 | ||
| 60770 | 407 | subsection \<open>Typechecking\<close> | 
| 20140 | 408 | |
| 60770 | 409 | ML \<open> | 
| 20140 | 410 | local | 
| 411 | ||
| 412 | val type_rls = | |
| 27221 
31328dc30196
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27208diff
changeset | 413 |   @{thms canTs} @ @{thms icanTs} @ @{thms applyT2} @ @{thms ncanTs} @ @{thms incanTs} @
 | 
| 
31328dc30196
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27208diff
changeset | 414 |   @{thms precTs} @ @{thms letrecTs} @ @{thms letT} @ @{thms Subtype_canTs};
 | 
| 20140 | 415 | |
| 74298 | 416 | fun bvars \<^Const_>\<open>Pure.all _ for \<open>Abs(s,_,t)\<close>\<close> l = bvars t (s::l) | 
| 20140 | 417 | | bvars _ l = l | 
| 418 | ||
| 74298 | 419 | fun get_bno l n \<^Const_>\<open>Pure.all _ for \<open>Abs(s,_,t)\<close>\<close> = get_bno (s::l) n t | 
| 420 | | get_bno l n \<^Const_>\<open>Trueprop for t\<close> = get_bno l n t | |
| 421 | | get_bno l n \<^Const_>\<open>Ball _ for _ \<open>Abs(s,_,t)\<close>\<close> = get_bno (s::l) (n+1) t | |
| 422 | | get_bno l n \<^Const_>\<open>mem _ for t _\<close> = get_bno l n t | |
| 20140 | 423 | | get_bno l n (t $ s) = get_bno l n t | 
| 424 | | get_bno l n (Bound m) = (m-length(l),n) | |
| 425 | ||
| 426 | (* Not a great way of identifying induction hypothesis! *) | |
| 59582 | 427 | fun could_IH x = Term.could_unify(x,hd (Thm.prems_of @{thm rcallT})) orelse
 | 
| 428 |                  Term.could_unify(x,hd (Thm.prems_of @{thm rcall2T})) orelse
 | |
| 429 |                  Term.could_unify(x,hd (Thm.prems_of @{thm rcall3T}))
 | |
| 20140 | 430 | |
| 431 | fun IHinst tac rls = SUBGOAL (fn (Bi,i) => | |
| 432 | let val bvs = bvars Bi [] | |
| 33317 | 433 | val ihs = filter could_IH (Logic.strip_assums_hyp Bi) | 
| 59755 | 434 | val rnames = map (fn x => | 
| 20140 | 435 | let val (a,b) = get_bno [] 0 x | 
| 42364 | 436 | in (nth bvs a, b) end) ihs | 
| 20140 | 437 | fun try_IHs [] = no_tac | 
| 59755 | 438 | | try_IHs ((x,y)::xs) = | 
| 439 |             tac [((("g", 0), Position.none), x)] (nth rls (y - 1)) i ORELSE (try_IHs xs)
 | |
| 20140 | 440 | in try_IHs rnames end) | 
| 441 | ||
| 442 | fun is_rigid_prog t = | |
| 58976 | 443 | (case (Logic.strip_assums_concl t) of | 
| 74375 | 444 | \<^Const_>\<open>Trueprop for \<^Const_>\<open>mem _ for a _\<close>\<close> => null (Term.add_vars a []) | 
| 58976 | 445 | | _ => false) | 
| 446 | ||
| 20140 | 447 | in | 
| 448 | ||
| 27221 
31328dc30196
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27208diff
changeset | 449 | fun rcall_tac ctxt i = | 
| 60754 | 450 | let fun tac ps rl i = Rule_Insts.res_inst_tac ctxt ps [] rl i THEN assume_tac ctxt i | 
| 27221 
31328dc30196
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27208diff
changeset | 451 |   in IHinst tac @{thms rcallTs} i end
 | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58977diff
changeset | 452 |   THEN eresolve_tac ctxt @{thms rcall_lemmas} i
 | 
| 20140 | 453 | |
| 58957 | 454 | fun raw_step_tac ctxt prems i = | 
| 59499 | 455 | assume_tac ctxt i ORELSE | 
| 456 | resolve_tac ctxt (prems @ type_rls) i ORELSE | |
| 58957 | 457 | rcall_tac ctxt i ORELSE | 
| 59499 | 458 |   ematch_tac ctxt @{thms SubtypeE} i ORELSE
 | 
| 459 |   match_tac ctxt @{thms SubtypeI} i
 | |
| 20140 | 460 | |
| 27221 
31328dc30196
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27208diff
changeset | 461 | fun tc_step_tac ctxt prems = SUBGOAL (fn (Bi,i) => | 
| 58976 | 462 | if is_rigid_prog Bi then raw_step_tac ctxt prems i else no_tac) | 
| 20140 | 463 | |
| 27221 
31328dc30196
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27208diff
changeset | 464 | fun typechk_tac ctxt rls i = SELECT_GOAL (REPEAT_FIRST (tc_step_tac ctxt rls)) i | 
| 20140 | 465 | |
| 466 | ||
| 467 | (*** Clean up Correctness Condictions ***) | |
| 468 | ||
| 27221 
31328dc30196
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27208diff
changeset | 469 | fun clean_ccs_tac ctxt = | 
| 60754 | 470 | let fun tac ps rl i = Rule_Insts.eres_inst_tac ctxt ps [] rl i THEN assume_tac ctxt i in | 
| 27221 
31328dc30196
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27208diff
changeset | 471 |     TRY (REPEAT_FIRST (IHinst tac @{thms hyprcallTs} ORELSE'
 | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58977diff
changeset | 472 |       eresolve_tac ctxt ([asm_rl, @{thm SubtypeE}] @ @{thms rmIHs}) ORELSE'
 | 
| 51798 | 473 | hyp_subst_tac ctxt)) | 
| 27221 
31328dc30196
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27208diff
changeset | 474 | end | 
| 20140 | 475 | |
| 27221 
31328dc30196
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27208diff
changeset | 476 | fun gen_ccs_tac ctxt rls i = | 
| 
31328dc30196
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27208diff
changeset | 477 | SELECT_GOAL (REPEAT_FIRST (tc_step_tac ctxt rls) THEN clean_ccs_tac ctxt) i | 
| 17456 | 478 | |
| 0 | 479 | end | 
| 60770 | 480 | \<close> | 
| 20140 | 481 | |
| 60770 | 482 | method_setup typechk = \<open> | 
| 58971 | 483 | Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (typechk_tac ctxt ths)) | 
| 60770 | 484 | \<close> | 
| 58971 | 485 | |
| 60770 | 486 | method_setup clean_ccs = \<open> | 
| 58971 | 487 | Scan.succeed (SIMPLE_METHOD o clean_ccs_tac) | 
| 60770 | 488 | \<close> | 
| 58971 | 489 | |
| 60770 | 490 | method_setup gen_ccs = \<open> | 
| 58971 | 491 | Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (gen_ccs_tac ctxt ths)) | 
| 60770 | 492 | \<close> | 
| 58971 | 493 | |
| 20140 | 494 | |
| 60770 | 495 | subsection \<open>Evaluation\<close> | 
| 20140 | 496 | |
| 57945 
cacb00a569e0
prefer 'named_theorems' over Named_Thms, with subtle change of semantics due to visual order vs. internal reverse order;
 wenzelm parents: 
57521diff
changeset | 497 | named_theorems eval "evaluation rules" | 
| 
cacb00a569e0
prefer 'named_theorems' over Named_Thms, with subtle change of semantics due to visual order vs. internal reverse order;
 wenzelm parents: 
57521diff
changeset | 498 | |
| 60770 | 499 | ML \<open> | 
| 32282 
853ef99c04cc
FOCUS_PREMS as full replacement for METAHYPS, where the conclusion may still contain schematic variables;
 wenzelm parents: 
32211diff
changeset | 500 | fun eval_tac ths = | 
| 57945 
cacb00a569e0
prefer 'named_theorems' over Named_Thms, with subtle change of semantics due to visual order vs. internal reverse order;
 wenzelm parents: 
57521diff
changeset | 501 |   Subgoal.FOCUS_PREMS (fn {context = ctxt, prems, ...} =>
 | 
| 69593 | 502 | let val eval_rules = Named_Theorems.get ctxt \<^named_theorems>\<open>eval\<close> | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58977diff
changeset | 503 | in DEPTH_SOLVE_1 (resolve_tac ctxt (ths @ prems @ rev eval_rules) 1) end) | 
| 60770 | 504 | \<close> | 
| 20140 | 505 | |
| 60770 | 506 | method_setup eval = \<open> | 
| 47432 | 507 | Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (CHANGED o eval_tac ths ctxt)) | 
| 60770 | 508 | \<close> | 
| 20140 | 509 | |
| 510 | ||
| 511 | lemmas eval_rls [eval] = trueV falseV pairV lamV caseVtrue caseVfalse caseVpair caseVlam | |
| 512 | ||
| 513 | lemma applyV [eval]: | |
| 61966 | 514 | assumes "f \<longlongrightarrow> lam x. b(x)" | 
| 515 | and "b(a) \<longlongrightarrow> c" | |
| 516 | shows "f ` a \<longlongrightarrow> c" | |
| 41526 | 517 | unfolding apply_def by (eval assms) | 
| 20140 | 518 | |
| 519 | lemma letV: | |
| 61966 | 520 | assumes 1: "t \<longlongrightarrow> a" | 
| 521 | and 2: "f(a) \<longlongrightarrow> c" | |
| 522 | shows "let x be t in f(x) \<longlongrightarrow> c" | |
| 20140 | 523 | apply (unfold let_def) | 
| 524 | apply (rule 1 [THEN canonical]) | |
| 60770 | 525 | apply (tactic \<open> | 
| 69593 | 526 |     REPEAT (DEPTH_SOLVE_1 (resolve_tac \<^context> (@{thms assms} @ @{thms eval_rls}) 1 ORELSE
 | 
| 527 |       eresolve_tac \<^context> @{thms substitute} 1))\<close>)
 | |
| 20140 | 528 | done | 
| 529 | ||
| 61966 | 530 | lemma fixV: "f(fix(f)) \<longlongrightarrow> c \<Longrightarrow> fix(f) \<longlongrightarrow> c" | 
| 20140 | 531 | apply (unfold fix_def) | 
| 532 | apply (rule applyV) | |
| 533 | apply (rule lamV) | |
| 534 | apply assumption | |
| 535 | done | |
| 536 | ||
| 537 | lemma letrecV: | |
| 61966 | 538 | "h(t,\<lambda>y. letrec g x be h(x,g) in g(y)) \<longlongrightarrow> c \<Longrightarrow> | 
| 539 | letrec g x be h(x,g) in g(t) \<longlongrightarrow> c" | |
| 20140 | 540 | apply (unfold letrec_def) | 
| 541 | apply (assumption | rule fixV applyV lamV)+ | |
| 542 | done | |
| 543 | ||
| 544 | lemmas [eval] = letV letrecV fixV | |
| 545 | ||
| 546 | lemma V_rls [eval]: | |
| 61966 | 547 | "true \<longlongrightarrow> true" | 
| 548 | "false \<longlongrightarrow> false" | |
| 549 | "\<And>b c t u. \<lbrakk>b\<longlongrightarrow>true; t\<longlongrightarrow>c\<rbrakk> \<Longrightarrow> if b then t else u \<longlongrightarrow> c" | |
| 550 | "\<And>b c t u. \<lbrakk>b\<longlongrightarrow>false; u\<longlongrightarrow>c\<rbrakk> \<Longrightarrow> if b then t else u \<longlongrightarrow> c" | |
| 551 | "\<And>a b. <a,b> \<longlongrightarrow> <a,b>" | |
| 552 | "\<And>a b c t h. \<lbrakk>t \<longlongrightarrow> <a,b>; h(a,b) \<longlongrightarrow> c\<rbrakk> \<Longrightarrow> split(t,h) \<longlongrightarrow> c" | |
| 553 | "zero \<longlongrightarrow> zero" | |
| 554 | "\<And>n. succ(n) \<longlongrightarrow> succ(n)" | |
| 555 | "\<And>c n t u. \<lbrakk>n \<longlongrightarrow> zero; t \<longlongrightarrow> c\<rbrakk> \<Longrightarrow> ncase(n,t,u) \<longlongrightarrow> c" | |
| 556 | "\<And>c n t u x. \<lbrakk>n \<longlongrightarrow> succ(x); u(x) \<longlongrightarrow> c\<rbrakk> \<Longrightarrow> ncase(n,t,u) \<longlongrightarrow> c" | |
| 557 | "\<And>c n t u. \<lbrakk>n \<longlongrightarrow> zero; t \<longlongrightarrow> c\<rbrakk> \<Longrightarrow> nrec(n,t,u) \<longlongrightarrow> c" | |
| 558 | "\<And>c n t u x. \<lbrakk>n\<longlongrightarrow>succ(x); u(x,nrec(x,t,u))\<longlongrightarrow>c\<rbrakk> \<Longrightarrow> nrec(n,t,u)\<longlongrightarrow>c" | |
| 559 | "[] \<longlongrightarrow> []" | |
| 560 | "\<And>h t. h$t \<longlongrightarrow> h$t" | |
| 561 | "\<And>c l t u. \<lbrakk>l \<longlongrightarrow> []; t \<longlongrightarrow> c\<rbrakk> \<Longrightarrow> lcase(l,t,u) \<longlongrightarrow> c" | |
| 562 | "\<And>c l t u x xs. \<lbrakk>l \<longlongrightarrow> x$xs; u(x,xs) \<longlongrightarrow> c\<rbrakk> \<Longrightarrow> lcase(l,t,u) \<longlongrightarrow> c" | |
| 563 | "\<And>c l t u. \<lbrakk>l \<longlongrightarrow> []; t \<longlongrightarrow> c\<rbrakk> \<Longrightarrow> lrec(l,t,u) \<longlongrightarrow> c" | |
| 564 | "\<And>c l t u x xs. \<lbrakk>l\<longlongrightarrow>x$xs; u(x,xs,lrec(xs,t,u))\<longlongrightarrow>c\<rbrakk> \<Longrightarrow> lrec(l,t,u)\<longlongrightarrow>c" | |
| 20140 | 565 | unfolding data_defs by eval+ | 
| 566 | ||
| 567 | ||
| 60770 | 568 | subsection \<open>Factorial\<close> | 
| 20140 | 569 | |
| 61337 | 570 | schematic_goal | 
| 58977 | 571 | "letrec f n be ncase(n,succ(zero),\<lambda>x. nrec(n,zero,\<lambda>y g. nrec(f(x),g,\<lambda>z h. succ(h)))) | 
| 61966 | 572 | in f(succ(succ(zero))) \<longlongrightarrow> ?a" | 
| 20140 | 573 | by eval | 
| 574 | ||
| 61337 | 575 | schematic_goal | 
| 58977 | 576 | "letrec f n be ncase(n,succ(zero),\<lambda>x. nrec(n,zero,\<lambda>y g. nrec(f(x),g,\<lambda>z h. succ(h)))) | 
| 61966 | 577 | in f(succ(succ(succ(zero)))) \<longlongrightarrow> ?a" | 
| 20140 | 578 | by eval | 
| 579 | ||
| 60770 | 580 | subsection \<open>Less Than Or Equal\<close> | 
| 20140 | 581 | |
| 61337 | 582 | schematic_goal | 
| 58977 | 583 | "letrec f p be split(p,\<lambda>m n. ncase(m,true,\<lambda>x. ncase(n,false,\<lambda>y. f(<x,y>)))) | 
| 61966 | 584 | in f(<succ(zero), succ(zero)>) \<longlongrightarrow> ?a" | 
| 20140 | 585 | by eval | 
| 586 | ||
| 61337 | 587 | schematic_goal | 
| 58977 | 588 | "letrec f p be split(p,\<lambda>m n. ncase(m,true,\<lambda>x. ncase(n,false,\<lambda>y. f(<x,y>)))) | 
| 61966 | 589 | in f(<succ(zero), succ(succ(succ(succ(zero))))>) \<longlongrightarrow> ?a" | 
| 20140 | 590 | by eval | 
| 591 | ||
| 61337 | 592 | schematic_goal | 
| 58977 | 593 | "letrec f p be split(p,\<lambda>m n. ncase(m,true,\<lambda>x. ncase(n,false,\<lambda>y. f(<x,y>)))) | 
| 61966 | 594 | in f(<succ(succ(succ(succ(succ(zero))))), succ(succ(succ(succ(zero))))>) \<longlongrightarrow> ?a" | 
| 20140 | 595 | by eval | 
| 596 | ||
| 597 | ||
| 60770 | 598 | subsection \<open>Reverse\<close> | 
| 20140 | 599 | |
| 61337 | 600 | schematic_goal | 
| 58977 | 601 | "letrec id l be lcase(l,[],\<lambda>x xs. x$id(xs)) | 
| 61966 | 602 | in id(zero$succ(zero)$[]) \<longlongrightarrow> ?a" | 
| 20140 | 603 | by eval | 
| 604 | ||
| 61337 | 605 | schematic_goal | 
| 58977 | 606 | "letrec rev l be lcase(l,[],\<lambda>x xs. lrec(rev(xs),x$[],\<lambda>y ys g. y$g)) | 
| 61966 | 607 | in rev(zero$succ(zero)$(succ((lam x. x)`succ(zero)))$([])) \<longlongrightarrow> ?a" | 
| 20140 | 608 | by eval | 
| 609 | ||
| 610 | end |