src/HOL/Relation.ML
author paulson
Fri, 28 Jun 1996 15:29:05 +0200
changeset 1842 a9c36056d320
parent 1786 8a31d85d27b8
child 1985 84cf16192e03
permissions -rw-r--r--
Removed the unused rel_eq_cs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1454
diff changeset
     1
(*  Title:      Relation.ML
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1454
diff changeset
     3
    Authors:    Riccardo Mattolini, Dip. Sistemi e Informatica
5d7a7e439cec expanded tabs
clasohm
parents: 1454
diff changeset
     4
                Lawrence C Paulson, Cambridge University Computer Laboratory
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     5
    Copyright   1994 Universita' di Firenze
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     6
    Copyright   1993  University of Cambridge
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     7
*)
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     8
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
     9
val RSLIST = curry (op MRS);
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    10
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    11
open Relation;
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    12
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    13
(** Identity relation **)
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    14
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    15
goalw Relation.thy [id_def] "(a,a) : id";  
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    16
by (rtac CollectI 1);
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    17
by (rtac exI 1);
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    18
by (rtac refl 1);
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    19
qed "idI";
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    20
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    21
val major::prems = goalw Relation.thy [id_def]
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    22
    "[| p: id;  !!x.[| p = (x,x) |] ==> P  \
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    23
\    |] ==>  P";  
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    24
by (rtac (major RS CollectE) 1);
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    25
by (etac exE 1);
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    26
by (eresolve_tac prems 1);
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    27
qed "idE";
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    28
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    29
goalw Relation.thy [id_def] "(a,b):id = (a=b)";
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
    30
by (Fast_tac 1);
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    31
qed "pair_in_id_conv";
1694
3452958f85a8 Added R_O_id and id_O_R
nipkow
parents: 1642
diff changeset
    32
Addsimps [pair_in_id_conv];
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    33
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    34
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    35
(** Composition of two relations **)
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    36
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    37
val prems = goalw Relation.thy [comp_def]
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    38
    "[| (a,b):s; (b,c):r |] ==> (a,c) : r O s";
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
    39
by (fast_tac (!claset addIs prems) 1);
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    40
qed "compI";
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    41
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    42
(*proof requires higher-level assumptions or a delaying of hyp_subst_tac*)
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    43
val prems = goalw Relation.thy [comp_def]
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    44
    "[| xz : r O s;  \
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    45
\       !!x y z. [| xz = (x,z);  (x,y):s;  (y,z):r |] ==> P \
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    46
\    |] ==> P";
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    47
by (cut_facts_tac prems 1);
1454
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1264
diff changeset
    48
by (REPEAT (eresolve_tac [CollectE, splitE, exE, conjE] 1 ORELSE ares_tac prems 1));
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    49
qed "compE";
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    50
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    51
val prems = goal Relation.thy
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    52
    "[| (a,c) : r O s;  \
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    53
\       !!y. [| (a,y):s;  (y,c):r |] ==> P \
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    54
\    |] ==> P";
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    55
by (rtac compE 1);
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    56
by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [Pair_inject,ssubst] 1));
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    57
qed "compEpair";
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    58
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
    59
AddIs [compI, idI];
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
    60
AddSEs [compE, idE];
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
    61
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    62
val comp_cs = prod_cs addIs [compI, idI] addSEs [compE, idE];
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    63
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    64
goal Relation.thy "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)";
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
    65
by (Fast_tac 1);
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    66
qed "comp_mono";
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    67
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    68
goal Relation.thy
1642
21db0cf9a1a4 Using new "Times" infix
paulson
parents: 1605
diff changeset
    69
    "!!r s. [| s <= A Times B;  r <= B Times C |] ==> \
21db0cf9a1a4 Using new "Times" infix
paulson
parents: 1605
diff changeset
    70
\           (r O s) <= A Times C";
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
    71
