src/ZF/Induct/Acc.thy
author wenzelm
Wed, 13 Jul 2005 16:07:35 +0200
changeset 16813 67140ae50e77
parent 16417 9bc16273c2d4
child 18414 560f89584ada
permissions -rw-r--r--
removed ad-hoc atp_hook, cal_atp; removed depth_of; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
     1
(*  Title:      ZF/Induct/Acc.thy
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     2
    ID:         $Id$
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
     5
*)
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     6
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
     7
header {* The accessible part of a relation *}
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
     8
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 12610
diff changeset
     9
theory Acc imports Main begin
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    10
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    11
text {*
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    12
  Inductive definition of @{text "acc(r)"}; see \cite{paulin-tlca}.
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    13
*}
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    14
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    15
consts
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    16
  acc :: "i => i"
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    17
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    18
inductive
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    19
  domains "acc(r)" \<subseteq> "field(r)"
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    20
  intros
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    21
    vimage:  "[| r-``{a}: Pow(acc(r)); a \<in> field(r) |] ==> a \<in> acc(r)"
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    22
  monos      Pow_mono
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    23
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    24
text {*
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    25
  The introduction rule must require @{prop "a \<in> field(r)"},
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    26
  otherwise @{text "acc(r)"} would be a proper class!
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    27
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    28
  \medskip
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    29
  The intended introduction rule:
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    30
*}
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    31
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    32
lemma accI: "[| !!b. <b,a>:r ==> b \<in> acc(r);  a \<in> field(r) |] ==> a \<in> acc(r)"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    33
  by (blast intro: acc.intros)
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    34
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    35
lemma acc_downward: "[| b \<in> acc(r);  <a,b>: r |] ==> a \<in> acc(r)"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    36
  by (erule acc.cases) blast
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    37
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    38
lemma acc_induct [induct set: acc]:
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    39
    "[| a \<in> acc(r);
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    40
        !!x. [| x \<in> acc(r);  \<forall>y. <y,x>:r --> P(y) |] ==> P(x)
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    41
     |] ==> P(a)"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    42
  by (erule acc.induct) (blast intro: acc.intros)
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    43
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    44
lemma wf_on_acc: "wf[acc(r)](r)"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    45
  apply (rule wf_onI2)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    46
  apply (erule acc_induct)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    47
  apply fast
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    48
  done
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    49
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    50
lemma acc_wfI: "field(r) \<subseteq> acc(r) \<Longrightarrow> wf(r)"
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    51
  by (erule wf_on_acc [THEN wf_on_subset_A, THEN wf_on_field_imp_wf])
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    52
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    53
lemma acc_wfD: "wf(r) ==> field(r) \<subseteq> acc(r)"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    54
  apply (rule subsetI)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    55
  apply (erule wf_induct2, assumption)
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    56
   apply (blast intro: accI)+
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    57
  done
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    58
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    59
lemma wf_acc_iff: "wf(r) <-> field(r) \<subseteq> acc(r)"
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12560
diff changeset
    60
  by (rule iffI, erule acc_wfD, erule acc_wfI)
12560
5820841f21fd converted some ZF/Induct examples to Isar
paulson
parents: 12088
diff changeset
    61
12088
6f463d16cbd0 reorganization of the ZF examples
paulson
parents:
diff changeset
    62
end