| author | wenzelm | 
| Fri, 04 Jan 2002 19:23:28 +0100 | |
| changeset 12630 | 6f2951938b66 | 
| parent 12566 | fe20540bcf93 | 
| child 13867 | 1fdecd15437f | 
| permissions | -rw-r--r-- | 
| 10213 | 1 | (* Title: HOL/Wellfounded_Recursion.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow, with minor changes by Konrad Slind | |
| 4 | Copyright 1992 University of Cambridge/1995 TU Munich | |
| 5 | ||
| 6 | Wellfoundedness, induction, and recursion | |
| 7 | *) | |
| 8 | ||
| 9 | Goal "x = y ==> H x z = H y z"; | |
| 10 | by (Asm_simp_tac 1); | |
| 11 | val H_cong2 = (*freeze H!*) | |
| 12 | 	      read_instantiate [("H","H")] (result());
 | |
| 13 | ||
| 14 | val [prem] = Goalw [wf_def] | |
| 15 | "(!!P x. (ALL x. (ALL y. (y,x) : r --> P(y)) --> P(x)) ==> P(x)) ==> wf(r)"; | |
| 16 | by (Clarify_tac 1); | |
| 17 | by (rtac prem 1); | |
| 18 | by (assume_tac 1); | |
| 19 | qed "wfUNIVI"; | |
| 20 | ||
| 21 | (*Restriction to domain A. If r is well-founded over A then wf(r)*) | |
| 22 | val [prem1,prem2] = Goalw [wf_def] | |
| 23 | "[| r <= A <*> A; \ | |
| 24 | \ !!x P. [| ALL x. (ALL y. (y,x) : r --> P y) --> P x; x:A |] ==> P x |] \ | |
| 25 | \ ==> wf r"; | |
| 26 | by (cut_facts_tac [prem1] 1); | |
| 27 | by (blast_tac (claset() addIs [prem2]) 1); | |
| 28 | qed "wfI"; | |
| 29 | ||
| 30 | val major::prems = Goalw [wf_def] | |
| 31 | "[| wf(r); \ | |
| 32 | \ !!x.[| ALL y. (y,x): r --> P(y) |] ==> P(x) \ | |
| 33 | \ |] ==> P(a)"; | |
| 34 | by (rtac (major RS spec RS mp RS spec) 1); | |
| 35 | by (blast_tac (claset() addIs prems) 1); | |
| 36 | qed "wf_induct"; | |
| 37 | ||
| 38 | (*Perform induction on i, then prove the wf(r) subgoal using prems. *) | |
| 39 | fun wf_ind_tac a prems i = | |
| 40 |     EVERY [res_inst_tac [("a",a)] wf_induct i,
 | |
| 41 | rename_last_tac a ["1"] (i+1), | |
| 42 | ares_tac prems i]; | |
| 43 | ||
| 44 | Goal "wf(r) ==> ALL x. (a,x):r --> (x,a)~:r"; | |
| 45 | by (wf_ind_tac "a" [] 1); | |
| 46 | by (Blast_tac 1); | |
| 47 | qed_spec_mp "wf_not_sym"; | |
| 48 | ||
| 49 | (* [| wf r; ~Z ==> (a,x) : r; (x,a) ~: r ==> Z |] ==> Z *) | |
| 50 | bind_thm ("wf_asym", cla_make_elim wf_not_sym);
 | |
| 51 | ||
| 52 | Goal "wf(r) ==> (a,a) ~: r"; | |
| 53 | by (blast_tac (claset() addEs [wf_asym]) 1); | |
| 54 | qed "wf_not_refl"; | |
| 55 | ||
| 56 | (* [| wf r; (a,a) ~: r ==> PROP W |] ==> PROP W *) | |
| 57 | bind_thm ("wf_irrefl", make_elim wf_not_refl);
 | |
| 58 | ||
| 59 | (*transitive closure of a wf relation is wf! *) | |
| 60 | Goal "wf(r) ==> wf(r^+)"; | |
| 61 | by (stac wf_def 1); | |
| 62 | by (Clarify_tac 1); | |
| 63 | (*must retain the universal formula for later use!*) | |
| 64 | by (rtac allE 1 THEN assume_tac 1); | |
| 65 | by (etac mp 1); | |
| 66 | by (eres_inst_tac [("a","x")] wf_induct 1);
 | |
