author | wenzelm |
Tue, 31 Oct 2023 16:49:03 +0100 | |
changeset 78869 | f464f6bc5809 |
parent 76419 | f20865ad6319 |
child 81125 | ec121999a9cb |
permissions | -rw-r--r-- |
13165 | 1 |
(* Title: ZF/WF.thy |
1478 | 2 |
Author: Tobias Nipkow and Lawrence C Paulson |
435 | 3 |
Copyright 1994 University of Cambridge |
0 | 4 |
|
13165 | 5 |
Derived first for transitive relations, and finally for arbitrary WF relations |
6 |
via wf_trancl and trans_trancl. |
|
7 |
||
8 |
It is difficult to derive this general case directly, using r^+ instead of |
|
9 |
r. In is_recfun, the two occurrences of the relation must have the same |
|
10 |
form. Inserting r^+ in the_recfun or wftrec yields a recursion rule with |
|
11 |
r^+ -`` {a} instead of r-``{a}. This recursion rule is stronger in |
|
12 |
principle, but harder to use, especially to prove wfrec_eclose_eq in |
|
13 |
epsilon.ML. Expanding out the definition of wftrec in wfrec would yield |
|
14 |
a mess. |
|
0 | 15 |
*) |
16 |
||
60770 | 17 |
section\<open>Well-Founded Recursion\<close> |
13356 | 18 |
|
16417 | 19 |
theory WF imports Trancl begin |
13165 | 20 |
|
24893 | 21 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
22 |
wf :: "i\<Rightarrow>o" where |
13165 | 23 |
(*r is a well-founded relation*) |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
24 |
"wf(r) \<equiv> \<forall>Z. Z=0 | (\<exists>x\<in>Z. \<forall>y. \<langle>y,x\<rangle>:r \<longrightarrow> \<not> y \<in> Z)" |
13165 | 25 |
|
24893 | 26 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
27 |
wf_on :: "[i,i]\<Rightarrow>o" (\<open>wf[_]'(_')\<close>) where |
13165 | 28 |
(*r is well-founded on A*) |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
29 |
"wf_on(A,r) \<equiv> wf(r \<inter> A*A)" |
13165 | 30 |
|
24893 | 31 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
32 |
is_recfun :: "[i, i, [i,i]\<Rightarrow>i, i] \<Rightarrow>o" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
33 |
"is_recfun(r,a,H,f) \<equiv> (f = (\<lambda>x\<in>r-``{a}. H(x, restrict(f, r-``{x}))))" |
13165 | 34 |
|
24893 | 35 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
36 |
the_recfun :: "[i, i, [i,i]\<Rightarrow>i] \<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
37 |
"the_recfun(r,a,H) \<equiv> (THE f. is_recfun(r,a,H,f))" |
13165 | 38 |
|
24893 | 39 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
40 |
wftrec :: "[i, i, [i,i]\<Rightarrow>i] \<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
41 |
"wftrec(r,a,H) \<equiv> H(a, the_recfun(r,a,H))" |
13165 | 42 |
|
24893 | 43 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
44 |
wfrec :: "[i, i, [i,i]\<Rightarrow>i] \<Rightarrow>i" where |
13165 | 45 |
(*public version. Does not require r to be transitive*) |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
46 |
"wfrec(r,a,H) \<equiv> wftrec(r^+, a, \<lambda>x f. H(x, restrict(f,r-``{x})))" |
13165 | 47 |
|
24893 | 48 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
49 |
wfrec_on :: "[i, i, i, [i,i]\<Rightarrow>i] \<Rightarrow>i" (\<open>wfrec[_]'(_,_,_')\<close>) where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
50 |
"wfrec[A](r,a,H) \<equiv> wfrec(r \<inter> A*A, a, H)" |
13165 | 51 |
|
52 |
||
60770 | 53 |
subsection\<open>Well-Founded Relations\<close> |
13165 | 54 |
|
69593 | 55 |
subsubsection\<open>Equivalences between \<^term>\<open>wf\<close> and \<^term>\<open>wf_on\<close>\<close> |
13165 | 56 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
57 |
lemma wf_imp_wf_on: "wf(r) \<Longrightarrow> wf[A](r)" |
46820 | 58 |
by (unfold wf_def wf_on_def, force) |
13165 | 59 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
60 |
lemma wf_on_imp_wf: "\<lbrakk>wf[A](r); r \<subseteq> A*A\<rbrakk> \<Longrightarrow> wf(r)" |
13248 | 61 |
by (simp add: wf_on_def subset_Int_iff) |
62 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
63 |
lemma wf_on_field_imp_wf: "wf[field(r)](r) \<Longrightarrow> wf(r)" |
13165 | 64 |
by (unfold wf_def wf_on_def, fast) |
65 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
46820
diff
changeset
|
66 |
lemma wf_iff_wf_on_field: "wf(r) \<longleftrightarrow> wf[field(r)](r)" |
13165 | 67 |
by (blast intro: wf_imp_wf_on wf_on_field_imp_wf) |
68 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
69 |
lemma wf_on_subset_A: "\<lbrakk>wf[A](r); B<=A\<rbrakk> \<Longrightarrow> wf[B](r)" |
13165 | 70 |
by (unfold wf_on_def wf_def, fast) |
71 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
72 |
lemma wf_on_subset_r: "\<lbrakk>wf[A](r); s<=r\<rbrakk> \<Longrightarrow> wf[A](s)" |
13165 | 73 |
by (unfold wf_on_def wf_def, fast) |
74 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
75 |
lemma wf_subset: "\<lbrakk>wf(s); r<=s\<rbrakk> \<Longrightarrow> wf(r)" |
13217 | 76 |
by (simp add: wf_def, fast) |
77 |
||
69593 | 78 |
subsubsection\<open>Introduction Rules for \<^term>\<open>wf_on\<close>\<close> |
13165 | 79 |
|
69593 | 80 |
text\<open>If every non-empty subset of \<^term>\<open>A\<close> has an \<^term>\<open>r\<close>-minimal element |
81 |
then we have \<^term>\<open>wf[A](r)\<close>.\<close> |
|
13165 | 82 |
lemma wf_onI: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
83 |
assumes prem: "\<And>Z u. \<lbrakk>Z<=A; u \<in> Z; \<forall>x\<in>Z. \<exists>y\<in>Z. \<langle>y,x\<rangle>:r\<rbrakk> \<Longrightarrow> False" |
13165 | 84 |
shows "wf[A](r)" |
76217 | 85 |
unfolding wf_on_def wf_def |
13165 | 86 |
apply (rule equals0I [THEN disjCI, THEN allI]) |
13784 | 87 |
apply (rule_tac Z = Z in prem, blast+) |
13165 | 88 |
done |
89 |
||
69593 | 90 |
text\<open>If \<^term>\<open>r\<close> allows well-founded induction over \<^term>\<open>A\<close> |
91 |
then we have \<^term>\<open>wf[A](r)\<close>. Premise is equivalent to |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
92 |
\<^prop>\<open>\<And>B. \<forall>x\<in>A. (\<forall>y. \<langle>y,x\<rangle>: r \<longrightarrow> y \<in> B) \<longrightarrow> x \<in> B \<Longrightarrow> A<=B\<close>\<close> |
13165 | 93 |
lemma wf_onI2: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
94 |
assumes prem: "\<And>y B. \<lbrakk>\<forall>x\<in>A. (\<forall>y\<in>A. \<langle>y,x\<rangle>:r \<longrightarrow> y \<in> B) \<longrightarrow> x \<in> B; y \<in> A\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
95 |
\<Longrightarrow> y \<in> B" |
13165 | 96 |
shows "wf[A](r)" |
97 |
apply (rule wf_onI) |
|
98 |
apply (rule_tac c=u in prem [THEN DiffE]) |
|
46820 | 99 |
prefer 3 apply blast |
13165 | 100 |
apply fast+ |
101 |
done |
|
102 |
||
103 |
||
60770 | 104 |
subsubsection\<open>Well-founded Induction\<close> |
13165 | 105 |
|
69593 | 106 |
text\<open>Consider the least \<^term>\<open>z\<close> in \<^term>\<open>domain(r)\<close> such that |
107 |
\<^term>\<open>P(z)\<close> does not hold...\<close> |
|
46993 | 108 |
lemma wf_induct_raw: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
109 |
"\<lbrakk>wf(r); |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
110 |
\<And>x.\<lbrakk>\<forall>y. \<langle>y,x\<rangle>: r \<longrightarrow> P(y)\<rbrakk> \<Longrightarrow> P(x)\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
111 |
\<Longrightarrow> P(a)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
112 |
unfolding wf_def |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
113 |
apply (erule_tac x = "{z \<in> domain(r). \<not> P(z)}" in allE) |
46820 | 114 |
apply blast |
13165 | 115 |
done |
435 | 116 |
|
46993 | 117 |
lemmas wf_induct = wf_induct_raw [rule_format, consumes 1, case_names step, induct set: wf] |
13203
fac77a839aa2
Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents:
13175
diff
changeset
|
118 |
|
61798 | 119 |
text\<open>The form of this rule is designed to match \<open>wfI\<close>\<close> |
13165 | 120 |
lemma wf_induct2: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
121 |
"\<lbrakk>wf(r); a \<in> A; field(r)<=A; |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
122 |
\<And>x.\<lbrakk>x \<in> A; \<forall>y. \<langle>y,x\<rangle>: r \<longrightarrow> P(y)\<rbrakk> \<Longrightarrow> P(x)\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
123 |
\<Longrightarrow> P(a)" |
46953 | 124 |
apply (erule_tac P="a \<in> A" in rev_mp) |
46820 | 125 |
apply (erule_tac a=a in wf_induct, blast) |
13165 | 126 |
done |
127 |
||
46820 | 128 |
lemma field_Int_square: "field(r \<inter> A*A) \<subseteq> A" |
13165 | 129 |
by blast |
130 |
||
46993 | 131 |
lemma wf_on_induct_raw [consumes 2, induct set: wf_on]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
132 |
"\<lbrakk>wf[A](r); a \<in> A; |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
133 |
\<And>x.\<lbrakk>x \<in> A; \<forall>y\<in>A. \<langle>y,x\<rangle>: r \<longrightarrow> P(y)\<rbrakk> \<Longrightarrow> P(x) |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
134 |
\<rbrakk> \<Longrightarrow> P(a)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
135 |
unfolding wf_on_def |
13165 | 136 |
apply (erule wf_induct2, assumption) |
137 |
apply (rule field_Int_square, blast) |
|
138 |
done |
|
139 |
||
71085
950e1cfe0fe4
tuned proofs -- more stable proof terms without [rule_format];
wenzelm
parents:
69593
diff
changeset
|
140 |
lemma wf_on_induct [consumes 2, case_names step, induct set: wf_on]: |
950e1cfe0fe4
tuned proofs -- more stable proof terms without [rule_format];
wenzelm
parents:
69593
diff
changeset
|
141 |
"wf[A](r) \<Longrightarrow> a \<in> A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> (\<And>y. y \<in> A \<Longrightarrow> \<langle>y, x\<rangle> \<in> r \<Longrightarrow> P(y)) \<Longrightarrow> P(x)) \<Longrightarrow> P(a)" |
950e1cfe0fe4
tuned proofs -- more stable proof terms without [rule_format];
wenzelm
parents:
69593
diff
changeset
|
142 |
using wf_on_induct_raw [of A r a P] by simp |
13203
fac77a839aa2
Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents:
13175
diff
changeset
|
143 |
|
69593 | 144 |
text\<open>If \<^term>\<open>r\<close> allows well-founded induction |
145 |
then we have \<^term>\<open>wf(r)\<close>.\<close> |
|
13165 | 146 |
lemma wfI: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
147 |
"\<lbrakk>field(r)<=A; |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
148 |
\<And>y B. \<lbrakk>\<forall>x\<in>A. (\<forall>y\<in>A. \<langle>y,x\<rangle>:r \<longrightarrow> y \<in> B) \<longrightarrow> x \<in> B; y \<in> A\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
149 |
\<Longrightarrow> y \<in> B\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
150 |
\<Longrightarrow> wf(r)" |
13165 | 151 |
apply (rule wf_on_subset_A [THEN wf_on_field_imp_wf]) |
152 |
apply (rule wf_onI2) |
|
46820 | 153 |
prefer 2 apply blast |
154 |
apply blast |
|
13165 | 155 |
done |
156 |
||
157 |
||
60770 | 158 |
subsection\<open>Basic Properties of Well-Founded Relations\<close> |
13165 | 159 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
160 |
lemma wf_not_refl: "wf(r) \<Longrightarrow> \<langle>a,a\<rangle> \<notin> r" |
13165 | 161 |
by (erule_tac a=a in wf_induct, blast) |
162 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
163 |
lemma wf_not_sym [rule_format]: "wf(r) \<Longrightarrow> \<forall>x. \<langle>a,x\<rangle>:r \<longrightarrow> \<langle>x,a\<rangle> \<notin> r" |
13165 | 164 |
by (erule_tac a=a in wf_induct, blast) |
165 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
166 |
(* @{term"\<lbrakk>wf(r); \<langle>a,x\<rangle> \<in> r; \<not>P \<Longrightarrow> \<langle>x,a\<rangle> \<in> r\<rbrakk> \<Longrightarrow> P"} *) |
45602 | 167 |
lemmas wf_asym = wf_not_sym [THEN swap] |
13165 | 168 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
169 |
lemma wf_on_not_refl: "\<lbrakk>wf[A](r); a \<in> A\<rbrakk> \<Longrightarrow> \<langle>a,a\<rangle> \<notin> r" |
13269 | 170 |
by (erule_tac a=a in wf_on_induct, assumption, blast) |
0 | 171 |
|
71085
950e1cfe0fe4
tuned proofs -- more stable proof terms without [rule_format];
wenzelm
parents:
69593
diff
changeset
|
172 |
lemma wf_on_not_sym: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
173 |
"\<lbrakk>wf[A](r); a \<in> A\<rbrakk> \<Longrightarrow> (\<And>b. b\<in>A \<Longrightarrow> \<langle>a,b\<rangle>:r \<Longrightarrow> \<langle>b,a\<rangle>\<notin>r)" |
71085
950e1cfe0fe4
tuned proofs -- more stable proof terms without [rule_format];
wenzelm
parents:
69593
diff
changeset
|
174 |
apply (atomize (full), intro impI) |
13269 | 175 |
apply (erule_tac a=a in wf_on_induct, assumption, blast) |
13165 | 176 |
done |
177 |
||
178 |
lemma wf_on_asym: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
179 |
"\<lbrakk>wf[A](r); \<not>Z \<Longrightarrow> \<langle>a,b\<rangle> \<in> r; |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
180 |
\<langle>b,a\<rangle> \<notin> r \<Longrightarrow> Z; \<not>Z \<Longrightarrow> a \<in> A; \<not>Z \<Longrightarrow> b \<in> A\<rbrakk> \<Longrightarrow> Z" |
46820 | 181 |
by (blast dest: wf_on_not_sym) |
13165 | 182 |
|
183 |
||
184 |
(*Needed to prove well_ordI. Could also reason that wf[A](r) means |
|
46820 | 185 |
wf(r \<inter> A*A); thus wf( (r \<inter> A*A)^+ ) and use wf_not_refl *) |
13165 | 186 |
lemma wf_on_chain3: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
187 |
"\<lbrakk>wf[A](r); \<langle>a,b\<rangle>:r; \<langle>b,c\<rangle>:r; \<langle>c,a\<rangle>:r; a \<in> A; b \<in> A; c \<in> A\<rbrakk> \<Longrightarrow> P" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
188 |
apply (subgoal_tac "\<forall>y\<in>A. \<forall>z\<in>A. \<langle>a,y\<rangle>:r \<longrightarrow> \<langle>y,z\<rangle>:r \<longrightarrow> \<langle>z,a\<rangle>:r \<longrightarrow> P", |
46820 | 189 |
blast) |
13269 | 190 |
apply (erule_tac a=a in wf_on_induct, assumption, blast) |
13165 | 191 |
done |
192 |
||
193 |
||
194 |
||
60770 | 195 |
text\<open>transitive closure of a WF relation is WF provided |
69593 | 196 |
\<^term>\<open>A\<close> is downward closed\<close> |
13165 | 197 |
lemma wf_on_trancl: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
198 |
"\<lbrakk>wf[A](r); r-``A \<subseteq> A\<rbrakk> \<Longrightarrow> wf[A](r^+)" |
13165 | 199 |
apply (rule wf_onI2) |
200 |
apply (frule bspec [THEN mp], assumption+) |
|
13784 | 201 |
apply (erule_tac a = y in wf_on_induct, assumption) |
46820 | 202 |
apply (blast elim: tranclE, blast) |
13165 | 203 |
done |
204 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
205 |
lemma wf_trancl: "wf(r) \<Longrightarrow> wf(r^+)" |
13165 | 206 |
apply (simp add: wf_iff_wf_on_field) |
46820 | 207 |
apply (rule wf_on_subset_A) |
13165 | 208 |
apply (erule wf_on_trancl) |
46820 | 209 |
apply blast |
13165 | 210 |
apply (rule trancl_type [THEN field_rel_subset]) |
211 |
done |
|
212 |
||
213 |
||
69593 | 214 |
text\<open>\<^term>\<open>r-``{a}\<close> is the set of everything under \<^term>\<open>a\<close> in \<^term>\<open>r\<close>\<close> |
13165 | 215 |
|
45602 | 216 |
lemmas underI = vimage_singleton_iff [THEN iffD2] |
217 |
lemmas underD = vimage_singleton_iff [THEN iffD1] |
|
13165 | 218 |
|
13634 | 219 |
|
69593 | 220 |
subsection\<open>The Predicate \<^term>\<open>is_recfun\<close>\<close> |
0 | 221 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
222 |
lemma is_recfun_type: "is_recfun(r,a,H,f) \<Longrightarrow> f \<in> r-``{a} -> range(f)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
223 |
unfolding is_recfun_def |
13165 | 224 |
apply (erule ssubst) |
225 |
apply (rule lamI [THEN rangeI, THEN lam_type], assumption) |
|
226 |
done |
|
227 |
||
13269 | 228 |
lemmas is_recfun_imp_function = is_recfun_type [THEN fun_is_function] |
229 |
||
13165 | 230 |
lemma apply_recfun: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
231 |
"\<lbrakk>is_recfun(r,a,H,f); \<langle>x,a\<rangle>:r\<rbrakk> \<Longrightarrow> f`x = H(x, restrict(f,r-``{x}))" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
232 |
unfolding is_recfun_def |
60770 | 233 |
txt\<open>replace f only on the left-hand side\<close> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
234 |
apply (erule_tac P = "\<lambda>x. t(x) = u" for t u in ssubst) |
46820 | 235 |
apply (simp add: underI) |
13165 | 236 |
done |
237 |
||
238 |
lemma is_recfun_equal [rule_format]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
239 |
"\<lbrakk>wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,b,H,g)\<rbrakk> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
240 |
\<Longrightarrow> \<langle>x,a\<rangle>:r \<longrightarrow> \<langle>x,b\<rangle>:r \<longrightarrow> f`x=g`x" |
13784 | 241 |
apply (frule_tac f = f in is_recfun_type) |
242 |
apply (frule_tac f = g in is_recfun_type) |
|
13165 | 243 |
apply (simp add: is_recfun_def) |
244 |
apply (erule_tac a=x in wf_induct) |
|
245 |
apply (intro impI) |
|
246 |
apply (elim ssubst) |
|
247 |
apply (simp (no_asm_simp) add: vimage_singleton_iff restrict_def) |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
248 |
apply (rule_tac t = "\<lambda>z. H (x, z)" for x in subst_context) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
249 |
apply (subgoal_tac "\<forall>y\<in>r-``{x}. \<forall>z. \<langle>y,z\<rangle>:f \<longleftrightarrow> \<langle>y,z\<rangle>:g") |
13165 | 250 |
apply (blast dest: transD) |
251 |
apply (simp add: apply_iff) |
|
252 |
apply (blast dest: transD intro: sym) |
|
253 |
done |
|
254 |
||
255 |
lemma is_recfun_cut: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
256 |
"\<lbrakk>wf(r); trans(r); |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
257 |
is_recfun(r,a,H,f); is_recfun(r,b,H,g); \<langle>b,a\<rangle>:r\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
258 |
\<Longrightarrow> restrict(f, r-``{b}) = g" |
13784 | 259 |
apply (frule_tac f = f in is_recfun_type) |
13165 | 260 |
apply (rule fun_extension) |
261 |
apply (blast dest: transD intro: restrict_type2) |
|
262 |
apply (erule is_recfun_type, simp) |
|
263 |
apply (blast dest: transD intro: is_recfun_equal) |
|
264 |
done |
|
265 |
||
60770 | 266 |
subsection\<open>Recursion: Main Existence Lemma\<close> |
435 | 267 |
|
13165 | 268 |
lemma is_recfun_functional: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
269 |
"\<lbrakk>wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,a,H,g)\<rbrakk> \<Longrightarrow> f=g" |
13165 | 270 |
by (blast intro: fun_extension is_recfun_type is_recfun_equal) |
271 |
||
13248 | 272 |
lemma the_recfun_eq: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
273 |
"\<lbrakk>is_recfun(r,a,H,f); wf(r); trans(r)\<rbrakk> \<Longrightarrow> the_recfun(r,a,H) = f" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
274 |
unfolding the_recfun_def |
13248 | 275 |
apply (blast intro: is_recfun_functional) |
276 |
done |
|
277 |
||
13165 | 278 |
(*If some f satisfies is_recfun(r,a,H,-) then so does the_recfun(r,a,H) *) |
279 |
lemma is_the_recfun: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
280 |
"\<lbrakk>is_recfun(r,a,H,f); wf(r); trans(r)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
281 |
\<Longrightarrow> is_recfun(r, a, H, the_recfun(r,a,H))" |
13248 | 282 |
by (simp add: the_recfun_eq) |
13165 | 283 |
|
284 |
lemma unfold_the_recfun: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
285 |
"\<lbrakk>wf(r); trans(r)\<rbrakk> \<Longrightarrow> is_recfun(r, a, H, the_recfun(r,a,H))" |
13165 | 286 |
apply (rule_tac a=a in wf_induct, assumption) |
46820 | 287 |
apply (rename_tac a1) |
288 |
apply (rule_tac f = "\<lambda>y\<in>r-``{a1}. wftrec (r,y,H)" in is_the_recfun) |
|
13165 | 289 |
apply typecheck |
76217 | 290 |
unfolding is_recfun_def wftrec_def |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
61798
diff
changeset
|
291 |
\<comment> \<open>Applying the substitution: must keep the quantified assumption!\<close> |
46820 | 292 |
apply (rule lam_cong [OF refl]) |
293 |
apply (drule underD) |
|
13165 | 294 |
apply (fold is_recfun_def) |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
295 |
apply (rule_tac t = "\<lambda>z. H(x, z)" for x in subst_context) |
13165 | 296 |
apply (rule fun_extension) |
297 |
apply (blast intro: is_recfun_type) |
|
298 |
apply (rule lam_type [THEN restrict_type2]) |
|
299 |
apply blast |
|
300 |
apply (blast dest: transD) |
|
46993 | 301 |
apply atomize |
13165 | 302 |
apply (frule spec [THEN mp], assumption) |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
303 |
apply (subgoal_tac "\<langle>xa,a1\<rangle> \<in> r") |
13784 | 304 |
apply (drule_tac x1 = xa in spec [THEN mp], assumption) |
46820 | 305 |
apply (simp add: vimage_singleton_iff |
13165 | 306 |
apply_recfun is_recfun_cut) |
307 |
apply (blast dest: transD) |
|
308 |
done |
|
309 |
||
310 |
||
69593 | 311 |
subsection\<open>Unfolding \<^term>\<open>wftrec(r,a,H)\<close>\<close> |
13165 | 312 |
|
313 |
lemma the_recfun_cut: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
314 |
"\<lbrakk>wf(r); trans(r); \<langle>b,a\<rangle>:r\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
315 |
\<Longrightarrow> restrict(the_recfun(r,a,H), r-``{b}) = the_recfun(r,b,H)" |
13269 | 316 |
by (blast intro: is_recfun_cut unfold_the_recfun) |
0 | 317 |
|
13165 | 318 |
(*NOT SUITABLE FOR REWRITING: it is recursive!*) |
319 |
lemma wftrec: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
320 |
"\<lbrakk>wf(r); trans(r)\<rbrakk> \<Longrightarrow> |
46820 | 321 |
wftrec(r,a,H) = H(a, \<lambda>x\<in>r-``{a}. wftrec(r,x,H))" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
322 |
unfolding wftrec_def |
13165 | 323 |
apply (subst unfold_the_recfun [unfolded is_recfun_def]) |
324 |
apply (simp_all add: vimage_singleton_iff [THEN iff_sym] the_recfun_cut) |
|
325 |
done |
|
326 |
||
13634 | 327 |
|
69593 | 328 |
subsubsection\<open>Removal of the Premise \<^term>\<open>trans(r)\<close>\<close> |
13165 | 329 |
|
330 |
(*NOT SUITABLE FOR REWRITING: it is recursive!*) |
|
331 |
lemma wfrec: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
332 |
"wf(r) \<Longrightarrow> wfrec(r,a,H) = H(a, \<lambda>x\<in>r-``{a}. wfrec(r,x,H))" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
333 |
unfolding wfrec_def |
13165 | 334 |
apply (erule wf_trancl [THEN wftrec, THEN ssubst]) |
335 |
apply (rule trans_trancl) |
|
336 |
apply (rule vimage_pair_mono [THEN restrict_lam_eq, THEN subst_context]) |
|
337 |
apply (erule r_into_trancl) |
|
338 |
apply (rule subset_refl) |
|
339 |
done |
|
0 | 340 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
341 |
(*This form avoids giant explosions in proofs. NOTE USE OF \<equiv> *) |
13165 | 342 |
lemma def_wfrec: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
343 |
"\<lbrakk>\<And>x. h(x)\<equiv>wfrec(r,x,H); wf(r)\<rbrakk> \<Longrightarrow> |
46820 | 344 |
h(a) = H(a, \<lambda>x\<in>r-``{a}. h(x))" |
13165 | 345 |
apply simp |
46820 | 346 |
apply (elim wfrec) |
13165 | 347 |
done |
348 |
||
349 |
lemma wfrec_type: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
350 |
"\<lbrakk>wf(r); a \<in> A; field(r)<=A; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
351 |
\<And>x u. \<lbrakk>x \<in> A; u \<in> Pi(r-``{x}, B)\<rbrakk> \<Longrightarrow> H(x,u) \<in> B(x) |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
352 |
\<rbrakk> \<Longrightarrow> wfrec(r,a,H) \<in> B(a)" |
13784 | 353 |
apply (rule_tac a = a in wf_induct2, assumption+) |
13165 | 354 |
apply (subst wfrec, assumption) |
46820 | 355 |
apply (simp add: lam_type underD) |
13165 | 356 |
done |
357 |
||
358 |
||
359 |
lemma wfrec_on: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71085
diff
changeset
|
360 |
"\<lbrakk>wf[A](r); a \<in> A\<rbrakk> \<Longrightarrow> |
46820 | 361 |
wfrec[A](r,a,H) = H(a, \<lambda>x\<in>(r-``{a}) \<inter> A. wfrec[A](r,x,H))" |
76217 | 362 |
unfolding wf_on_def wfrec_on_def |
13165 | 363 |
apply (erule wfrec [THEN trans]) |
76419 | 364 |
apply (simp add: vimage_Int_square) |
13165 | 365 |
done |
0 | 366 |
|
60770 | 367 |
text\<open>Minimal-element characterization of well-foundedness\<close> |
76419 | 368 |
lemma wf_eq_minimal: "wf(r) \<longleftrightarrow> (\<forall>Q x. x \<in> Q \<longrightarrow> (\<exists>z\<in>Q. \<forall>y. \<langle>y,z\<rangle>:r \<longrightarrow> y\<notin>Q))" |
369 |
unfolding wf_def by blast |
|
13634 | 370 |
|
0 | 371 |
end |