src/ZF/WF.thy
author wenzelm
Tue, 16 Apr 2024 17:28:58 +0200
changeset 80126 b73df63e0f52
parent 76419 f20865ad6319
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
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
    17
section\<open>Well-Founded Recursion\<close>
13356
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
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
    22
  wf           :: "i\<Rightarrow>o"  where
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
    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
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
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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
    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
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
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
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
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
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
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
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
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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
    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
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
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
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
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
    53
subsection\<open>Well-Founded Relations\<close>
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
    54
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
    55
subsubsection\<open>Equivalences between \<^term>\<open>wf\<close> and \<^term>\<open>wf_on\<close>\<close>
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
    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
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
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
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
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
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
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
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
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
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
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
bc5dc2392578 new theorems
paulson
parents: 13203
diff changeset
    76
by (simp add: wf_def, fast)
bc5dc2392578 new theorems
paulson
parents: 13203
diff changeset
    77
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
    78
subsubsection\<open>Introduction Rules for \<^term>\<open>wf_on\<close>\<close>
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
    79
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
    80
text\<open>If every non-empty subset of \<^term>\<open>A\<close> has an \<^term>\<open>r\<close>-minimal element
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
    81
   then we have \<^term>\<open>wf[A](r)\<close>.\<close>
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
    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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
    84
 shows         "wf[A](r)"
76217
8655344f1cf6 More obsolete "unfold" calls
paulson <lp15@cam.ac.uk>
parents: 76216
diff changeset
    85
  unfolding wf_on_def wf_def
13165
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
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
    90
text\<open>If \<^term>\<open>r\<close> allows well-founded induction over \<^term>\<open>A\<close>
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
    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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
    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
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
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   104
subsubsection\<open>Well-founded Induction\<close>
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   105
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
   106
text\<open>Consider the least \<^term>\<open>z\<close> in \<^term>\<open>domain(r)\<close> such that
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
   107
  \<^term>\<open>P(z)\<close> does not hold...\<close>
46993
7371429c527d tidying and structured proofs
paulson
parents: 46953
diff changeset
   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
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
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   119
text\<open>The form of this rule is designed to match \<open>wfI\<close>\<close>
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
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]:
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
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
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
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
   144
text\<open>If \<^term>\<open>r\<close> allows well-founded induction
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
   145
   then we have \<^term>\<open>wf(r)\<close>.\<close>
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
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
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   158
subsection\<open>Basic Properties of Well-Founded Relations\<close>
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
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
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
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
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
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
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
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
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
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13252
diff changeset
   175
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
   176
done
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   177
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   181
by (blast dest: wf_on_not_sym)
13165
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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   184
(*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
   185
  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
   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
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   189
       blast)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13252
diff changeset
   190
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
   191
done
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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   194
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   195
text\<open>transitive closure of a WF relation is WF provided
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
   196
  \<^term>\<open>A\<close> is downward closed\<close>
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   199
apply (rule wf_onI2)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   200
apply (frule bspec [THEN mp], assumption+)
13784
b9f6154427a4 tidying (by script)
paulson
parents: 13780
diff changeset
   201
apply (erule_tac a = y in wf_on_induct, assumption)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   202
apply (blast elim: tranclE, blast)
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   203
done
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   206
apply (simp add: wf_iff_wf_on_field)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   207
apply (rule wf_on_subset_A)
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   208
 apply (erule wf_on_trancl)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   209
 apply blast
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   210
apply (rule trancl_type [THEN field_rel_subset])
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   211
done
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   212
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   213
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
   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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   215
45602
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 35762
diff changeset
   216
lemmas underI = vimage_singleton_iff [THEN iffD2]
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 35762
diff changeset
   217
lemmas underD = vimage_singleton_iff [THEN iffD1]
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   218
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13534
diff changeset
   219
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
   220
subsection\<open>The Predicate \<^term>\<open>is_recfun\<close>\<close>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   224
apply (erule ssubst)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   225
apply (rule lamI [THEN rangeI, THEN lam_type], assumption)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   226
done
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   227
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13252
diff changeset
   228
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
   229
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   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
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   235
apply (simp add: underI)
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   236
done
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   237
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
b9f6154427a4 tidying (by script)
paulson
parents: 13780
diff changeset
   241
apply (frule_tac f = f in is_recfun_type)
b9f6154427a4 tidying (by script)
paulson
parents: 13780
diff changeset
   242
apply (frule_tac f = g in is_recfun_type)
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   243
apply (simp add: is_recfun_def)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   244
apply (erule_tac a=x in wf_induct)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   245
apply (intro impI)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   246
apply (elim ssubst)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   250
 apply (blast dest: transD)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   251
