12560
|
1 |
(* Title: ZF/Induct/Acc.thy
|
12088
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
4 |
Copyright 1994 University of Cambridge
|
12610
|
5 |
*)
|
12088
|
6 |
|
12610
|
7 |
header {* The accessible part of a relation *}
|
12088
|
8 |
|
16417
|
9 |
theory Acc imports Main begin
|
12088
|
10 |
|
12610
|
11 |
text {*
|
|
12 |
Inductive definition of @{text "acc(r)"}; see \cite{paulin-tlca}.
|
|
13 |
*}
|
|
14 |
|
12088
|
15 |
consts
|
12610
|
16 |
acc :: "i => i"
|
12088
|
17 |
|
|
18 |
inductive
|
12610
|
19 |
domains "acc(r)" \<subseteq> "field(r)"
|
12560
|
20 |
intros
|
|
21 |
vimage: "[| r-``{a}: Pow(acc(r)); a \<in> field(r) |] ==> a \<in> acc(r)"
|
12088
|
22 |
monos Pow_mono
|
|
23 |
|
12610
|
24 |
text {*
|
|
25 |
The introduction rule must require @{prop "a \<in> field(r)"},
|
|
26 |
otherwise @{text "acc(r)"} would be a proper class!
|
12560
|
27 |
|
12610
|
28 |
\medskip
|
|
29 |
The intended introduction rule:
|
|
30 |
*}
|
|
31 |
|
12560
|
32 |
lemma accI: "[| !!b. <b,a>:r ==> b \<in> acc(r); a \<in> field(r) |] ==> a \<in> acc(r)"
|
12610
|
33 |
by (blast intro: acc.intros)
|
12560
|
34 |
|
|
35 |
lemma acc_downward: "[| b \<in> acc(r); <a,b>: r |] ==> a \<in> acc(r)"
|
12610
|
36 |
by (erule acc.cases) blast
|
12560
|
37 |
|
18414
|
38 |
lemma acc_induct [consumes 1, case_names vimage, induct set: acc]:
|
12610
|
39 |
"[| a \<in> acc(r);
|
|
40 |
!!x. [| x \<in> acc(r); \<forall>y. <y,x>:r --> P(y) |] ==> P(x)
|
12560
|
41 |
|] ==> P(a)"
|
12610
|
42 |
by (erule acc.induct) (blast intro: acc.intros)
|
12560
|
43 |
|
|
44 |
lemma wf_on_acc: "wf[acc(r)](r)"
|
12610
|
45 |
apply (rule wf_onI2)
|
|
46 |
apply (erule acc_induct)
|
|
47 |
apply fast
|
|
48 |
done
|
12560
|
49 |
|
12610
|
50 |
lemma acc_wfI: "field(r) \<subseteq> acc(r) \<Longrightarrow> wf(r)"
|
|
51 |
by (erule wf_on_acc [THEN wf_on_subset_A, THEN wf_on_field_imp_wf])
|
12560
|
52 |
|
|
53 |
lemma acc_wfD: "wf(r) ==> field(r) \<subseteq> acc(r)"
|
12610
|
54 |
apply (rule subsetI)
|
|
55 |
apply (erule wf_induct2, assumption)
|
|
56 |
apply (blast intro: accI)+
|
|
57 |
done
|
12560
|
58 |
|
|
59 |
lemma wf_acc_iff: "wf(r) <-> field(r) \<subseteq> acc(r)"
|
12610
|
60 |
by (rule iffI, erule acc_wfD, erule acc_wfI)
|
12560
|
61 |
|
12088
|
62 |
end
|