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