| author | wenzelm | 
| Wed, 10 Mar 1999 10:55:12 +0100 | |
| changeset 6340 | 7d5cbd5819a0 | 
| parent 5579 | 32f99ca617b7 | 
| child 6433 | 228237ec56e5 | 
| permissions | -rw-r--r-- | 
| 1475 | 1 | (* Title: HOL/wf.ML | 
| 923 | 2 | ID: $Id$ | 
| 1475 | 3 | Author: Tobias Nipkow, with minor changes by Konrad Slind | 
| 4 | Copyright 1992 University of Cambridge/1995 TU Munich | |
| 923 | 5 | |
| 3198 | 6 | Wellfoundedness, induction, and recursion | 
| 923 | 7 | *) | 
| 8 | ||
| 950 | 9 | val H_cong = read_instantiate [("f","H")] (standard(refl RS cong RS cong));
 | 
| 923 | 10 | val H_cong1 = refl RS H_cong; | 
| 11 | ||
| 5579 | 12 | val [prem] = Goalw [wf_def] | 
| 13 | "[| !!P x. [| !x. (!y. (y,x) : r --> P(y)) --> P(x) |] ==> P(x) |] ==> wf(r)"; | |
| 14 | by (Clarify_tac 1); | |
| 15 | by (rtac prem 1); | |
| 16 | by (assume_tac 1); | |
| 17 | qed "wfUNIVI"; | |
| 18 | ||
| 923 | 19 | (*Restriction to domain A. If r is well-founded over A then wf(r)*) | 
| 5316 | 20 | val [prem1,prem2] = Goalw [wf_def] | 
| 1642 | 21 | "[| r <= A Times A; \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
950diff
changeset | 22 | \ !!x P. [| ! x. (! y. (y,x) : r --> P(y)) --> P(x); x:A |] ==> P(x) |] \ | 
| 923 | 23 | \ ==> wf(r)"; | 
| 3708 | 24 | by (Clarify_tac 1); | 
| 923 | 25 | by (rtac allE 1); | 
| 26 | by (assume_tac 1); | |
| 4089 | 27 | by (best_tac (claset() addSEs [prem1 RS subsetD RS SigmaE2] addIs [prem2]) 1); | 
| 923 | 28 | qed "wfI"; | 
| 29 | ||
| 5316 | 30 | val major::prems = Goalw [wf_def] | 
| 923 | 31 | "[| wf(r); \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
950diff
changeset | 32 | \ !!x.[| ! y. (y,x): r --> P(y) |] ==> P(x) \ | 
| 923 | 33 | \ |] ==> P(a)"; | 
| 34 | by (rtac (major RS spec RS mp RS spec) 1); | |
| 4089 | 35 | by (blast_tac (claset() addIs prems) 1); | 
| 923 | 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,
 | |
| 1465 | 41 | rename_last_tac a ["1"] (i+1), | 
| 42 | ares_tac prems i]; | |
| 923 | 43 | |
| 5452 | 44 | Goal "wf(r) ==> ! x. (a,x):r --> (x,a)~:r"; | 
| 5316 | 45 | by (wf_ind_tac "a" [] 1); | 
| 2935 | 46 | by (Blast_tac 1); | 
| 5452 | 47 | qed_spec_mp "wf_not_sym"; | 
| 48 | ||
| 49 | (* [| wf(r); (a,x):r; ~P ==> (x,a):r |] ==> P *) | |
| 50 | bind_thm ("wf_asym", wf_not_sym RS swap);
 | |
| 923 | 51 | |
| 5316 | 52 | Goal "[| wf(r); (a,a): r |] ==> P"; | 
| 53 | by (blast_tac (claset() addEs [wf_asym]) 1); | |
| 1618 | 54 | qed "wf_irrefl"; | 
| 923 | 55 | |
| 1475 | 56 | (*transitive closure of a wf relation is wf! *) | 
| 5316 | 57 | Goal "wf(r) ==> wf(r^+)"; | 
| 58 | by (stac wf_def 1); | |
| 3708 | 59 | by (Clarify_tac 1); | 
| 923 | 60 | (*must retain the universal formula for later use!*) | 
| 61 | by (rtac allE 1 THEN assume_tac 1); | |
| 62 | by (etac mp 1); | |
| 5316 | 63 | by (eres_inst_tac [("a","x")] wf_induct 1);
 | 
