src/ZF/Trancl.ML
author wenzelm
Sun, 15 Oct 2000 19:50:35 +0200
changeset 10220 2a726de6e124
parent 10216 e928bdf62014
child 12089 34e7693271a9
permissions -rw-r--r--
proper symbol markup with \isamath, \isatext; support sub/super scripts:
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
     1
(*  Title:      ZF/trancl.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2493
diff changeset
     6
Transitive closure of a relation
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
     9
Goal "bnd_mono(field(r)*field(r), %s. id(field(r)) Un (r O s))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
by (rtac bnd_monoI 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
by (REPEAT (ares_tac [subset_refl, Un_mono, comp_mono] 2));
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2493
diff changeset
    12
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
    13
qed "rtrancl_bnd_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    15
Goalw [rtrancl_def] "r<=s ==> r^* <= s^*";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
by (rtac lfp_mono 1);
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    17
by (REPEAT (ares_tac [rtrancl_bnd_mono, subset_refl, id_mono,
8318
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
    18
		      comp_mono, Un_mono, field_mono, Sigma_mono] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
    19
qed "rtrancl_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
(* r^* = id(field(r)) Un ( r O r^* )    *)
10216
e928bdf62014 renamed fp_Tarski to fp_unfold
paulson
parents: 9907
diff changeset
    22
bind_thm ("rtrancl_unfold", rtrancl_bnd_mono RS (rtrancl_def RS def_lfp_unfold));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
(** The relation rtrancl **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
8318
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
    26
(*  r^* <= field(r) * field(r)  *)
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
    27
bind_thm ("rtrancl_type", rtrancl_def RS def_lfp_subset);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
(*Reflexivity of rtrancl*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    30
Goal "[| a: field(r) |] ==> <a,a> : r^*";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
by (resolve_tac [rtrancl_unfold RS ssubst] 1);
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    32
by (etac (idI RS UnI1) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
    33
qed "rtrancl_refl";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
(*Closure under composition with r  *)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    36
Goal "[| <a,b> : r^*;  <b,c> : r |] ==> <a,c> : r^*";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
by (resolve_tac [rtrancl_unfold RS ssubst] 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    38
by (rtac (compI RS UnI2) 1);
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    39
by (assume_tac 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    40
by (assume_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
    41
qed "rtrancl_into_rtrancl";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
(*rtrancl of r contains all pairs in r  *)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    44
Goal "<a,b> : r ==> <a,b> : r^*";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
by (resolve_tac [rtrancl_refl RS rtrancl_into_rtrancl] 1);
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    46
by (REPEAT (ares_tac [fieldI1] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
    47
qed "r_into_rtrancl";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    49
(*The premise ensures that r consists entirely of pairs*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    50
Goal "r <= Sigma(A,B) ==> r <= r^*";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
    51
by (blast_tac (claset() addIs [r_into_rtrancl]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
    52
qed "r_subset_rtrancl";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
    54
Goal "field(r^*) = field(r)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
    55
by (blast_tac (claset() addIs [r_into_rtrancl] 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    56
                    addSDs [rtrancl_type RS subsetD]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
    57
qed "rtrancl_field";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    58
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    59
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    60
(** standard induction rule **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    61
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    62
val major::prems = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    63
  "[| <a,b> : r^*; \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
\     !!x. x: field(r) ==> P(<x,x>); \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65
\     !!x y z.[| P(<x,y>); <x,y>: r^*; <y,z>: r |]  ==>  P(<x,z>) |] \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
\  ==>  P(<a,b>)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
by (rtac ([rtrancl_def, rtrancl_bnd_mono, major] MRS def_induct) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
    68
by (blast_tac (claset() addIs prems) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
    69
qed "rtrancl_full_induct";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    70
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
(*nice induction rule.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
  Tried adding the typing hypotheses y,z:field(r), but these
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    73
  caused expensive case splits!*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    74
val major::prems = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    75
  "[| <a,b> : r^*;                                              \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    76
\     P(a);                                                     \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    77
\     !!y z.[| <a,y> : r^*;  <y,z> : r;  P(y) |] ==> P(z)       \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
\  |] ==> P(b)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    79
(*by induction on this formula*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    80
by (subgoal_tac "ALL y. <a,b> = <a,y> --> P(y)" 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    81
(*now solve first subgoal: this formula is sufficient*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    82
by (EVERY1 [etac (spec RS mp), rtac refl]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    83
(*now do the induction*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    84
by (resolve_tac [major RS rtrancl_full_induct] 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
    85
by (ALLGOALS (blast_tac (claset() addIs prems)));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
    86
qed "rtrancl_induct";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    87
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    88
(*transitivity of transitive closure!! -- by induction.*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
    89
Goalw [trans_def] "trans(r^*)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    90
by (REPEAT (resolve_tac [allI,impI] 1));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
by (eres_inst_tac [("b","z")] rtrancl_induct 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
by (DEPTH_SOLVE (eresolve_tac [asm_rl, rtrancl_into_rtrancl] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
    93
qed "trans_rtrancl";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    94
8318
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
    95
bind_thm ("rtrancl_trans", trans_rtrancl RS transD);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
    96
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    97
(*elimination of rtrancl -- by induction on a special formula*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
    98
val major::prems = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    99
    "[| <a,b> : r^*;  (a=b) ==> P;                       \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   100
\       !!y.[| <a,y> : r^*;   <y,b> : r |] ==> P |]      \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
\    ==> P";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
by (subgoal_tac "a = b  | (EX y. <a,y> : r^* & <y,b> : r)" 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
(*see HOL/trancl*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
by (rtac (major RS rtrancl_induct) 2);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
   105
by (ALLGOALS (fast_tac (claset() addSEs prems)));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
   106
qed "rtranclE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   107
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   108
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   109
(**** The relation trancl ****)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   110
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   111
(*Transitivity of r^+ is proved by transitivity of r^*  *)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
   112
Goalw [trans_def,trancl_def] "trans(r^+)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
   113
by (blast_tac (claset() addIs [rtrancl_into_rtrancl RS 
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2929
diff changeset
   114
			      (trans_rtrancl RS transD RS compI)]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
   115
qed "trans_trancl";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   116
8318
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   117
bind_thm ("trancl_trans", trans_trancl RS transD);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   118
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   119
(** Conversions between trancl and rtrancl **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   120
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   121
Goalw [trancl_def] "<a,b> : r^+ ==> <a,b> : r^*";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
   122
by (blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
   123
qed "trancl_into_rtrancl";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   124
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   125
(*r^+ contains all pairs in r  *)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   126
Goalw [trancl_def] "<a,b> : r ==> <a,b> : r^+";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
   127
by (blast_tac (claset() addSIs [rtrancl_refl]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
   128
qed "r_into_trancl";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   129
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   130
(*The premise ensures that r consists entirely of pairs*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   131
Goal "r <= Sigma(A,B) ==> r <= r^+";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
   132
by (blast_tac (claset() addIs [r_into_trancl]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
   133
qed "r_subset_trancl";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   134
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   135
(*intro rule by definition: from r^* and r  *)
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 8318
diff changeset
   136
Goalw [trancl_def] "[| <a,b> : r^*;  <b,c> : r |]   ==>  <a,c> : r^+";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2929
diff changeset
   137
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
   138
qed "rtrancl_into_trancl1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   139
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   140
(*intro rule from r and r^*  *)
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9211
diff changeset
   141
val prems = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   142
    "[| <a,b> : r;  <b,c> : r^* |]   ==>  <a,c> : r^+";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   143
by (resolve_tac (prems RL [rtrancl_induct]) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   144
by (resolve_tac (prems RL [r_into_trancl]) 1);
8318
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   145
by (etac trancl_trans 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   146
by (etac r_into_trancl 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
   147
qed "rtrancl_into_trancl2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   148
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   149
(*Nice induction rule for trancl*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
   150
val major::prems = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   151
  "[| <a,b> : r^+;                                      \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   152
\     !!y.  [| <a,y> : r |] ==> P(y);                   \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   153
\     !!y z.[| <a,y> : r^+;  <y,z> : r;  P(y) |] ==> P(z)       \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   154
\  |] ==> P(b)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   155
by (rtac (rewrite_rule [trancl_def] major  RS  compEpair) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   156
(*by induction on this formula*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   157
by (subgoal_tac "ALL z. <y,z> : r --> P(z)" 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   158
(*now solve first subgoal: this formula is sufficient*)
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2493
diff changeset
   159
by (Blast_tac 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   160
by (etac rtrancl_induct 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
   161
by (ALLGOALS (fast_tac (claset() addIs (rtrancl_into_trancl1::prems))));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
   162
qed "trancl_induct";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   163
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   164
(*elimination of r^+ -- NOT an induction rule*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
   165
val major::prems = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   166
    "[| <a,b> : r^+;  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   167
\       <a,b> : r ==> P; \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   168
\       !!y.[| <a,y> : r^+; <y,b> : r |] ==> P  \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   169
\    |] ==> P";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   170
by (subgoal_tac "<a,b> : r | (EX y. <a,y> : r^+  &  <y,b> : r)" 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
   171
by (fast_tac (claset() addIs prems) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   172
by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   173
by (etac rtranclE 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
   174
by (ALLGOALS (blast_tac (claset() addIs [rtrancl_into_trancl1])));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
   175
qed "tranclE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   176
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4091
diff changeset
   177
Goalw [trancl_def] "r^+ <= field(r)*field(r)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3016
diff changeset
   178
by (blast_tac (claset() addEs [rtrancl_type RS subsetD RS SigmaE2]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
   179
qed "trancl_type";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   180
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
   181
Goalw [trancl_def] "r<=s ==> r^+ <= s^+";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5137
diff changeset
   182
by (REPEAT (ares_tac [comp_mono, rtrancl_mono] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 435
diff changeset
   183
qed "trancl_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   184
8318
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   185
(** Suggested by Sidi Ould Ehmety **)
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   186
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   187
Goal "(r^*)^* = r^*";
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   188
by (rtac equalityI 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   189
by Auto_tac;
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   190
by (ALLGOALS (forward_tac [impOfSubs rtrancl_type]));
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   191
by (ALLGOALS Clarify_tac);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   192
by (etac r_into_rtrancl 2);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   193
by (etac rtrancl_induct 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   194
by (asm_full_simp_tac (simpset() addsimps [rtrancl_refl, rtrancl_field]) 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   195
by (blast_tac (claset() addIs [rtrancl_trans]) 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   196
qed "rtrancl_idemp";
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   197
Addsimps [rtrancl_idemp];
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   198
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   199
Goal "[| R <= S; S <= R^* |] ==> S^* = R^*";
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   200
by (dtac rtrancl_mono 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   201
by (dtac rtrancl_mono 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   202
by (ALLGOALS Asm_full_simp_tac);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   203
by (Blast_tac 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   204
qed "rtrancl_subset";
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   205
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   206
Goal "[| r<= Sigma(A,B); s<=Sigma(C,D) |] ==> (r^* Un s^*)^* = (r Un s)^*";
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   207
by (rtac rtrancl_subset 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   208
by (blast_tac (claset() addDs [r_subset_rtrancl]) 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   209
by (blast_tac (claset() addIs [rtrancl_mono RS subsetD]) 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   210
qed "rtrancl_Un_rtrancl";
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   211
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   212
(** "converse" laws by Sidi Ould Ehmety **)
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   213
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   214
Goal "<x,y>:converse(r)^* ==> <x,y>:converse(r^*)";
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   215
by (rtac converseI 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   216
by (forward_tac [rtrancl_type RS subsetD] 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   217
by (etac rtrancl_induct 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   218
by (blast_tac (claset() addIs [rtrancl_refl]) 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   219
by (blast_tac (claset() addIs [r_into_rtrancl,rtrancl_trans]) 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   220
qed "rtrancl_converseD";
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   221
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   222
Goal "<x,y>:converse(r^*) ==> <x,y>:converse(r)^*";
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   223
by (dtac converseD 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   224
by (forward_tac [rtrancl_type RS subsetD] 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   225
by (etac rtrancl_induct 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   226
by (blast_tac (claset() addIs [rtrancl_refl]) 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   227
by (blast_tac (claset() addIs [r_into_rtrancl,rtrancl_trans]) 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   228
qed "rtrancl_converseI";
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   229
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   230
Goal "converse(r)^* = converse(r^*)";
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   231
by (safe_tac (claset() addSIs [equalityI]));
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   232
by (forward_tac [rtrancl_type RS subsetD] 1);
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   233
by (safe_tac (claset() addSDs [rtrancl_converseD] addSIs [rtrancl_converseI]));
54d69141a17f new theorems from Sidi Ould Ehmety
paulson
parents: 5321
diff changeset
   234
qed "rtrancl_converse";