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