| 923 | 64 | by (rtac (impI RS allI) 1); | 
| 65 | by (etac tranclE 1); | |
| 2935 | 66 | by (Blast_tac 1); | 
| 67 | by (Blast_tac 1); | |
| 923 | 68 | qed "wf_trancl"; | 
| 69 | ||
| 70 | ||
| 4762 | 71 | val wf_converse_trancl = prove_goal thy | 
| 72 | "!!X. wf (r^-1) ==> wf ((r^+)^-1)" (K [ | |
| 73 | stac (trancl_converse RS sym) 1, | |
| 74 | etac wf_trancl 1]); | |
| 75 | ||
| 3198 | 76 | (*---------------------------------------------------------------------------- | 
| 77 | * Minimal-element characterization of well-foundedness | |
| 78 | *---------------------------------------------------------------------------*) | |
| 79 | ||
| 5316 | 80 | Goalw [wf_def] "wf r ==> x:Q --> (? z:Q. ! y. (y,z):r --> y~:Q)"; | 
| 5318 | 81 | by (dtac spec 1); | 
| 5316 | 82 | by (etac (mp RS spec) 1); | 
| 3198 | 83 | by (Blast_tac 1); | 
| 84 | val lemma1 = result(); | |
| 85 | ||
| 5316 | 86 | Goalw [wf_def] "(! Q x. x:Q --> (? z:Q. ! y. (y,z):r --> y~:Q)) ==> wf r"; | 
| 3708 | 87 | by (Clarify_tac 1); | 
| 3198 | 88 | by (dres_inst_tac [("x", "{x. ~ P x}")] spec 1);
 | 
| 89 | by (Blast_tac 1); | |
| 90 | val lemma2 = result(); | |
| 91 | ||
| 5069 | 92 | Goal "wf r = (! Q x. x:Q --> (? z:Q. ! y. (y,z):r --> y~:Q))"; | 
| 4089 | 93 | by (blast_tac (claset() addSIs [lemma1, lemma2]) 1); | 
| 3198 | 94 | qed "wf_eq_minimal"; | 
| 95 | ||
| 3413 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 96 | (*--------------------------------------------------------------------------- | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 97 | * Wellfoundedness of subsets | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 98 | *---------------------------------------------------------------------------*) | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 99 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 100 | Goal "[| wf(r); p<=r |] ==> wf(p)"; | 
| 4089 | 101 | by (full_simp_tac (simpset() addsimps [wf_eq_minimal]) 1); | 
| 3413 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 102 | by (Fast_tac 1); | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 103 | qed "wf_subset"; | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 104 | |
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 105 | (*--------------------------------------------------------------------------- | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 106 | * Wellfoundedness of the empty relation. | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 107 | *---------------------------------------------------------------------------*) | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 108 | |
| 5069 | 109 | Goal "wf({})";
 | 
