author | wenzelm |
Tue, 28 Oct 1997 17:58:35 +0100 | |
changeset 4029 | 22f2d1b17f97 |
parent 2469 | b50b8c0eec01 |
child 4091 | 771b1f6422a8 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/ex/acc |
0 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1993 University of Cambridge |
5 |
||
6 |
Inductive definition of acc(r) |
|
7 |
||
8 |
See Ch. Paulin-Mohring, Inductive Definitions in the System Coq. |
|
9 |
Research Report 92-49, LIP, ENS Lyon. Dec 1992. |
|
10 |
*) |
|
11 |
||
515 | 12 |
open Acc; |
0 | 13 |
|
434 | 14 |
(*The introduction rule must require a:field(r) |
15 |
Otherwise acc(r) would be a proper class! *) |
|
16 |
||
515 | 17 |
(*The intended introduction rule*) |
18 |
val prems = goal Acc.thy |
|
19 |
"[| !!b. <b,a>:r ==> b: acc(r); a: field(r) |] ==> a: acc(r)"; |
|
2469 | 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 | 22 |
|
0 | 23 |
goal Acc.thy "!!a b r. [| b: acc(r); <a,b>: r |] ==> a: acc(r)"; |
515 | 24 |
by (etac acc.elim 1); |
2469 | 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 | 27 |
|
434 | 28 |
val [major,indhyp] = goal Acc.thy |
1461 | 29 |
"[| a : acc(r); \ |
30 |
\ !!x. [| x: acc(r); ALL y. <y,x>:r --> P(y) |] ==> P(x) \ |
|
434 | 31 |
\ |] ==> P(a)"; |
515 | 32 |
by (rtac (major RS acc.induct) 1); |
438 | 33 |
by (rtac indhyp 1); |
2469 | 34 |
by (Fast_tac 2); |
515 | 35 |
by (resolve_tac acc.intrs 1); |
438 | 36 |
by (assume_tac 2); |
804 | 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 | 39 |
|
434 | 40 |
goal Acc.thy "wf[acc(r)](r)"; |
438 | 41 |
by (rtac wf_onI2 1); |
42 |
by (etac acc_induct 1); |
|
2469 | 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 | 45 |
|
46 |
(* field(r) <= acc(r) ==> wf(r) *) |
|
47 |
val acc_wfI = wf_on_acc RS wf_on_subset_A RS wf_on_field_imp_wf; |
|
0 | 48 |
|
49 |
val [major] = goal Acc.thy "wf(r) ==> field(r) <= acc(r)"; |
|
50 |
by (rtac subsetI 1); |
|
51 |
by (etac (major RS wf_induct2) 1); |
|
52 |
by (rtac subset_refl 1); |
|
804 | 53 |
by (rtac accI 1); |
0 | 54 |
by (assume_tac 2); |
2469 | 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 | 57 |
|
58 |
goal Acc.thy "wf(r) <-> field(r) <= acc(r)"; |
|
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"; |