author | nipkow |
Tue, 12 Jan 1999 15:48:59 +0100 | |
changeset 6099 | d4866f6ff2f9 |
parent 5321 | f8848433d240 |
child 9907 | 473a6604da94 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/Rel.ML |
435 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
435 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
6 |
Relations in Zermelo-Fraenkel Set Theory |
435 | 7 |
*) |
8 |
||
9 |
open Rel; |
|
10 |
||
11 |
(*** Some properties of relations -- useful? ***) |
|
12 |
||
13 |
(* irreflexivity *) |
|
14 |
||
5321 | 15 |
val prems = Goalw [irrefl_def] |
435 | 16 |
"[| !!x. x:A ==> <x,x> ~: r |] ==> irrefl(A,r)"; |
17 |
by (REPEAT (ares_tac (prems @ [ballI]) 1)); |
|
760 | 18 |
qed "irreflI"; |
435 | 19 |
|
5321 | 20 |
Goalw [irrefl_def] "[| irrefl(A,r); x:A |] ==> <x,x> ~: r"; |
21 |
by (Blast_tac 1); |
|
760 | 22 |
qed "irreflE"; |
435 | 23 |
|
24 |
(* symmetry *) |
|
25 |
||
5321 | 26 |
val prems = Goalw [sym_def] |
435 | 27 |
"[| !!x y.<x,y>: r ==> <y,x>: r |] ==> sym(r)"; |
28 |
by (REPEAT (ares_tac (prems @ [allI,impI]) 1)); |
|
760 | 29 |
qed "symI"; |
435 | 30 |
|
5137 | 31 |
Goalw [sym_def] "[| sym(r); <x,y>: r |] ==> <y,x>: r"; |
3016 | 32 |
by (Blast_tac 1); |
760 | 33 |
qed "symE"; |
435 | 34 |
|
35 |
(* antisymmetry *) |
|
36 |
||
5321 | 37 |
val prems = Goalw [antisym_def] |
435 | 38 |
"[| !!x y.[| <x,y>: r; <y,x>: r |] ==> x=y |] ==> \ |
39 |
\ antisym(r)"; |
|
40 |
by (REPEAT (ares_tac (prems @ [allI,impI]) 1)); |
|
760 | 41 |
qed "antisymI"; |
435 | 42 |
|
5321 | 43 |
Goalw [antisym_def] "[| antisym(r); <x,y>: r; <y,x>: r |] ==> x=y"; |
3016 | 44 |
by (Blast_tac 1); |
760 | 45 |
qed "antisymE"; |
435 | 46 |
|
47 |
(* transitivity *) |
|
48 |
||
5321 | 49 |
Goalw [trans_def] "[| trans(r); <a,b>:r; <b,c>:r |] ==> <a,c>:r"; |
3016 | 50 |
by (Blast_tac 1); |
760 | 51 |
qed "transD"; |
435 | 52 |
|
5067 | 53 |
Goalw [trans_on_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
54 |
"[| trans[A](r); <a,b>:r; <b,c>:r; a:A; b:A; c:A |] ==> <a,c>:r"; |
3016 | 55 |
by (Blast_tac 1); |
760 | 56 |
qed "trans_onD"; |
435 | 57 |