| 4089 | 110 | by (simp_tac (simpset() addsimps [wf_def]) 1); | 
| 3413 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 111 | qed "wf_empty"; | 
| 5281 | 112 | AddIffs [wf_empty]; | 
| 3413 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 113 | |
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 114 | (*--------------------------------------------------------------------------- | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 115 | * Wellfoundedness of `insert' | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 116 | *---------------------------------------------------------------------------*) | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 117 | |
| 5069 | 118 | Goal "wf(insert (y,x) r) = (wf(r) & (x,y) ~: r^*)"; | 
| 3457 | 119 | by (rtac iffI 1); | 
| 4350 
1983e4054fd8
updated for latest Blast_tac, which treats equality differently
 paulson parents: 
4153diff
changeset | 120 | by (blast_tac (claset() addEs [wf_trancl RS wf_irrefl] | 
| 
1983e4054fd8
updated for latest Blast_tac, which treats equality differently
 paulson parents: 
4153diff
changeset | 121 | addIs [rtrancl_into_trancl1,wf_subset,impOfSubs rtrancl_mono]) 1); | 
| 4089 | 122 | by (asm_full_simp_tac (simpset() addsimps [wf_eq_minimal]) 1); | 
| 4153 | 123 | by Safe_tac; | 
| 3457 | 124 | by (EVERY1[rtac allE, atac, etac impE, Blast_tac]); | 
| 125 | by (etac bexE 1); | |
| 126 | by (rename_tac "a" 1); | |
| 127 | by (case_tac "a = x" 1); | |
| 128 |  by (res_inst_tac [("x","a")]bexI 2);
 | |
| 129 | by (assume_tac 3); | |
| 130 | by (Blast_tac 2); | |
| 131 | by (case_tac "y:Q" 1); | |
| 132 | by (Blast_tac 2); | |
| 4059 | 133 | by (res_inst_tac [("x","{z. z:Q & (z,y) : r^*}")] allE 1);
 | 
| 3457 | 134 | by (assume_tac 1); | 
| 4059 | 135 | by (thin_tac "! Q. (? x. x : Q) --> ?P Q" 1); (*essential for speed*) | 
| 4350 
1983e4054fd8
updated for latest Blast_tac, which treats equality differently
 paulson parents: 
4153diff
changeset | 136 | (*Blast_tac with new substOccur fails*) | 
| 
1983e4054fd8
updated for latest Blast_tac, which treats equality differently
 paulson parents: 
4153diff
changeset | 137 | by (best_tac (claset() addIs [rtrancl_into_rtrancl2]) 1); | 
| 3413 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 138 | qed "wf_insert"; | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 139 | AddIffs [wf_insert]; | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 140 | |
| 5281 | 141 | (*--------------------------------------------------------------------------- | 
| 142 | * Wellfoundedness of `disjoint union' | |
| 143 | *---------------------------------------------------------------------------*) | |
| 144 | ||
| 5330 | 145 | (*Intuition behind this proof for the case of binary union: | 
| 146 | ||
| 147 | Goal: find an (R u S)-min element of a nonempty subset A. | |
| 148 | by case distinction: | |
| 149 | 1. There is a step a -R-> b with a,b : A. | |
| 150 |      Pick an R-min element z of the (nonempty) set {a:A | EX b:A. a -R-> b}.
 | |
| 151 | By definition, there is z':A s.t. z -R-> z'. Because z is R-min in the | |
| 152 | subset, z' must be R-min in A. Because z' has an R-predecessor, it cannot | |
| 153 | have an S-successor and is thus S-min in A as well. | |
| 154 | 2. There is no such step. | |
| 155 | Pick an S-min element of A. In this case it must be an R-min | |
| 156 | element of A as well. | |
| 157 | ||
| 158 | *) | |
| 159 | ||
| 5316 | 160 | Goal "[| !i:I. wf(r i); \ | 
| 161 | \        !i:I.!j:I. r i ~= r j --> Domain(r i) Int Range(r j) = {} & \
 | |
| 162 | \                                  Domain(r j) Int Range(r i) = {} \
 | |
| 163 | \ |] ==> wf(UN i:I. r i)"; | |
| 5318 | 164 | by (asm_full_simp_tac (HOL_basic_ss addsimps [wf_eq_minimal]) 1); | 
| 165 | by (Clarify_tac 1); | |
| 166 | by (rename_tac "A a" 1); | |
| 167 | by (case_tac "? i:I. ? a:A. ? b:A. (b,a) : r i" 1); | |
| 168 | by (Clarify_tac 1); | |
| 169 | by (EVERY1[dtac bspec, atac, | |
| 5281 | 170 |            eres_inst_tac[("x","{a. a:A & (? b:A. (b,a) : r i)}")]allE]);
 | 
| 5318 | 171 | by (EVERY1[etac allE,etac impE]); | 
| 172 | by (Blast_tac 1); | |
| 173 | by (Clarify_tac 1); | |
| 174 | by (rename_tac "z'" 1); | |
| 175 |  by (res_inst_tac [("x","z'")] bexI 1);
 | |