| 67 | by (blast_tac (claset() addEs [tranclE]) 1); | |
| 68 | qed "wf_trancl"; | |
| 69 | ||
| 70 | Goal "wf (r^-1) ==> wf ((r^+)^-1)"; | |
| 71 | by (stac (trancl_converse RS sym) 1); | |
| 72 | by (etac wf_trancl 1); | |
| 73 | qed "wf_converse_trancl"; | |
| 74 | ||
| 75 | ||
| 76 | (*---------------------------------------------------------------------------- | |
| 77 | * Minimal-element characterization of well-foundedness | |
| 78 | *---------------------------------------------------------------------------*) | |
| 79 | ||
| 80 | Goalw [wf_def] "wf r ==> x:Q --> (EX z:Q. ALL y. (y,z):r --> y~:Q)"; | |
| 81 | by (dtac spec 1); | |
| 82 | by (etac (mp RS spec) 1); | |
| 83 | by (Blast_tac 1); | |
| 84 | val lemma1 = result(); | |
| 85 | ||
| 86 | Goalw [wf_def] "(ALL Q x. x:Q --> (EX z:Q. ALL y. (y,z):r --> y~:Q)) ==> wf r"; | |
| 87 | by (Clarify_tac 1); | |
| 88 | by (dres_inst_tac [("x", "{x. ~ P x}")] spec 1);
 | |
| 89 | by (Blast_tac 1); | |
| 90 | val lemma2 = result(); | |
| 91 | ||
| 92 | Goal "wf r = (ALL Q x. x:Q --> (EX z:Q. ALL y. (y,z):r --> y~:Q))"; | |
| 93 | by (blast_tac (claset() addSIs [lemma1, lemma2]) 1); | |
| 94 | qed "wf_eq_minimal"; | |
| 95 | ||
| 96 | (*--------------------------------------------------------------------------- | |
| 97 | * Wellfoundedness of subsets | |
| 98 | *---------------------------------------------------------------------------*) | |
| 99 | ||
| 100 | Goal "[| wf(r); p<=r |] ==> wf(p)"; | |
| 101 | by (full_simp_tac (simpset() addsimps [wf_eq_minimal]) 1); | |
| 102 | by (Fast_tac 1); | |
| 103 | qed "wf_subset"; | |
| 104 | ||
| 105 | (*--------------------------------------------------------------------------- | |
| 106 | * Wellfoundedness of the empty relation. | |
| 107 | *---------------------------------------------------------------------------*) | |
| 108 | ||
| 109 | Goal "wf({})";
 | |
| 110 | by (simp_tac (simpset() addsimps [wf_def]) 1); | |
| 111 | qed "wf_empty"; | |
| 112 | AddIffs [wf_empty]; | |
| 113 | ||
| 114 | (*--------------------------------------------------------------------------- | |
| 115 | * Wellfoundedness of `insert' | |
| 116 | *---------------------------------------------------------------------------*) | |
| 117 | ||
| 118 | Goal "wf(insert (y,x) r) = (wf(r) & (x,y) ~: r^*)"; | |
| 119 | by (rtac iffI 1); | |
| 120 | by (blast_tac (claset() addEs [wf_trancl RS wf_irrefl] | |
| 121 | addIs [rtrancl_into_trancl1,wf_subset,impOfSubs rtrancl_mono]) 1); | |
| 122 | by (asm_full_simp_tac (simpset() addsimps [wf_eq_minimal]) 1); | |
| 123 | by Safe_tac; | |
| 124 | by (EVERY1[rtac allE, assume_tac, etac impE, Blast_tac]); | |
| 125 | by (etac bexE 1); | |
| 11141 | 126 | by (rename_tac "a" 1 THEN case_tac "a = x" 1); | 
| 10213 | 127 |  by (res_inst_tac [("x","a")]bexI 2);
 | 
| 128 | by (assume_tac 3); | |
| 129 | by (Blast_tac 2); | |
| 130 | by (case_tac "y:Q" 1); | |
| 131 | by (Blast_tac 2); | |
| 132 | by (res_inst_tac [("x","{z. z:Q & (z,y) : r^*}")] allE 1);
 | |
| 133 | by (assume_tac 1); | |
| 134 | by (thin_tac "ALL Q. (EX x. x : Q) --> ?P Q" 1); (*essential for speed*) | |
| 135 | (*Blast_tac with new substOccur fails*) | |
| 12566 
fe20540bcf93
renamed rtrancl_into_rtrancl2 to converse_rtrancl_into_rtrancl
 nipkow parents: 
12486diff
changeset | 136 | by (best_tac (claset() addIs [converse_rtrancl_into_rtrancl]) 1); | 
| 10213 | 137 | qed "wf_insert"; | 
| 138 | AddIffs [wf_insert]; | |
| 139 | ||
| 140 | (*--------------------------------------------------------------------------- | |
| 141 | * Wellfoundedness of `disjoint union' | |
| 142 | *---------------------------------------------------------------------------*) | |
| 143 | ||
| 144 | (*Intuition behind this proof for the case of binary union: | |
| 145 | ||
| 146 | Goal: find an (R u S)-min element of a nonempty subset A. | |
| 147 | by case distinction: | |
| 148 | 1. There is a step a -R-> b with a,b : A. | |
| 149 |      Pick an R-min element z of the (nonempty) set {a:A | EX b:A. a -R-> b}.
 | |
| 150 | By definition, there is z':A s.t. z -R-> z'. Because z is R-min in the | |
| 151 | subset, z' must be R-min in A. Because z' has an R-predecessor, it cannot | |
| 152 | have an S-successor and is thus S-min in A as well. | |
| 153 | 2. There is no such step. | |
| 154 | Pick an S-min element of A. In this case it must be an R-min | |
| 155 | element of A as well. | |
| 156 | ||
| 157 | *) | |
| 158 | ||
| 159 | Goal "[| ALL i:I. wf(r i); \ | |
| 160 | \        ALL i:I. ALL j:I. r i ~= r j --> Domain(r i) Int Range(r j) = {} & \
 | |