by (Fast_tac 1);
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    72
qed "comp_subset_Sigma";
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    73
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    74
(** Natural deduction for trans(r) **)
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    75
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    76
val prems = goalw Relation.thy [trans_def]
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    77
    "(!! x y z. [| (x,y):r;  (y,z):r |] ==> (x,z):r) ==> trans(r)";
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    78
by (REPEAT (ares_tac (prems@[allI,impI]) 1));
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    79
qed "transI";
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    80
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    81
val major::prems = goalw Relation.thy [trans_def]
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    82
    "[| trans(r);  (a,b):r;  (b,c):r |] ==> (a,c):r";
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    83
by (cut_facts_tac [major] 1);
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
    84
by (fast_tac (!claset addIs prems) 1);
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    85
qed "transD";
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    86
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    87
(** Natural deduction for converse(r) **)
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    88
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    89
goalw Relation.thy [converse_def] "!!a b r. (a,b):r ==> (b,a):converse(r)";
1264
3eb91524b938 added local simpsets; removed IOA from 'make test'
clasohm
parents: 1128
diff changeset
    90
by (Simp_tac 1);
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    91
qed "converseI";
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    92
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    93
goalw Relation.thy [converse_def] "!!a b r. (a,b) : converse(r) ==> (b,a) : r";
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
    94
by (Fast_tac 1);
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    95
qed "converseD";
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    96
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    97
qed_goalw "converseE" Relation.thy [converse_def]
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    98
    "[| yx : converse(r);  \
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
    99
\       !!x y. [| yx=(y,x);  (x,y):r |] ==> P \
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   100
\    |] ==> P"
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   101
 (fn [major,minor]=>
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   102
  [ (rtac (major RS CollectE) 1),
1454
d0266c81a85e Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents: 1264
diff changeset
   103
    (REPEAT (eresolve_tac [splitE, bexE,exE, conjE, minor] 1)),
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   104
    (assume_tac 1) ]);
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   105
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
   106
AddSIs [converseI];
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
   107
AddSEs [converseD,converseE];
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
   108
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
   109
val converse_cs = comp_cs addSIs [converseI]
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1454
diff changeset
   110
                          addSEs [converseD,converseE];
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   111
1605
248e1e125ca0 added converse_converse
nipkow
parents: 1552
diff changeset
   112
