author | wenzelm |
Thu, 03 Apr 2008 16:03:57 +0200 | |
changeset 26527 | c392354a1b79 |
parent 26235 | 96b804999ca7 |
permissions | -rw-r--r-- |
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
1 |
(* ID: $Id$ |
10213 | 2 |
Author: Tobias Nipkow |
3 |
Copyright 1992 University of Cambridge |
|
4 |
*) |
|
5 |
||
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
6 |
header {*Well-founded Recursion*} |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
7 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
8 |
theory Wellfounded_Recursion |
26072
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
9 |
imports Transitive_Closure Nat |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
10 |
uses ("Tools/function_package/size.ML") |
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
11 |
begin |
10213 | 12 |
|
23744
7c9e6e2fe249
Adapted to changes in infrastructure for converting between
berghofe
parents:
23389
diff
changeset
|
13 |
inductive |
22263 | 14 |
wfrec_rel :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b => bool" |
15 |
for R :: "('a * 'a) set" |
|
16 |
and F :: "('a => 'b) => 'a => 'b" |
|
17 |
where |
|
18 |
wfrecI: "ALL z. (z, x) : R --> wfrec_rel R F z (g z) ==> |
|
19 |
wfrec_rel R F x (F g x)" |
|
11328 | 20 |
|
10213 | 21 |
constdefs |
22 |
wf :: "('a * 'a)set => bool" |
|
23 |
"wf(r) == (!P. (!x. (!y. (y,x):r --> P(y)) --> P(x)) --> (!x. P(x)))" |
|
24 |
||
22263 | 25 |
wfP :: "('a => 'a => bool) => bool" |
23744
7c9e6e2fe249
Adapted to changes in infrastructure for converting between
berghofe
parents:
23389
diff
changeset
|
26 |
"wfP r == wf {(x, y). r x y}" |
22263 | 27 |
|
10213 | 28 |
acyclic :: "('a*'a)set => bool" |
29 |
"acyclic r == !x. (x,x) ~: r^+" |
|
30 |
||
31 |
cut :: "('a => 'b) => ('a * 'a)set => 'a => 'a => 'b" |
|
32 |
"cut f r x == (%y. if (y,x):r then f y else arbitrary)" |
|
33 |
||
11328 | 34 |
adm_wf :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => bool" |
35 |
"adm_wf R F == ALL f g x. |
|
36 |
(ALL z. (z, x) : R --> f z = g z) --> F f x = F g x" |
|
10213 | 37 |
|
11328 | 38 |
wfrec :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b" |
22845 | 39 |
[code func del]: "wfrec R F == %x. THE y. wfrec_rel R (%f x. F (cut f R x) x) x y" |
22263 | 40 |
|
41 |
abbreviation acyclicP :: "('a => 'a => bool) => bool" where |
|
23744
7c9e6e2fe249
Adapted to changes in infrastructure for converting between
berghofe
parents:
23389
diff
changeset
|
42 |
"acyclicP r == acyclic {(x, y). r x y}" |
10213 | 43 |
|
22390 | 44 |
class wellorder = linorder + |
25207 | 45 |
assumes wf: "wf {(x, y). x < y}" |
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
46 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
47 |
|
23744
7c9e6e2fe249
Adapted to changes in infrastructure for converting between
berghofe
parents:
23389
diff
changeset
|
48 |
lemma wfP_wf_eq [pred_set_conv]: "wfP (\<lambda>x y. (x, y) \<in> r) = wf r" |
22263 | 49 |
by (simp add: wfP_def) |
50 |
||
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
51 |
lemma wfUNIVI: |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
52 |
"(!!P x. (ALL x. (ALL y. (y,x) : r --> P(y)) --> P(x)) ==> P(x)) ==> wf(r)" |
26175 | 53 |
unfolding wf_def by blast |
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
54 |
|
22263 | 55 |
lemmas wfPUNIVI = wfUNIVI [to_pred] |
56 |
||
19766 | 57 |
text{*Restriction to domain @{term A} and range @{term B}. If @{term r} is |
58 |
well-founded over their intersection, then @{term "wf r"}*} |
|
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
59 |
lemma wfI: |
19766 | 60 |
"[| r \<subseteq> A <*> B; |
61 |
!!x P. [|\<forall>x. (\<forall>y. (y,x) : r --> P y) --> P x; x : A; x : B |] ==> P x |] |
|
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
62 |
==> wf r" |
26175 | 63 |
unfolding wf_def by blast |
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
64 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
65 |
lemma wf_induct: |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
66 |
"[| wf(r); |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
67 |
!!x.[| ALL y. (y,x): r --> P(y) |] ==> P(x) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
68 |
|] ==> P(a)" |
26175 | 69 |
unfolding wf_def by blast |
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
70 |
|
22263 | 71 |
lemmas wfP_induct = wf_induct [to_pred] |
72 |
||
18458 | 73 |
lemmas wf_induct_rule = wf_induct [rule_format, consumes 1, case_names less, induct set: wf] |
17042
da5cfaa258f7
moved wf_induct_rule from PreList.thy to Wellfounded_Recursion.thy
nipkow
parents:
15950
diff
changeset
|
74 |
|
25425
9191942c4ead
Removed some case_names and consumes attributes that are now no longer
berghofe
parents:
25207
diff
changeset
|
75 |
lemmas wfP_induct_rule = wf_induct_rule [to_pred, induct set: wfP] |
22263 | 76 |
|
26175 | 77 |
lemma wf_not_sym: "wf r ==> (a, x) : r ==> (x, a) ~: r" |
78 |
by (induct a arbitrary: x set: wf) blast |
|
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
79 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
80 |
(* [| wf r; ~Z ==> (a,x) : r; (x,a) ~: r ==> Z |] ==> Z *) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
81 |
lemmas wf_asym = wf_not_sym [elim_format] |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
82 |
|
26175 | 83 |
lemma wf_not_refl [simp]: "wf r ==> (a, a) ~: r" |
84 |
by (blast elim: wf_asym) |
|
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
85 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
86 |
(* [| wf r; (a,a) ~: r ==> PROP W |] ==> PROP W *) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
87 |
lemmas wf_irrefl = wf_not_refl [elim_format] |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
88 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
89 |
text{*transitive closure of a well-founded relation is well-founded! *} |
26175 | 90 |
lemma wf_trancl: |
91 |
assumes "wf r" |
|
92 |
shows "wf (r^+)" |
|
93 |
proof - |
|
94 |
{ |
|
95 |
fix P and x |
|
96 |
assume induct_step: "!!x. (!!y. (y, x) : r^+ ==> P y) ==> P x" |
|
97 |
have "P x" |
|
98 |
proof (rule induct_step) |
|
99 |
fix y assume "(y, x) : r^+" |
|
100 |
with `wf r` show "P y" |
|
101 |
proof (induct x arbitrary: y) |
|
102 |
case (less x) |
|
103 |
note hyp = `\<And>x' y'. (x', x) : r ==> (y', x') : r^+ ==> P y'` |
|
104 |
from `(y, x) : r^+` show "P y" |
|
105 |
proof cases |
|
106 |
case base |
|
107 |
show "P y" |
|
108 |
proof (rule induct_step) |
|
109 |
fix y' assume "(y', y) : r^+" |
|
110 |
with `(y, x) : r` show "P y'" by (rule hyp [of y y']) |
|
111 |
qed |
|
112 |
next |
|
113 |
case step |
|
114 |
then obtain x' where "(x', x) : r" and "(y, x') : r^+" by simp |
|
115 |
then show "P y" by (rule hyp [of x' y]) |
|
116 |
qed |
|
117 |
qed |
|
118 |
qed |
|
119 |
} then show ?thesis unfolding wf_def by blast |
|
120 |
qed |
|
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
121 |
|
22263 | 122 |
lemmas wfP_trancl = wf_trancl [to_pred] |
123 |
||
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
124 |
lemma wf_converse_trancl: "wf (r^-1) ==> wf ((r^+)^-1)" |
26175 | 125 |
apply (subst trancl_converse [symmetric]) |
126 |
apply (erule wf_trancl) |
|
127 |
done |
|
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
128 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
129 |
|
26175 | 130 |
subsubsection {* Other simple well-foundedness results *} |
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
131 |
|
19870 | 132 |
text{*Minimal-element characterization of well-foundedness*} |
133 |
lemma wf_eq_minimal: "wf r = (\<forall>Q x. x\<in>Q --> (\<exists>z\<in>Q. \<forall>y. (y,z)\<in>r --> y\<notin>Q))" |
|
134 |
proof (intro iffI strip) |
|
26175 | 135 |
fix Q :: "'a set" and x |
19870 | 136 |
assume "wf r" and "x \<in> Q" |
26175 | 137 |
then show "\<exists>z\<in>Q. \<forall>y. (y, z) \<in> r \<longrightarrow> y \<notin> Q" |
138 |
unfolding wf_def |
|
139 |
by (blast dest: spec [of _ "%x. x\<in>Q \<longrightarrow> (\<exists>z\<in>Q. \<forall>y. (y,z) \<in> r \<longrightarrow> y\<notin>Q)"]) |
|
19870 | 140 |
next |
22766 | 141 |
assume 1: "\<forall>Q x. x \<in> Q \<longrightarrow> (\<exists>z\<in>Q. \<forall>y. (y, z) \<in> r \<longrightarrow> y \<notin> Q)" |
142 |
show "wf r" |
|
143 |
proof (rule wfUNIVI) |
|
144 |
fix P :: "'a \<Rightarrow> bool" and x |
|
145 |
assume 2: "\<forall>x. (\<forall>y. (y, x) \<in> r \<longrightarrow> P y) \<longrightarrow> P x" |
|
146 |
let ?Q = "{x. \<not> P x}" |
|
26175 | 147 |
have "x \<in> ?Q \<longrightarrow> (\<exists>z \<in> ?Q. \<forall>y. (y, z) \<in> r \<longrightarrow> y \<notin> ?Q)" |
22766 | 148 |
by (rule 1 [THEN spec, THEN spec]) |
26175 | 149 |
then have "\<not> P x \<longrightarrow> (\<exists>z. \<not> P z \<and> (\<forall>y. (y, z) \<in> r \<longrightarrow> P y))" by simp |
22766 | 150 |
with 2 have "\<not> P x \<longrightarrow> (\<exists>z. \<not> P z \<and> P z)" by fast |
26175 | 151 |
then show "P x" by simp |
22766 | 152 |
qed |
19870 | 153 |
qed |
154 |
||
23186
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
155 |
lemma wfE_min: |
26175 | 156 |
assumes "wf R" "x \<in> Q" |
23186
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
157 |
obtains z where "z \<in> Q" "\<And>y. (y, z) \<in> R \<Longrightarrow> y \<notin> Q" |
26175 | 158 |
using assms unfolding wf_eq_minimal by blast |
23186
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
159 |
|
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
160 |
lemma wfI_min: |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
161 |
"(\<And>x Q. x \<in> Q \<Longrightarrow> \<exists>z\<in>Q. \<forall>y. (y, z) \<in> R \<longrightarrow> y \<notin> Q) |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
162 |
\<Longrightarrow> wf R" |
26175 | 163 |
unfolding wf_eq_minimal by blast |
23186
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
164 |
|
22263 | 165 |
lemmas wfP_eq_minimal = wf_eq_minimal [to_pred] |
166 |
||
26175 | 167 |
text {* Well-foundedness of subsets *} |
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
168 |
lemma wf_subset: "[| wf(r); p<=r |] ==> wf(p)" |
26175 | 169 |
apply (simp (no_asm_use) add: wf_eq_minimal) |
170 |
apply fast |
|
171 |
done |
|
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
172 |
|
22263 | 173 |
lemmas wfP_subset = wf_subset [to_pred] |
174 |
||
26175 | 175 |
text {* Well-foundedness of the empty relation *} |
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
176 |
lemma wf_empty [iff]: "wf({})" |
26175 | 177 |
by (simp add: wf_def) |
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
178 |
|
23744
7c9e6e2fe249
Adapted to changes in infrastructure for converting between
berghofe
parents:
23389
diff
changeset
|
179 |
lemmas wfP_empty [iff] = |
7c9e6e2fe249
Adapted to changes in infrastructure for converting between
berghofe
parents:
23389
diff
changeset
|
180 |
wf_empty [to_pred bot_empty_eq2, simplified bot_fun_eq bot_bool_eq] |
22263 | 181 |
|
19870 | 182 |
lemma wf_Int1: "wf r ==> wf (r Int r')" |
26175 | 183 |
apply (erule wf_subset) |
184 |
apply (rule Int_lower1) |
|
185 |
done |
|
19870 | 186 |
|
187 |
lemma wf_Int2: "wf r ==> wf (r' Int r)" |
|
26175 | 188 |
apply (erule wf_subset) |
189 |
apply (rule Int_lower2) |
|
190 |
done |
|
19870 | 191 |
|
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
192 |
text{*Well-foundedness of insert*} |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
193 |
lemma wf_insert [iff]: "wf(insert (y,x) r) = (wf(r) & (x,y) ~: r^*)" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
194 |
apply (rule iffI) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
195 |
apply (blast elim: wf_trancl [THEN wf_irrefl] |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
196 |
intro: rtrancl_into_trancl1 wf_subset |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
197 |
rtrancl_mono [THEN [2] rev_subsetD]) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
198 |
apply (simp add: wf_eq_minimal, safe) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
199 |
apply (rule allE, assumption, erule impE, blast) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
200 |
apply (erule bexE) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
201 |
apply (rename_tac "a", case_tac "a = x") |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
202 |
prefer 2 |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
203 |
apply blast |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
204 |
apply (case_tac "y:Q") |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
205 |
prefer 2 apply blast |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
206 |
apply (rule_tac x = "{z. z:Q & (z,y) : r^*}" in allE) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
207 |
apply assumption |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
208 |
apply (erule_tac V = "ALL Q. (EX x. x : Q) --> ?P Q" in thin_rl) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
209 |
--{*essential for speed*} |
15343 | 210 |
txt{*Blast with new substOccur fails*} |
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
211 |
apply (fast intro: converse_rtrancl_into_rtrancl) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
212 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
213 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
214 |
text{*Well-foundedness of image*} |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
215 |
lemma wf_prod_fun_image: "[| wf r; inj f |] ==> wf(prod_fun f f ` r)" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
216 |
apply (simp only: wf_eq_minimal, clarify) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
217 |
apply (case_tac "EX p. f p : Q") |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
218 |
apply (erule_tac x = "{p. f p : Q}" in allE) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
219 |
apply (fast dest: inj_onD, blast) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
220 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
221 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
222 |
|
26175 | 223 |
subsubsection {* Well-Foundedness Results for Unions *} |
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
224 |
|
26044
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
225 |
lemma wf_union_compatible: |
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
226 |
assumes "wf R" "wf S" |
26175 | 227 |
assumes "S O R \<subseteq> R" |
26044
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
228 |
shows "wf (R \<union> S)" |
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
229 |
proof (rule wfI_min) |
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
230 |
fix x :: 'a and Q |
26175 | 231 |
let ?Q' = "{x \<in> Q. \<forall>y. (y, x) \<in> R \<longrightarrow> y \<notin> Q}" |
26044
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
232 |
assume "x \<in> Q" |
26175 | 233 |
obtain a where "a \<in> ?Q'" |
234 |
by (rule wfE_min [OF `wf R` `x \<in> Q`]) blast |
|
26044
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
235 |
with `wf S` |
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
236 |
obtain z where "z \<in> ?Q'" and zmin: "\<And>y. (y, z) \<in> S \<Longrightarrow> y \<notin> ?Q'" by (erule wfE_min) |
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
237 |
{ |
26175 | 238 |
fix y assume "(y, z) \<in> S" |
26044
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
239 |
then have "y \<notin> ?Q'" by (rule zmin) |
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
240 |
|
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
241 |
have "y \<notin> Q" |
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
242 |
proof |
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
243 |
assume "y \<in> Q" |
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
244 |
with `y \<notin> ?Q'` |
26175 | 245 |
obtain w where "(w, y) \<in> R" and "w \<in> Q" by auto |
246 |
from `(w, y) \<in> R` `(y, z) \<in> S` have "(w, z) \<in> S O R" by (rule rel_compI) |
|
247 |
with `S O R \<subseteq> R` have "(w, z) \<in> R" .. |
|
26044
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
248 |
with `z \<in> ?Q'` have "w \<notin> Q" by blast |
26175 | 249 |
with `w \<in> Q` show False by contradiction |
26044
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
250 |
qed |
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
251 |
} |
26175 | 252 |
with `z \<in> ?Q'` show "\<exists>z\<in>Q. \<forall>y. (y, z) \<in> R \<union> S \<longrightarrow> y \<notin> Q" by blast |
26044
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
253 |
qed |
32889481ec4c
lemma wf_union_compatible: "wf R ==> wf S ==> S O R <= R ==> wf (R Un S)"
krauss
parents:
25425
diff
changeset
|
254 |
|
26175 | 255 |
|
256 |
text {* Well-foundedness of indexed union with disjoint domains and ranges *} |
|
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
257 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
258 |
lemma wf_UN: "[| ALL i:I. wf(r i); |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
259 |
ALL i:I. ALL j:I. r i ~= r j --> Domain(r i) Int Range(r j) = {} |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
260 |
|] ==> wf(UN i:I. r i)" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
261 |
apply (simp only: wf_eq_minimal, clarify) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
262 |
apply (rename_tac A a, case_tac "EX i:I. EX a:A. EX b:A. (b,a) : r i") |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
263 |
prefer 2 |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
264 |
apply force |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
265 |
apply clarify |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
266 |
apply (drule bspec, assumption) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
267 |
apply (erule_tac x="{a. a:A & (EX b:A. (b,a) : r i) }" in allE) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
268 |
apply (blast elim!: allE) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
269 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
270 |
|
23744
7c9e6e2fe249
Adapted to changes in infrastructure for converting between
berghofe
parents:
23389
diff
changeset
|
271 |
lemmas wfP_SUP = wf_UN [where I=UNIV and r="\<lambda>i. {(x, y). r i x y}", |
7c9e6e2fe249
Adapted to changes in infrastructure for converting between
berghofe
parents:
23389
diff
changeset
|
272 |
to_pred SUP_UN_eq2 bot_empty_eq, simplified, standard] |
22263 | 273 |
|
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
274 |
lemma wf_Union: |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
275 |
"[| ALL r:R. wf r; |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
276 |
ALL r:R. ALL s:R. r ~= s --> Domain r Int Range s = {} |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
277 |
|] ==> wf(Union R)" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
278 |
apply (simp add: Union_def) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
279 |
apply (blast intro: wf_UN) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
280 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
281 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
282 |
(*Intuition: we find an (R u S)-min element of a nonempty subset A |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
283 |
by case distinction. |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
284 |
1. There is a step a -R-> b with a,b : A. |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
285 |
Pick an R-min element z of the (nonempty) set {a:A | EX b:A. a -R-> b}. |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
286 |
By definition, there is z':A s.t. z -R-> z'. Because z is R-min in the |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
287 |
subset, z' must be R-min in A. Because z' has an R-predecessor, it cannot |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
288 |
have an S-successor and is thus S-min in A as well. |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
289 |
2. There is no such step. |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
290 |
Pick an S-min element of A. In this case it must be an R-min |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
291 |
element of A as well. |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
292 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
293 |
*) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
294 |
lemma wf_Un: |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
295 |
"[| wf r; wf s; Domain r Int Range s = {} |] ==> wf(r Un s)" |
26175 | 296 |
using wf_union_compatible[of s r] |
297 |
by (auto simp: Un_ac) |
|
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
298 |
|
23186
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
299 |
lemma wf_union_merge: |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
300 |
"wf (R \<union> S) = wf (R O R \<union> R O S \<union> S)" (is "wf ?A = wf ?B") |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
301 |
proof |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
302 |
assume "wf ?A" |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
303 |
with wf_trancl have wfT: "wf (?A^+)" . |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
304 |
moreover have "?B \<subseteq> ?A^+" |
26175 | 305 |
by (subst trancl_unfold, subst trancl_unfold) blast |
23186
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
306 |
ultimately show "wf ?B" by (rule wf_subset) |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
307 |
next |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
308 |
assume "wf ?B" |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
309 |
|
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
310 |
show "wf ?A" |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
311 |
proof (rule wfI_min) |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
312 |
fix Q :: "'a set" and x |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
313 |
assume "x \<in> Q" |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
314 |
|
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
315 |
with `wf ?B` |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
316 |
obtain z where "z \<in> Q" and "\<And>y. (y, z) \<in> ?B \<Longrightarrow> y \<notin> Q" |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
317 |
by (erule wfE_min) |
26175 | 318 |
then have A1: "\<And>y. (y, z) \<in> R O R \<Longrightarrow> y \<notin> Q" |
23186
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
319 |
and A2: "\<And>y. (y, z) \<in> R O S \<Longrightarrow> y \<notin> Q" |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
320 |
and A3: "\<And>y. (y, z) \<in> S \<Longrightarrow> y \<notin> Q" |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
321 |
by auto |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
322 |
|
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
323 |
show "\<exists>z\<in>Q. \<forall>y. (y, z) \<in> ?A \<longrightarrow> y \<notin> Q" |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
324 |
proof (cases "\<forall>y. (y, z) \<in> R \<longrightarrow> y \<notin> Q") |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
325 |
case True |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
326 |
with `z \<in> Q` A3 show ?thesis by blast |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
327 |
next |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
328 |
case False |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
329 |
then obtain z' where "z'\<in>Q" "(z', z) \<in> R" by blast |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
330 |
|
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
331 |
have "\<forall>y. (y, z') \<in> ?A \<longrightarrow> y \<notin> Q" |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
332 |
proof (intro allI impI) |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
333 |
fix y assume "(y, z') \<in> ?A" |
26175 | 334 |
then show "y \<notin> Q" |
23186
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
335 |
proof |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
336 |
assume "(y, z') \<in> R" |
26175 | 337 |
then have "(y, z) \<in> R O R" using `(z', z) \<in> R` .. |
23186
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
338 |
with A1 show "y \<notin> Q" . |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
339 |
next |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
340 |
assume "(y, z') \<in> S" |
26175 | 341 |
then have "(y, z) \<in> R O S" using `(z', z) \<in> R` .. |
23186
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
342 |
with A2 show "y \<notin> Q" . |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
343 |
qed |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
344 |
qed |
23389 | 345 |
with `z' \<in> Q` show ?thesis .. |
23186
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
346 |
qed |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
347 |
qed |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
348 |
qed |
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
349 |
|
26175 | 350 |
lemma wf_comp_self: "wf R = wf (R O R)" -- {* special case *} |
351 |
by (rule wf_union_merge [where S = "{}", simplified]) |
|
23186
f948708bc100
Added "merge-lemma" about wellfoundedness of unions (useful for some termination proofs)
krauss
parents:
22845
diff
changeset
|
352 |
|
26175 | 353 |
|
354 |
subsubsection {* acyclic *} |
|
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
355 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
356 |
lemma acyclicI: "ALL x. (x, x) ~: r^+ ==> acyclic r" |
26175 | 357 |
by (simp add: acyclic_def) |
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
358 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
359 |
lemma wf_acyclic: "wf r ==> acyclic r" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
360 |
apply (simp add: acyclic_def) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
361 |
apply (blast elim: wf_trancl [THEN wf_irrefl]) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
362 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
363 |
|
22263 | 364 |
lemmas wfP_acyclicP = wf_acyclic [to_pred] |
365 |
||
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
366 |
lemma acyclic_insert [iff]: |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
367 |
"acyclic(insert (y,x) r) = (acyclic r & (x,y) ~: r^*)" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
368 |
apply (simp add: acyclic_def trancl_insert) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
369 |
apply (blast intro: rtrancl_trans) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
370 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
371 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
372 |
lemma acyclic_converse [iff]: "acyclic(r^-1) = acyclic r" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
373 |
by (simp add: acyclic_def trancl_converse) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
374 |
|
22263 | 375 |
lemmas acyclicP_converse [iff] = acyclic_converse [to_pred] |
376 |
||
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
377 |
lemma acyclic_impl_antisym_rtrancl: "acyclic r ==> antisym(r^*)" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
378 |
apply (simp add: acyclic_def antisym_def) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
379 |
apply (blast elim: rtranclE intro: rtrancl_into_trancl1 rtrancl_trancl_trancl) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
380 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
381 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
382 |
(* Other direction: |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
383 |
acyclic = no loops |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
384 |
antisym = only self loops |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
385 |
Goalw [acyclic_def,antisym_def] "antisym( r^* ) ==> acyclic(r - Id) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
386 |
==> antisym( r^* ) = acyclic(r - Id)"; |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
387 |
*) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
388 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
389 |
lemma acyclic_subset: "[| acyclic s; r <= s |] ==> acyclic r" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
390 |
apply (simp add: acyclic_def) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
391 |
apply (blast intro: trancl_mono) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
392 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
393 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
394 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
395 |
subsection{*Well-Founded Recursion*} |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
396 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
397 |
text{*cut*} |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
398 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
399 |
lemma cuts_eq: "(cut f r x = cut g r x) = (ALL y. (y,x):r --> f(y)=g(y))" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
400 |
by (simp add: expand_fun_eq cut_def) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
401 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
402 |
lemma cut_apply: "(x,a):r ==> (cut f r a)(x) = f(x)" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
403 |
by (simp add: cut_def) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
404 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
405 |
text{*Inductive characterization of wfrec combinator; for details see: |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
406 |
John Harrison, "Inductive definitions: automation and application"*} |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
407 |
|
22263 | 408 |
lemma wfrec_unique: "[| adm_wf R F; wf R |] ==> EX! y. wfrec_rel R F x y" |
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
409 |
apply (simp add: adm_wf_def) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
410 |
apply (erule_tac a=x in wf_induct) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
411 |
apply (rule ex1I) |
22263 | 412 |
apply (rule_tac g = "%x. THE y. wfrec_rel R F x y" in wfrec_rel.wfrecI) |
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
413 |
apply (fast dest!: theI') |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
414 |
apply (erule wfrec_rel.cases, simp) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
415 |
apply (erule allE, erule allE, erule allE, erule mp) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
416 |
apply (fast intro: the_equality [symmetric]) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
417 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
418 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
419 |
lemma adm_lemma: "adm_wf R (%f x. F (cut f R x) x)" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
420 |
apply (simp add: adm_wf_def) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
421 |
apply (intro strip) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
422 |
apply (rule cuts_eq [THEN iffD2, THEN subst], assumption) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
423 |
apply (rule refl) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
424 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
425 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
426 |
lemma wfrec: "wf(r) ==> wfrec r H a = H (cut (wfrec r H) r a) a" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
427 |
apply (simp add: wfrec_def) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
428 |
apply (rule adm_lemma [THEN wfrec_unique, THEN the1_equality], assumption) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
429 |
apply (rule wfrec_rel.wfrecI) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
430 |
apply (intro strip) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
431 |
apply (erule adm_lemma [THEN wfrec_unique, THEN theI']) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
432 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
433 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
434 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
435 |
text{** This form avoids giant explosions in proofs. NOTE USE OF ==*} |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
436 |
lemma def_wfrec: "[| f==wfrec r H; wf(r) |] ==> f(a) = H (cut f r a) a" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
437 |
apply auto |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
438 |
apply (blast intro: wfrec) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
439 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
440 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
441 |
|
17459 | 442 |
subsection {* Code generator setup *} |
443 |
||
444 |
consts_code |
|
17654
38496187809d
Renamed wf_rec to wfrec in consts_code declaration.
berghofe
parents:
17459
diff
changeset
|
445 |
"wfrec" ("\<module>wfrec?") |
17459 | 446 |
attach {* |
17654
38496187809d
Renamed wf_rec to wfrec in consts_code declaration.
berghofe
parents:
17459
diff
changeset
|
447 |
fun wfrec f x = f (wfrec f) x; |
17459 | 448 |
*} |
449 |
||
450 |
||
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
451 |
subsection{*Variants for TFL: the Recdef Package*} |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
452 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
453 |
lemma tfl_wf_induct: "ALL R. wf R --> |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
454 |
(ALL P. (ALL x. (ALL y. (y,x):R --> P y) --> P x) --> (ALL x. P x))" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
455 |
apply clarify |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
456 |
apply (rule_tac r = R and P = P and a = x in wf_induct, assumption, blast) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
457 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
458 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
459 |
lemma tfl_cut_apply: "ALL f R. (x,a):R --> (cut f R a)(x) = f(x)" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
460 |
apply clarify |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
461 |
apply (rule cut_apply, assumption) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
462 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
463 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
464 |
lemma tfl_wfrec: |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
465 |
"ALL M R f. (f=wfrec R M) --> wf R --> (ALL x. f x = M (cut f R x) x)" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
466 |
apply clarify |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
467 |
apply (erule wfrec) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
468 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
469 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
470 |
subsection {*LEAST and wellorderings*} |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
471 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
472 |
text{* See also @{text wf_linord_ex_has_least} and its consequences in |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
473 |
@{text Wellfounded_Relations.ML}*} |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
474 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
475 |
lemma wellorder_Least_lemma [rule_format]: |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
476 |
"P (k::'a::wellorder) --> P (LEAST x. P(x)) & (LEAST x. P(x)) <= k" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
477 |
apply (rule_tac a = k in wf [THEN wf_induct]) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
478 |
apply (rule impI) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
479 |
apply (rule classical) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
480 |
apply (rule_tac s = x in Least_equality [THEN ssubst], auto) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
481 |
apply (auto simp add: linorder_not_less [symmetric]) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
482 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
483 |
|
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
484 |
lemmas LeastI = wellorder_Least_lemma [THEN conjunct1, standard] |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
485 |
lemmas Least_le = wellorder_Least_lemma [THEN conjunct2, standard] |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
486 |
|
15950 | 487 |
-- "The following 3 lemmas are due to Brian Huffman" |
488 |
lemma LeastI_ex: "EX x::'a::wellorder. P x ==> P (Least P)" |
|
489 |
apply (erule exE) |
|
490 |
apply (erule LeastI) |
|
491 |
done |
|
492 |
||
493 |
lemma LeastI2: |
|
494 |
"[| P (a::'a::wellorder); !!x. P x ==> Q x |] ==> Q (Least P)" |
|
495 |
by (blast intro: LeastI) |
|
496 |
||
497 |
lemma LeastI2_ex: |
|
498 |
"[| EX a::'a::wellorder. P a; !!x. P x ==> Q x |] ==> Q (Least P)" |
|
499 |
by (blast intro: LeastI_ex) |
|
500 |
||
15341
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
501 |
lemma not_less_Least: "[| k < (LEAST x. P x) |] ==> ~P (k::'a::wellorder)" |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
502 |
apply (simp (no_asm_use) add: linorder_not_le [symmetric]) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
503 |
apply (erule contrapos_nn) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
504 |
apply (erule Least_le) |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
505 |
done |
254f6f00b60e
converted to Isar script, simplifying some results
paulson
parents:
11451
diff
changeset
|
506 |
|
26072
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
507 |
subsection {* @{typ nat} is well-founded *} |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
508 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
509 |
lemma less_nat_rel: "op < = (\<lambda>m n. n = Suc m)^++" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
510 |
proof (rule ext, rule ext, rule iffI) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
511 |
fix n m :: nat |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
512 |
assume "m < n" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
513 |
then show "(\<lambda>m n. n = Suc m)^++ m n" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
514 |
proof (induct n) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
515 |
case 0 then show ?case by auto |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
516 |
next |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
517 |
case (Suc n) then show ?case |
26175 | 518 |
by (auto simp add: less_Suc_eq_le le_less intro: tranclp.trancl_into_trancl) |
26072
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
519 |
qed |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
520 |
next |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
521 |
fix n m :: nat |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
522 |
assume "(\<lambda>m n. n = Suc m)^++ m n" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
523 |
then show "m < n" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
524 |
by (induct n) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
525 |
(simp_all add: less_Suc_eq_le reflexive le_less) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
526 |
qed |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
527 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
528 |
definition |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
529 |
pred_nat :: "(nat * nat) set" where |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
530 |
"pred_nat = {(m, n). n = Suc m}" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
531 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
532 |
definition |
26235 | 533 |
less_than :: "(nat * nat) set" where |
26072
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
534 |
"less_than = pred_nat^+" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
535 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
536 |
lemma less_eq: "(m, n) \<in> pred_nat^+ \<longleftrightarrow> m < n" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
537 |
unfolding less_nat_rel pred_nat_def trancl_def by simp |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
538 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
539 |
lemma pred_nat_trancl_eq_le: |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
540 |
"(m, n) \<in> pred_nat^* \<longleftrightarrow> m \<le> n" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
541 |
unfolding less_eq rtrancl_eq_or_trancl by auto |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
542 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
543 |
lemma wf_pred_nat: "wf pred_nat" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
544 |
apply (unfold wf_def pred_nat_def, clarify) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
545 |
apply (induct_tac x, blast+) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
546 |
done |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
547 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
548 |
lemma wf_less_than [iff]: "wf less_than" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
549 |
by (simp add: less_than_def wf_pred_nat [THEN wf_trancl]) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
550 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
551 |
lemma trans_less_than [iff]: "trans less_than" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
552 |
by (simp add: less_than_def trans_trancl) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
553 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
554 |
lemma less_than_iff [iff]: "((x,y): less_than) = (x<y)" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
555 |
by (simp add: less_than_def less_eq) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
556 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
557 |
lemma wf_less: "wf {(x, y::nat). x < y}" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
558 |
using wf_less_than by (simp add: less_than_def less_eq [symmetric]) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
559 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
560 |
text {* Complete induction, aka course-of-values induction *} |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
561 |
lemma nat_less_induct: |
26175 | 562 |
assumes "!!n. \<forall>m::nat. m < n --> P m ==> P n" shows "P n" |
26072
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
563 |
apply (induct n rule: wf_induct [OF wf_pred_nat [THEN wf_trancl]]) |
26175 | 564 |
apply (rule assms) |
26072
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
565 |
apply (unfold less_eq [symmetric], assumption) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
566 |
done |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
567 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
568 |
lemmas less_induct = nat_less_induct [rule_format, case_names less] |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
569 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
570 |
text {* Type @{typ nat} is a wellfounded order *} |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
571 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
572 |
instance nat :: wellorder |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
573 |
by intro_classes |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
574 |
(assumption | |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
575 |
rule le_refl le_trans le_anti_sym nat_less_le nat_le_linear wf_less)+ |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
576 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
577 |
lemma nat_induct2: "[|P 0; P (Suc 0); !!k. P k ==> P (Suc (Suc k))|] ==> P n" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
578 |
apply (rule nat_less_induct) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
579 |
apply (case_tac n) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
580 |
apply (case_tac [2] nat) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
581 |
apply (blast intro: less_trans)+ |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
582 |
done |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
583 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
584 |
text {* The method of infinite descent, frequently used in number theory. |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
585 |
Provided by Roelof Oosterhuis. |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
586 |
$P(n)$ is true for all $n\in\mathbb{N}$ if |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
587 |
\begin{itemize} |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
588 |
\item case ``0'': given $n=0$ prove $P(n)$, |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
589 |
\item case ``smaller'': given $n>0$ and $\neg P(n)$ prove there exists |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
590 |
a smaller integer $m$ such that $\neg P(m)$. |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
591 |
\end{itemize} *} |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
592 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
593 |
lemma infinite_descent0[case_names 0 smaller]: |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
594 |
"\<lbrakk> P 0; !!n. n>0 \<Longrightarrow> \<not> P n \<Longrightarrow> (\<exists>m::nat. m < n \<and> \<not>P m) \<rbrakk> \<Longrightarrow> P n" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
595 |
by (induct n rule: less_induct, case_tac "n>0", auto) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
596 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
597 |
text{* A compact version without explicit base case: *} |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
598 |
lemma infinite_descent: |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
599 |
"\<lbrakk> !!n::nat. \<not> P n \<Longrightarrow> \<exists>m<n. \<not> P m \<rbrakk> \<Longrightarrow> P n" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
600 |
by (induct n rule: less_induct, auto) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
601 |
|
26175 | 602 |
text {* |
603 |
Infinite descent using a mapping to $\mathbb{N}$: |
|
26072
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
604 |
$P(x)$ is true for all $x\in D$ if there exists a $V: D \to \mathbb{N}$ and |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
605 |
\begin{itemize} |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
606 |
\item case ``0'': given $V(x)=0$ prove $P(x)$, |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
607 |
\item case ``smaller'': given $V(x)>0$ and $\neg P(x)$ prove there exists a $y \in D$ such that $V(y)<V(x)$ and $~\neg P(y)$. |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
608 |
\end{itemize} |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
609 |
NB: the proof also shows how to use the previous lemma. *} |
26175 | 610 |
|
26072
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
611 |
corollary infinite_descent0_measure [case_names 0 smaller]: |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
612 |
assumes A0: "!!x. V x = (0::nat) \<Longrightarrow> P x" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
613 |
and A1: "!!x. V x > 0 \<Longrightarrow> \<not>P x \<Longrightarrow> (\<exists>y. V y < V x \<and> \<not>P y)" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
614 |
shows "P x" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
615 |
proof - |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
616 |
obtain n where "n = V x" by auto |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
617 |
moreover have "\<And>x. V x = n \<Longrightarrow> P x" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
618 |
proof (induct n rule: infinite_descent0) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
619 |
case 0 -- "i.e. $V(x) = 0$" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
620 |
with A0 show "P x" by auto |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
621 |
next -- "now $n>0$ and $P(x)$ does not hold for some $x$ with $V(x)=n$" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
622 |
case (smaller n) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
623 |
then obtain x where vxn: "V x = n " and "V x > 0 \<and> \<not> P x" by auto |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
624 |
with A1 obtain y where "V y < V x \<and> \<not> P y" by auto |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
625 |
with vxn obtain m where "m = V y \<and> m<n \<and> \<not> P y" by auto |
26175 | 626 |
then show ?case by auto |
26072
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
627 |
qed |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
628 |
ultimately show "P x" by auto |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
629 |
qed |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
630 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
631 |
text{* Again, without explicit base case: *} |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
632 |
lemma infinite_descent_measure: |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
633 |
assumes "!!x. \<not> P x \<Longrightarrow> \<exists>y. (V::'a\<Rightarrow>nat) y < V x \<and> \<not> P y" shows "P x" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
634 |
proof - |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
635 |
from assms obtain n where "n = V x" by auto |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
636 |
moreover have "!!x. V x = n \<Longrightarrow> P x" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
637 |
proof (induct n rule: infinite_descent, auto) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
638 |
fix x assume "\<not> P x" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
639 |
with assms show "\<exists>m < V x. \<exists>y. V y = m \<and> \<not> P y" by auto |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
640 |
qed |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
641 |
ultimately show "P x" by auto |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
642 |
qed |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
643 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
644 |
text {* @{text LEAST} theorems for type @{typ nat}*} |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
645 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
646 |
lemma Least_Suc: |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
647 |
"[| P n; ~ P 0 |] ==> (LEAST n. P n) = Suc (LEAST m. P(Suc m))" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
648 |
apply (case_tac "n", auto) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
649 |
apply (frule LeastI) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
650 |
apply (drule_tac P = "%x. P (Suc x) " in LeastI) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
651 |
apply (subgoal_tac " (LEAST x. P x) \<le> Suc (LEAST x. P (Suc x))") |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
652 |
apply (erule_tac [2] Least_le) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
653 |
apply (case_tac "LEAST x. P x", auto) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
654 |
apply (drule_tac P = "%x. P (Suc x) " in Least_le) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
655 |
apply (blast intro: order_antisym) |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
656 |
done |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
657 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
658 |
lemma Least_Suc2: |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
659 |
"[|P n; Q m; ~P 0; !k. P (Suc k) = Q k|] ==> Least P = Suc (Least Q)" |
26175 | 660 |
apply (erule (1) Least_Suc [THEN ssubst]) |
661 |
apply simp |
|
662 |
done |
|
26072
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
663 |
|
26105
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
nipkow
parents:
26072
diff
changeset
|
664 |
lemma ex_least_nat_le: "\<not>P(0) \<Longrightarrow> P(n::nat) \<Longrightarrow> \<exists>k\<le>n. (\<forall>i<k. \<not>P i) & P(k)" |
26175 | 665 |
apply (cases n) |
666 |
apply blast |
|
667 |
apply (rule_tac x="LEAST k. P(k)" in exI) |
|
668 |
apply (blast intro: Least_le dest: not_less_Least intro: LeastI_ex) |
|
669 |
done |
|
26105
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
nipkow
parents:
26072
diff
changeset
|
670 |
|
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
nipkow
parents:
26072
diff
changeset
|
671 |
lemma ex_least_nat_less: "\<not>P(0) \<Longrightarrow> P(n::nat) \<Longrightarrow> \<exists>k<n. (\<forall>i\<le>k. \<not>P i) & P(k+1)" |
26175 | 672 |
apply (cases n) |
673 |
apply blast |
|
674 |
apply (frule (1) ex_least_nat_le) |
|
675 |
apply (erule exE) |
|
676 |
apply (case_tac k) |
|
677 |
apply simp |
|
678 |
apply (rename_tac k1) |
|
679 |
apply (rule_tac x=k1 in exI) |
|
680 |
apply fastsimp |
|
681 |
done |
|
26105
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
nipkow
parents:
26072
diff
changeset
|
682 |
|
26072
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
683 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
684 |
subsection {* size of a datatype value *} |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
685 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
686 |
use "Tools/function_package/size.ML" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
687 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
688 |
setup Size.setup |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
689 |
|
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
690 |
lemma nat_size [simp, code func]: "size (n\<Colon>nat) = n" |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
691 |
by (induct n) simp_all |
f65a7fa2da6c
<= and < on nat no longer depend on wellfounded relations
haftmann
parents:
26044
diff
changeset
|
692 |
|
10213 | 693 |
end |