| 161 | \                                         Domain(r j) Int Range(r i) = {} \
 | |
| 162 | \ |] ==> wf(UN i:I. r i)"; | |
| 163 | by (asm_full_simp_tac (HOL_basic_ss addsimps [wf_eq_minimal]) 1); | |
| 164 | by (Clarify_tac 1); | |
| 11141 | 165 | by (rename_tac "A a" 1 THEN case_tac "EX i:I. EX a:A. EX b:A. (b,a) : r i" 1); | 
| 10213 | 166 | by (Asm_full_simp_tac 2); | 
| 167 | by (Best_tac 2); (*much faster than Blast_tac*) | |
| 168 | by (Clarify_tac 1); | |
| 169 | by (EVERY1[dtac bspec, assume_tac, | |
| 170 | 	   eres_inst_tac [("x","{a. a:A & (EX b:A. (b,a) : r i)}")] allE]);
 | |
| 171 | by (EVERY1[etac allE, etac impE]); | |
| 172 | by (ALLGOALS Blast_tac); | |
| 173 | qed "wf_UN"; | |
| 174 | ||
| 175 | Goalw [Union_def] | |
| 176 | "[| ALL r:R. wf r; \ | |
| 177 | \    ALL r:R. ALL s:R. r ~= s --> Domain r Int Range s = {} & \
 | |
| 178 | \                                 Domain s Int Range r = {} \
 | |
| 179 | \ |] ==> wf(Union R)"; | |
| 180 | by (blast_tac (claset() addIs [wf_UN]) 1); | |
| 181 | qed "wf_Union"; | |
| 182 | ||
| 183 | Goal "[| wf r; wf s; Domain r Int Range s = {}; Domain s Int Range r = {} \
 | |
| 184 | \ |] ==> wf(r Un s)"; | |
| 185 | by (rtac (simplify (simpset()) (read_instantiate[("R","{r,s}")]wf_Union)) 1);
 | |
| 186 | by (Blast_tac 1); | |
| 187 | by (Blast_tac 1); | |
| 188 | qed "wf_Un"; | |
| 189 | ||
| 190 | (*--------------------------------------------------------------------------- | |
| 191 | * Wellfoundedness of `image' | |
| 192 | *---------------------------------------------------------------------------*) | |
| 193 | ||
| 10832 | 194 | Goal "[| wf r; inj f |] ==> wf(prod_fun f f ` r)"; | 
| 10213 | 195 | by (asm_full_simp_tac (HOL_basic_ss addsimps [wf_eq_minimal]) 1); | 
| 196 | by (Clarify_tac 1); | |
| 197 | by (case_tac "EX p. f p : Q" 1); | |
| 198 | by (eres_inst_tac [("x","{p. f p : Q}")]allE 1);
 | |