apply (simp add: apply_iff)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   252
apply (blast dest: transD intro: sym)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   253
done
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   254
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
b9f6154427a4 tidying (by script)
paulson
parents: 13780
diff changeset
   259
apply (frule_tac f = f in is_recfun_type)
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   260
apply (rule fun_extension)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   261
  apply (blast dest: transD intro: restrict_type2)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   262
 apply (erule is_recfun_type, simp)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   263
apply (blast dest: transD intro: is_recfun_equal)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   264
done
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   265
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   266
subsection\<open>Recursion: Main Existence Lemma\<close>
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 124
diff changeset
   267
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   270
by (blast intro: fun_extension is_recfun_type is_recfun_equal)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   271
13248
ae66c22ed52e new theorems
paulson
parents: 13219
diff changeset
   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
ae66c22ed52e new theorems
paulson
parents: 13219
diff changeset
   275
apply (blast intro: is_recfun_functional)
ae66c22ed52e new theorems
paulson
parents: 13219
diff changeset
   276
done
ae66c22ed52e new theorems
paulson
parents: 13219
diff changeset
   277
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   278
(*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
   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
ae66c22ed52e new theorems
paulson
parents: 13219
diff changeset
   282
by (simp add: the_recfun_eq)
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   283
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   286
apply (rule_tac a=a in wf_induct, assumption)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   287
apply (rename_tac a1)
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   288
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
   289
  apply typecheck
76217
8655344f1cf6 More obsolete "unfold" calls
paulson <lp15@cam.ac.uk>
parents: 76216
diff changeset
   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
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   292
apply (rule lam_cong [OF refl])
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   293
apply (drule underD)
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   296
apply (rule fun_extension)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   297
  apply (blast intro: is_recfun_type)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   298
 apply (rule lam_type [THEN restrict_type2])
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   299
  apply blast
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   300
 apply (blast dest: transD)
46993
7371429c527d tidying and structured proofs
paulson
parents: 46953
diff changeset
   301
apply atomize
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
b9f6154427a4 tidying (by script)
paulson
parents: 13780
diff changeset
   304
 apply (drule_tac x1 = xa in spec [THEN mp], assumption)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   305
apply (simp add: vimage_singleton_iff
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   306
                 apply_recfun is_recfun_cut)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   307
apply (blast dest: transD)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   308
done
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   309
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   310
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
   311
subsection\<open>Unfolding \<^term>\<open>wftrec(r,a,H)\<close>\<close>
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   312
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13252
diff changeset
   316
by (blast intro: is_recfun_cut unfold_the_recfun)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   317
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   318
(*NOT SUITABLE FOR REWRITING: it is recursive!*)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   323
apply (subst unfold_the_recfun [unfolded is_recfun_def])
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   324
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
   325
done
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   326
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13534
diff changeset
   327
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69587
diff changeset
   328
subsubsection\<open>Removal of the Premise \<^term>\<open>trans(r)\<close>\<close>
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   329
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   330
(*NOT SUITABLE FOR REWRITING: it is recursive!*)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   334
apply (erule wf_trancl [THEN wftrec, THEN ssubst])
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   335
 apply (rule trans_trancl)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   336
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
   337
 apply (erule r_into_trancl)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   338
apply (rule subset_refl)
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   339
done
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   344
     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
   345
apply simp
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   346
apply (elim wfrec)
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   347
done
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   348
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
b9f6154427a4 tidying (by script)
paulson
parents: 13780
diff changeset
   353
apply (rule_tac a = a in wf_induct2, assumption+)
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   354
apply (subst wfrec, assumption)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   355
apply (simp add: lam_type underD)
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   356
done
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
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   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
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 45602
diff changeset
   361
         wfrec[A](r,a,H) = H(a, \<lambda>x\<in>(r-``{a}) \<inter> A. wfrec[A](r,x,H))"
76217
8655344f1cf6 More obsolete "unfold" calls
paulson <lp15@cam.ac.uk>
parents: 76216
diff changeset
   362
  unfolding wf_on_def wfrec_on_def
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   363
apply (erule wfrec [THEN trans])
76419
f20865ad6319 tuned proofs;
wenzelm
parents: 76217
diff changeset
   364
apply (simp add: vimage_Int_square)
13165
31d020705aff conversion of equalities and WF to Isar
paulson
parents: 3840
diff changeset
   365
done
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   366
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   367
text\<open>Minimal-element characterization of well-foundedness\<close>
76419
f20865ad6319 tuned proofs;
wenzelm
parents: 76217
diff changeset
   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))"
f20865ad6319 tuned proofs;
wenzelm
parents: 76217
diff changeset
   369
  unfolding wf_def by blast
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13534
diff changeset
   370
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   371
end