| 176 | by (assume_tac 2); | |
| 177 | by (Clarify_tac 1); | |
| 178 | by (rename_tac "j" 1); | |
| 179 | by (case_tac "r j = r i" 1); | |
| 180 | by (EVERY1[etac allE,etac impE,atac]); | |
| 181 | by (Asm_full_simp_tac 1); | |
| 182 | by (Blast_tac 1); | |
| 183 | by (blast_tac (claset() addEs [equalityE]) 1); | |
| 184 | by (Asm_full_simp_tac 1); | |
| 5521 | 185 | by (fast_tac (claset() delWrapper "bspec") 1); (*faster than Blast_tac*) | 
| 5281 | 186 | qed "wf_UN"; | 
| 187 | ||
| 188 | Goalw [Union_def] | |
| 189 | "[| !r:R. wf r; \ | |
| 190 | \    !r:R.!s:R. r ~= s --> Domain r Int Range s = {} & \
 | |
| 191 | \                          Domain s Int Range r = {} \
 | |
| 192 | \ |] ==> wf(Union R)"; | |
| 5318 | 193 | by (rtac wf_UN 1); | 
| 194 | by (Blast_tac 1); | |
| 195 | by (Blast_tac 1); | |
| 5281 | 196 | qed "wf_Union"; | 
| 197 | ||
| 5316 | 198 | Goal "[| wf r; wf s; Domain r Int Range s = {}; Domain s Int Range r = {} \
 | 
| 199 | \ |] ==> wf(r Un s)"; | |
| 5318 | 200 | by (rtac (simplify (simpset()) (read_instantiate[("R","{r,s}")]wf_Union)) 1);
 | 