| 199 | by (fast_tac (claset() addDs [injD]) 1); | |
| 200 | by (Blast_tac 1); | |
| 201 | qed "wf_prod_fun_image"; | |
| 202 | ||
| 203 | (*** acyclic ***) | |
| 204 | ||
| 205 | Goalw [acyclic_def] "ALL x. (x, x) ~: r^+ ==> acyclic r"; | |
| 206 | by (assume_tac 1); | |
| 207 | qed "acyclicI"; | |
| 208 | ||
| 209 | Goalw [acyclic_def] "wf r ==> acyclic r"; | |
| 210 | by (blast_tac (claset() addEs [wf_trancl RS wf_irrefl]) 1); | |
| 211 | qed "wf_acyclic"; | |
| 212 | ||
| 213 | Goalw [acyclic_def] "acyclic(insert (y,x) r) = (acyclic r & (x,y) ~: r^*)"; | |
| 214 | by (simp_tac (simpset() addsimps [trancl_insert]) 1); | |
| 215 | by (blast_tac (claset() addIs [rtrancl_trans]) 1); | |
| 216 | qed "acyclic_insert"; | |
| 217 | AddIffs [acyclic_insert]; | |
| 218 | ||
| 219 | Goalw [acyclic_def] "acyclic(r^-1) = acyclic r"; | |
| 220 | by (simp_tac (simpset() addsimps [trancl_converse]) 1); | |
| 221 | qed "acyclic_converse"; | |
| 222 | AddIffs [acyclic_converse]; | |
| 223 | ||
| 224 | Goalw [acyclic_def,antisym_def] "acyclic r ==> antisym(r^*)"; | |
| 12486 | 225 | by (blast_tac (claset() addEs [rtranclE] | 
| 10213 | 226 | addIs [rtrancl_into_trancl1,rtrancl_trancl_trancl]) 1); | 
| 227 | qed "acyclic_impl_antisym_rtrancl"; | |
| 228 | ||
| 229 | (* Other direction: | |
| 230 | acyclic = no loops | |
| 231 | antisym = only self loops | |
| 232 | Goalw [acyclic_def,antisym_def] "antisym(r^* ) ==> acyclic(r - Id)"; | |
| 233 | ==> "antisym(r^* ) = acyclic(r - Id)"; | |
| 234 | *) | |
| 235 | ||
| 236 | Goalw [acyclic_def] "[| acyclic s; r <= s |] ==> acyclic r"; | |
| 237 | by (blast_tac (claset() addIs [trancl_mono]) 1); | |
| 238 | qed "acyclic_subset"; | |
| 239 | ||
| 240 | (** cut **) | |
| 241 | ||
| 242 | (*This rewrite rule works upon formulae; thus it requires explicit use of | |
| 243 | H_cong to expose the equality*) | |
| 244 | Goalw [cut_def] "(cut f r x = cut g r x) = (ALL y. (y,x):r --> f(y)=g(y))"; | |
| 245 | by (simp_tac (HOL_ss addsimps [expand_fun_eq]) 1); | |
| 246 | qed "cuts_eq"; | |
| 247 | ||
| 248 | Goalw [cut_def] "(x,a):r ==> (cut f r a)(x) = f(x)"; | |
| 249 | by (asm_simp_tac HOL_ss 1); | |
| 250 | qed "cut_apply"; | |
| 251 | ||
| 11328 | 252 | (*** Inductive characterization of wfrec combinator; for details see: ***) | 
| 253 | (*** John Harrison, "Inductive definitions: automation and application" ***) | |
| 10213 | 254 | |
| 11328 | 255 | Goalw [adm_wf_def] | 
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11328diff
changeset | 256 | "[| adm_wf R F; wf R |] ==> EX! y. (x, y) : wfrec_rel R F"; | 
| 11328 | 257 | by (wf_ind_tac "x" [] 1); | 
| 258 | by (rtac ex1I 1); | |
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11328diff
changeset | 259 | by (res_inst_tac [("g","%x. THE y. (x, y) : wfrec_rel R F")] wfrec_rel.wfrecI 1);
 | 
