src/ZF/ex/Acc.ML
author wenzelm
Tue, 20 Oct 1998 16:26:20 +0200
changeset 5686 1f053d05f571
parent 5618 721671c68324
child 9491 1a36151ee2fc
permissions -rw-r--r--
Symtab.foldl;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 804
diff changeset
     1
(*  Title:      ZF/ex/acc
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 804
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Inductive definition of acc(r)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
See Ch. Paulin-Mohring, Inductive Definitions in the System Coq.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
Research Report 92-49, LIP, ENS Lyon.  Dec 1992.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    12
open Acc;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
434
89d45187f04d Various updates and tidying
lcp
parents: 279
diff changeset
    14
(*The introduction rule must require  a:field(r)  
89d45187f04d Various updates and tidying
lcp
parents: 279
diff changeset
    15
  Otherwise acc(r) would be a proper class!    *)
89d45187f04d Various updates and tidying
lcp
parents: 279
diff changeset
    16
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    17
(*The intended introduction rule*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    18
val prems = goal Acc.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    19
    "[| !!b. <b,a>:r ==> b: acc(r);  a: field(r) |] ==> a: acc(r)";
5618
paulson
parents: 5137
diff changeset
    20
by (fast_tac (claset() addIs prems@acc.intrs) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 515
diff changeset
    21
qed "accI";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    22
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
    23
Goal "[| b: acc(r);  <a,b>: r |] ==> a: acc(r)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    24
by (etac acc.elim 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    25
by (Fast_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 515
diff changeset
    26
qed "acc_downward";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    27
434
89d45187f04d Various updates and tidying
lcp
parents: 279
diff changeset
    28
val [major,indhyp] = goal Acc.thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 804
diff changeset
    29
    "[| a : acc(r);                                             \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 804
diff changeset
    30
\       !!x. [| x: acc(r);  ALL y. <y,x>:r --> P(y) |] ==> P(x) \
434
89d45187f04d Various updates and tidying
lcp
parents: 279
diff changeset
    31
\    |] ==> P(a)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    32
by (rtac (major RS acc.induct) 1);
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 434
diff changeset
    33
by (rtac indhyp 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    34
by (Fast_tac 2);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    35
by (resolve_tac acc.intrs 1);
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 434
diff changeset
    36
by (assume_tac 2);
804
02430d273ebf ran expandshort script
lcp
parents: 782
diff changeset
    37
by (etac (Collect_subset RS Pow_mono RS subsetD) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 515
diff changeset
    38
qed "acc_induct";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
    40
Goal "wf[acc(r)](r)";
438
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 434
diff changeset
    41
by (rtac wf_onI2 1);
52e8393ccd77 minor tidying up (ordered rewriting in Integ.ML)
lcp
parents: 434
diff changeset
    42
by (etac acc_induct 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    43
by (Fast_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 515
diff changeset
    44
qed "wf_on_acc";
434
89d45187f04d Various updates and tidying
lcp
parents: 279
diff changeset
    45
89d45187f04d Various updates and tidying
lcp
parents: 279
diff changeset
    46
(* field(r) <= acc(r) ==> wf(r) *)
89d45187f04d Various updates and tidying
lcp
parents: 279
diff changeset
    47
val acc_wfI = wf_on_acc RS wf_on_subset_A RS wf_on_field_imp_wf;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    49
val [major] = goal Acc.thy "wf(r) ==> field(r) <= acc(r)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
by (rtac subsetI 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    51
by (etac (major RS wf_induct2) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    52
by (rtac subset_refl 1);
804
02430d273ebf ran expandshort script
lcp
parents: 782
diff changeset
    53
by (rtac accI 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
by (assume_tac 2);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    55
by (Fast_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 515
diff changeset
    56
qed "acc_wfD";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    57
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
    58
Goal "wf(r) <-> field(r) <= acc(r)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    59
by (EVERY1 [rtac iffI, etac acc_wfD, etac acc_wfI]);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 515
diff changeset
    60
qed "wf_acc_iff";