| 201 | by (Blast_tac 1); | |
| 202 | by (Blast_tac 1); | |
| 5281 | 203 | qed "wf_Un"; | 
| 204 | ||
| 205 | (*--------------------------------------------------------------------------- | |
| 206 | * Wellfoundedness of `image' | |
| 207 | *---------------------------------------------------------------------------*) | |
| 208 | ||
| 209 | Goal "[| wf r; inj f |] ==> wf(prod_fun f f `` r)"; | |
| 5318 | 210 | by (asm_full_simp_tac (HOL_basic_ss addsimps [wf_eq_minimal]) 1); | 
| 211 | by (Clarify_tac 1); | |
| 212 | by (case_tac "? p. f p : Q" 1); | |
| 213 | by (eres_inst_tac [("x","{p. f p : Q}")]allE 1);
 | |
| 214 | by (fast_tac (claset() addDs [injD]) 1); | |
| 215 | by (Blast_tac 1); | |
| 5281 | 216 | qed "wf_prod_fun_image"; | 
| 217 | ||
| 3413 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 218 | (*** acyclic ***) | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 219 | |
| 4750 | 220 | val acyclicI = prove_goalw WF.thy [acyclic_def] | 
| 221 | "!!r. !x. (x, x) ~: r^+ ==> acyclic r" (K [atac 1]); | |
| 222 | ||
| 5069 | 223 | Goalw [acyclic_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 224 | "wf r ==> acyclic r"; | 
| 4089 | 225 | by (blast_tac (claset() addEs [wf_trancl RS wf_irrefl]) 1); | 
| 3413 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 226 | qed "wf_acyclic"; | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 227 | |
| 5452 | 228 | Goalw [acyclic_def] "acyclic(insert (y,x) r) = (acyclic r & (x,y) ~: r^*)"; | 
| 4089 | 229 | by (simp_tac (simpset() addsimps [trancl_insert]) 1); | 
| 5452 | 230 | by (blast_tac (claset() addIs [rtrancl_trans]) 1); | 
| 3413 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 231 | qed "acyclic_insert"; | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 232 | AddIffs [acyclic_insert]; | 
| 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
3320diff
changeset | 233 | |
| 5069 | 234 | Goalw [acyclic_def] "acyclic(r^-1) = acyclic r"; | 
| 4746 | 235 | by (simp_tac (simpset() addsimps [trancl_converse]) 1); | 
| 236 | qed "acyclic_converse"; | |
| 3198 | 237 | |
| 923 | 238 | (** cut **) | 
| 239 | ||
| 240 | (*This rewrite rule works upon formulae; thus it requires explicit use of | |
| 241 | H_cong to expose the equality*) | |
| 5069 | 242 | Goalw [cut_def] | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
950diff
changeset | 243 | "(cut f r x = cut g r x) = (!y. (y,x):r --> f(y)=g(y))"; | 
| 4686 | 244 | by (simp_tac (HOL_ss addsimps [expand_fun_eq]) 1); | 
| 1475 | 245 | qed "cuts_eq"; | 
| 923 | 246 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5132diff
changeset | 247 | Goalw [cut_def] "(x,a):r ==> (cut f r a)(x) = f(x)"; | 
| 1552 | 248 | by (asm_simp_tac HOL_ss 1); | 
| 923 | 249 | qed "cut_apply"; | 
| 250 | ||
| 251 | (*** is_recfun ***) | |
| 252 | ||
| 5069 | 253 | Goalw [is_recfun_def,cut_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 254 | "[| is_recfun r H a f; ~(b,a):r |] ==> f(b) = arbitrary"; | 
| 923 | 255 | by (etac ssubst 1); | 
| 1552 | 256 | by (asm_simp_tac HOL_ss 1); | 
| 923 | 257 | qed "is_recfun_undef"; | 
| 258 | ||
| 259 | (*** NOTE! some simplifications need a different finish_tac!! ***) | |
| 260 | fun indhyp_tac hyps = | |
| 261 | (cut_facts_tac hyps THEN' | |
| 262 | DEPTH_SOLVE_1 o (ares_tac [TrueI] ORELSE' | |
| 1465 | 263 | eresolve_tac [transD, mp, allE])); | 
| 2637 
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
 oheimb parents: 
2031diff
changeset | 264 | val wf_super_ss = HOL_ss addSolver indhyp_tac; | 
| 923 | 265 | |
| 5316 | 266 | Goalw [is_recfun_def,cut_def] | 
| 1475 | 267 | "[| wf(r); trans(r); is_recfun r H a f; is_recfun r H b g |] ==> \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
950diff
changeset | 268 | \ (x,a):r --> (x,b):r --> f(x)=g(x)"; | 
| 923 | 269 | by (etac wf_induct 1); | 
| 270 | by (REPEAT (rtac impI 1 ORELSE etac ssubst 1)); | |
| 271 | by (asm_simp_tac (wf_super_ss addcongs [if_cong]) 1); | |
| 1485 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 nipkow parents: 
1475diff
changeset | 272 | qed_spec_mp "is_recfun_equal"; | 
| 923 | 273 | |
| 274 | ||
| 275 | val prems as [wfr,transr,recfa,recgb,_] = goalw WF.thy [cut_def] | |
| 276 | "[| wf(r); trans(r); \ | |
| 1475 | 277 | \ is_recfun r H a f; is_recfun r H b g; (b,a):r |] ==> \ | 
| 923 | 278 | \ cut f r b = g"; | 
| 279 | val gundef = recgb RS is_recfun_undef | |
| 280 | and fisg = recgb RS (recfa RS (transr RS (wfr RS is_recfun_equal))); | |
| 281 | by (cut_facts_tac prems 1); | |
| 282 | by (rtac ext 1); | |
| 4686 | 283 | by (asm_simp_tac (wf_super_ss addsimps [gundef,fisg]) 1); | 
| 923 | 284 | qed "is_recfun_cut"; | 
| 285 | ||
| 286 | (*** Main Existence Lemma -- Basic Properties of the_recfun ***) | |
| 287 | ||
| 5316 | 288 | Goalw [the_recfun_def] | 
| 1475 | 289 | "is_recfun r H a f ==> is_recfun r H a (the_recfun r H a)"; | 
| 5316 | 290 | by (eres_inst_tac [("P", "is_recfun r H a")] selectI 1);
 | 
| 923 | 291 | qed "is_the_recfun"; | 
| 292 | ||
| 5316 | 293 | Goal "[| wf(r); trans(r) |] ==> is_recfun r H a (the_recfun r H a)"; | 
| 294 | by (wf_ind_tac "a" [] 1); | |
| 4821 | 295 | by (res_inst_tac [("f","cut (%y. H (the_recfun r H y) y) r a1")]
 | 
| 296 | is_the_recfun 1); | |
| 297 | by (rewtac is_recfun_def); | |
| 298 | by (stac cuts_eq 1); | |
| 299 | by (Clarify_tac 1); | |
| 300 | by (rtac (refl RSN (2,H_cong)) 1); | |
| 301 | by (subgoal_tac | |
| 1475 | 302 | "the_recfun r H y = cut(%x. H(cut(the_recfun r H y) r x) x) r y" 1); | 
| 4821 | 303 | by (etac allE 2); | 
| 304 | by (dtac impE 2); | |
| 305 | by (atac 2); | |
| 1475 | 306 | by (atac 3); | 
| 4821 | 307 | by (atac 2); | 
| 308 | by (etac ssubst 1); | |
| 309 | by (simp_tac (HOL_ss addsimps [cuts_eq]) 1); | |
| 310 | by (Clarify_tac 1); | |
| 311 | by (stac cut_apply 1); | |
| 5132 | 312 | by (fast_tac (claset() addDs [transD]) 1); | 
| 4821 | 313 | by (rtac (refl RSN (2,H_cong)) 1); | 
| 314 | by (fold_tac [is_recfun_def]); | |
| 315 | by (asm_simp_tac (wf_super_ss addsimps[is_recfun_cut]) 1); | |
| 923 | 316 | qed "unfold_the_recfun"; | 
| 317 | ||
| 1475 | 318 | val unwind1_the_recfun = rewrite_rule[is_recfun_def] unfold_the_recfun; | 
| 923 | 319 | |
| 1475 | 320 | (*--------------Old proof----------------------------------------------------- | 
| 5316 | 321 | val prems = Goal | 
| 1475 | 322 | "[| wf(r); trans(r) |] ==> is_recfun r H a (the_recfun r H a)"; | 
| 323 | by (cut_facts_tac prems 1); | |
| 324 | by (wf_ind_tac "a" prems 1); | |
| 325 | by (res_inst_tac [("f", "cut (%y. wftrec r H y) r a1")] is_the_recfun 1); 
 | |
| 326 | by (rewrite_goals_tac [is_recfun_def, wftrec_def]); | |
| 2031 | 327 | by (stac cuts_eq 1); | 
| 1475 | 328 | (*Applying the substitution: must keep the quantified assumption!!*) | 
| 3708 | 329 | by (EVERY1 [Clarify_tac, rtac H_cong1, rtac allE, atac, | 
| 1475 | 330 | etac (mp RS ssubst), atac]); | 
| 331 | by (fold_tac [is_recfun_def]); | |
| 332 | by (asm_simp_tac (wf_super_ss addsimps[cut_apply,is_recfun_cut,cuts_eq]) 1); | |
| 333 | qed "unfold_the_recfun"; | |
| 334 | ---------------------------------------------------------------------------*) | |
| 923 | 335 | |
| 336 | (** Removal of the premise trans(r) **) | |
| 1475 | 337 | val th = rewrite_rule[is_recfun_def] | 
| 338 | (trans_trancl RSN (2,(wf_trancl RS unfold_the_recfun))); | |
| 923 | 339 | |
| 5069 | 340 | Goalw [wfrec_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 341 | "wf(r) ==> wfrec r H a = H (cut (wfrec r H) r a) a"; | 
| 1475 | 342 | by (rtac H_cong 1); | 
| 343 | by (rtac refl 2); | |
| 344 | by (simp_tac (HOL_ss addsimps [cuts_eq]) 1); | |
| 345 | by (rtac allI 1); | |
| 346 | by (rtac impI 1); | |
| 347 | by (simp_tac(HOL_ss addsimps [wfrec_def]) 1); | |
| 348 | by (res_inst_tac [("a1","a")] (th RS ssubst) 1);
 | |
| 349 | by (atac 1); | |
| 350 | by (forward_tac[wf_trancl] 1); | |
| 351 | by (forward_tac[r_into_trancl] 1); | |
| 352 | by (asm_simp_tac (HOL_ss addsimps [cut_apply]) 1); | |
| 353 | by (rtac H_cong 1); (*expose the equality of cuts*) | |
| 354 | by (rtac refl 2); | |
| 355 | by (simp_tac (HOL_ss addsimps [cuts_eq, cut_apply, r_into_trancl]) 1); | |
| 3708 | 356 | by (Clarify_tac 1); | 
| 1485 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 nipkow parents: 
1475diff
changeset | 357 | by (res_inst_tac [("r","r^+")] is_recfun_equal 1);
 | 
| 1475 | 358 | by (atac 1); | 
| 359 | by (rtac trans_trancl 1); | |
| 360 | by (rtac unfold_the_recfun 1); | |
| 361 | by (atac 1); | |
| 362 | by (rtac trans_trancl 1); | |
| 363 | by (rtac unfold_the_recfun 1); | |
| 364 | by (atac 1); | |
| 365 | by (rtac trans_trancl 1); | |
| 366 | by (rtac transD 1); | |
| 367 | by (rtac trans_trancl 1); | |
| 4762 | 368 | by (forw_inst_tac [("p","(ya,y)")] r_into_trancl 1);
 | 
| 1475 | 369 | by (atac 1); | 
| 370 | by (atac 1); | |
| 4762 | 371 | by (forw_inst_tac [("p","(ya,y)")] r_into_trancl 1);
 | 
| 1475 | 372 | by (atac 1); | 
| 373 | qed "wfrec"; | |
| 374 | ||
| 375 | (*--------------Old proof----------------------------------------------------- | |
| 5069 | 376 | Goalw [wfrec_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 377 | "wf(r) ==> wfrec r H a = H (cut (wfrec r H) r a) a"; | 
| 923 | 378 | by (etac (wf_trancl RS wftrec RS ssubst) 1); | 
| 379 | by (rtac trans_trancl 1); | |
| 380 | by (rtac (refl RS H_cong) 1); (*expose the equality of cuts*) | |
| 1475 | 381 | by (simp_tac (HOL_ss addsimps [cuts_eq, cut_apply, r_into_trancl]) 1); | 
| 923 | 382 | qed "wfrec"; | 
| 1475 | 383 | ---------------------------------------------------------------------------*) | 
| 923 | 384 | |
| 1475 | 385 | (*--------------------------------------------------------------------------- | 
| 386 | * This form avoids giant explosions in proofs. NOTE USE OF == | |
| 387 | *---------------------------------------------------------------------------*) | |
| 5316 | 388 | val rew::prems = goal thy | 
| 1475 | 389 | "[| f==wfrec r H; wf(r) |] ==> f(a) = H (cut f r a) a"; | 
| 923 | 390 | by (rewtac rew); | 
| 391 | by (REPEAT (resolve_tac (prems@[wfrec]) 1)); | |
| 392 | qed "def_wfrec"; | |
| 1475 | 393 | |
| 3198 | 394 | |
| 395 | (**** TFL variants ****) | |
| 396 | ||
| 5278 | 397 | Goal "!R. wf R --> (!P. (!x. (!y. (y,x):R --> P y) --> P x) --> (!x. P x))"; | 
| 3708 | 398 | by (Clarify_tac 1); | 
| 3198 | 399 | by (res_inst_tac [("r","R"),("P","P"), ("a","x")] wf_induct 1);
 | 
| 400 | by (assume_tac 1); | |
| 401 | by (Blast_tac 1); | |
| 402 | qed"tfl_wf_induct"; | |
| 403 | ||
| 5069 | 404 | Goal "!f R. (x,a):R --> (cut f R a)(x) = f(x)"; | 
| 3708 | 405 | by (Clarify_tac 1); | 
| 3198 | 406 | by (rtac cut_apply 1); | 
| 407 | by (assume_tac 1); | |
| 408 | qed"tfl_cut_apply"; | |
| 409 | ||
| 5069 | 410 | Goal "!M R f. (f=wfrec R M) --> wf R --> (!x. f x = M (cut f R x) x)"; | 
| 3708 | 411 | by (Clarify_tac 1); | 
| 4153 | 412 | by (etac wfrec 1); | 
| 3198 | 413 | qed "tfl_wfrec"; |