| 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11328diff
changeset | 260 | by (fast_tac (claset() addSDs [theI']) 1); | 
| 11328 | 261 | by (etac wfrec_rel.elim 1); | 
| 262 | by (Asm_full_simp_tac 1); | |
| 263 | byev [etac allE 1, etac allE 1, etac allE 1, etac mp 1]; | |
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11328diff
changeset | 264 | by (fast_tac (claset() addIs [the_equality RS sym]) 1); | 
| 11328 | 265 | qed "wfrec_unique"; | 
| 10213 | 266 | |
| 11328 | 267 | Goalw [adm_wf_def] "adm_wf R (%f x. F (cut f R x) x)"; | 
| 268 | by (strip_tac 1); | |
| 269 | by (rtac (cuts_eq RS iffD2 RS subst) 1); | |
| 270 | by (atac 1); | |
| 271 | by (rtac refl 1); | |
| 272 | qed "adm_lemma"; | |
| 10213 | 273 | |
| 274 | Goalw [wfrec_def] | |
| 275 | "wf(r) ==> wfrec r H a = H (cut (wfrec r H) r a) a"; | |
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11328diff
changeset | 276 | by (rtac (adm_lemma RS wfrec_unique RS the1_equality) 1); | 
| 11328 | 277 | by (atac 1); | 
| 278 | by (rtac wfrec_rel.wfrecI 1); | |
| 279 | by (strip_tac 1); | |
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11328diff
changeset | 280 | by (etac (adm_lemma RS wfrec_unique RS theI') 1); | 
| 10213 | 281 | qed "wfrec"; | 
| 282 | ||
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11328diff
changeset | 283 | |
| 10213 | 284 | (*--------------------------------------------------------------------------- | 
| 285 | * This form avoids giant explosions in proofs. NOTE USE OF == | |
| 286 | *---------------------------------------------------------------------------*) | |
| 287 | Goal "[| f==wfrec r H; wf(r) |] ==> f(a) = H (cut f r a) a"; | |
| 288 | by Auto_tac; | |
| 289 | by (blast_tac (claset() addIs [wfrec]) 1); | |
| 290 | qed "def_wfrec"; | |
| 291 | ||
| 292 | ||
| 293 | (**** TFL variants ****) | |
| 294 | ||
| 295 | Goal "ALL R. wf R --> \ | |
| 296 | \ (ALL P. (ALL x. (ALL y. (y,x):R --> P y) --> P x) --> (ALL x. P x))"; | |
| 297 | by (Clarify_tac 1); | |
| 298 | by (res_inst_tac [("r","R"),("P","P"), ("a","x")] wf_induct 1);
 | |
| 299 | by (assume_tac 1); | |
| 300 | by (Blast_tac 1); | |
| 301 | qed"tfl_wf_induct"; | |
| 302 | ||
| 303 | Goal "ALL f R. (x,a):R --> (cut f R a)(x) = f(x)"; | |
| 304 | by (Clarify_tac 1); | |
| 305 | by (rtac cut_apply 1); | |
| 306 | by (assume_tac 1); | |
| 307 | qed"tfl_cut_apply"; | |
| 308 | ||
| 309 | Goal "ALL M R f. (f=wfrec R M) --> wf R --> (ALL x. f x = M (cut f R x) x)"; | |
| 310 | by (Clarify_tac 1); | |
| 311 | by (etac wfrec 1); | |
| 312 | qed "tfl_wfrec"; | |
| 11141 | 313 | |
| 314 | (*LEAST and wellorderings*) | |
| 315 | (* ### see also wf_linord_ex_has_least and its consequences in Wellfounded_Relations.ML *) | |
| 316 | ||
| 317 | Goal "P (k::'a::wellorder) --> P (LEAST x. P(x)) & (LEAST x. P(x)) <= k"; | |
| 318 | by (res_inst_tac [("a","k")] (wf RS wf_induct) 1);
 | |
| 319 | by (rtac impI 1); | |
| 320 | by (rtac classical 1); | |
| 321 | by (res_inst_tac [("s","x")] (Least_equality RS ssubst) 1);
 | |
| 322 | by Auto_tac; | |
| 323 | by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym])); | |
| 324 | by (blast_tac (claset() addIs [order_less_trans]) 1); | |
| 325 | bind_thm("wellorder_LeastI",   result() RS mp RS conjunct1);
 | |
| 326 | bind_thm("wellorder_Least_le", result() RS mp RS conjunct2);
 | |
| 327 | ||
| 328 | Goal "[| k < (LEAST x. P x) |] ==> ~P (k::'a::wellorder)"; | |
| 329 | by (full_simp_tac (simpset() addsimps [linorder_not_le RS sym]) 1); | |
| 330 | by (etac contrapos_nn 1); | |
| 331 | by (etac wellorder_Least_le 1); | |
| 332 | qed "wellorder_not_less_Least"; | |
| 333 |