goalw Relation.thy [converse_def] "converse(converse R) = R";
1761
29e08d527ba1 Removed equalityI from some proofs (because it is now included
berghofe
parents: 1754
diff changeset
   113
by(Fast_tac 1);
1605
248e1e125ca0 added converse_converse
nipkow
parents: 1552
diff changeset
   114
qed "converse_converse";
248e1e125ca0 added converse_converse
nipkow
parents: 1552
diff changeset
   115
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   116
(** Domain **)
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   117
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   118
qed_goalw "Domain_iff" Relation.thy [Domain_def]
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   119
    "a: Domain(r) = (EX y. (a,y): r)"
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
   120
 (fn _=> [ (Fast_tac 1) ]);
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   121
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   122
qed_goal "DomainI" Relation.thy "!!a b r. (a,b): r ==> a: Domain(r)"
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   123
 (fn _ => [ (etac (exI RS (Domain_iff RS iffD2)) 1) ]);
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   124
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   125
qed_goal "DomainE" Relation.thy
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   126
    "[| a : Domain(r);  !!y. (a,y): r ==> P |] ==> P"
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   127
 (fn prems=>
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   128
  [ (rtac (Domain_iff RS iffD1 RS exE) 1),
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   129
    (REPEAT (ares_tac prems 1)) ]);
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   130
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   131
(** Range **)
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   132
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   133
qed_goalw "RangeI" Relation.thy [Range_def] "!!a b r.(a,b): r ==> b : Range(r)"
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   134
 (fn _ => [ (etac (converseI RS DomainI) 1) ]);
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   135
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   136
qed_goalw "RangeE" Relation.thy [Range_def]
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   137
    "[| b : Range(r);  !!x. (x,b): r ==> P |] ==> P"
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   138
 (fn major::prems=>
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   139
  [ (rtac (major RS DomainE) 1),
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   140
    (resolve_tac prems 1),
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   141
    (etac converseD 1) ]);
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   142
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   143
(*** Image of a set under a relation ***)
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   144
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   145
qed_goalw "Image_iff" Relation.thy [Image_def]
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   146
    "b : r^^A = (? x:A. (x,b):r)"
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
   147
 (fn _ => [ fast_tac (!claset addIs [RangeI]) 1 ]);
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   148
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   149
qed_goal "Image_singleton_iff" Relation.thy
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   150
    "(b : r^^{a}) = ((a,b):r)"
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   151
 (fn _ => [ rtac (Image_iff RS trans) 1,
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
   152
            Fast_tac 1 ]);
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   153
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   154
qed_goalw "ImageI" Relation.thy [Image_def]
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   155
    "!!a b r. [| (a,b): r;  a:A |] ==> b : r^^A"
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   156
 (fn _ => [ (REPEAT (ares_tac [CollectI,RangeI,bexI] 1)),
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   157
            (resolve_tac [conjI ] 1),
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1454
diff changeset
   158
            (rtac RangeI 1),
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
   159
            (REPEAT (Fast_tac 1))]);
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   160
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   161
qed_goalw "ImageE" Relation.thy [Image_def]
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   162
    "[| b: r^^A;  !!x.[| (x,b): r;  x:A |] ==> P |] ==> P"
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   163
 (fn major::prems=>
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   164
  [ (rtac (major RS CollectE) 1),
1786
8a31d85d27b8 best_tac, deepen_tac and safe_tac now also use default claset.
berghofe
parents: 1761
diff changeset
   165
    (safe_tac (!claset)),
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   166
    (etac RangeE 1),
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   167
    (rtac (hd prems) 1),
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   168
    (REPEAT (etac bexE 1 ORELSE ares_tac prems 1)) ]);
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   169
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   170
qed_goal "Image_subset" Relation.thy
1642
21db0cf9a1a4 Using new "Times" infix
paulson
parents: 1605
diff changeset
   171
    "!!A B r. r <= A Times B ==> r^^C <= B"
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   172
 (fn _ =>
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   173
  [ (rtac subsetI 1),
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   174
    (REPEAT (eresolve_tac [asm_rl, ImageE, subsetD RS SigmaD2] 1)) ]);
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   175
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
   176
AddSIs [converseI]; 
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
   177
AddIs  [ImageI, DomainI, RangeI];
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
   178
AddSEs [ImageE, DomainE, RangeE];
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
   179
1128
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   180
val rel_cs = converse_cs addSIs [converseI] 
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   181
                         addIs  [ImageI, DomainI, RangeI]
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   182
                         addSEs [ImageE, DomainE, RangeE];
64b30e3cc6d4 Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff changeset
   183
1694
3452958f85a8 Added R_O_id and id_O_R
nipkow
parents: 1642
diff changeset
   184
goal Relation.thy "R O id = R";
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
   185
by(fast_tac (!claset addIs [set_ext] addbefore (split_all_tac 1)) 1);
1694
3452958f85a8 Added R_O_id and id_O_R
nipkow
parents: 1642
diff changeset
   186
qed "R_O_id";
3452958f85a8 Added R_O_id and id_O_R
nipkow
parents: 1642
diff changeset
   187
3452958f85a8 Added R_O_id and id_O_R
nipkow
parents: 1642
diff changeset
   188
goal Relation.thy "id O R = R";
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1694
diff changeset
   189
by(fast_tac (!claset addIs [set_ext] addbefore (split_all_tac 1)) 1);
1694
3452958f85a8 Added R_O_id and id_O_R
nipkow
parents: 1642
diff changeset
   190
qed "id_O_R";
3452958f85a8 Added R_O_id and id_O_R
nipkow
parents: 1642
diff changeset
   191
3452958f85a8 Added R_O_id and id_O_R
nipkow
parents: 1642
diff changeset
   192
Addsimps [R_O_id,id_O_R];