src/HOL/Trancl.ML
author lcp
Tue, 25 Apr 1995 11:14:03 +0200
changeset 1072 0140ff702b23
parent 972 e61b058d58d2
child 1121 485b49694253
permissions -rw-r--r--
updated version
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     1
(*  Title: 	HOL/trancl
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
For trancl.thy.  Theorems about the transitive closure of a relation
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
open Trancl;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
(** Natural deduction for trans(r) **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
val prems = goalw Trancl.thy [trans_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    14
    "(!! x y z. [| (x,y):r;  (y,z):r |] ==> (x,z):r) ==> trans(r)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
by (REPEAT (ares_tac (prems@[allI,impI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
qed "transI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
val major::prems = goalw Trancl.thy [trans_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    19
    "[| trans(r);  (a,b):r;  (b,c):r |] ==> (a,c):r";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
by (cut_facts_tac [major] 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
by (fast_tac (HOL_cs addIs prems) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
qed "transD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
(** Identity relation **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    26
goalw Trancl.thy [id_def] "(a,a) : id";  
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
by (rtac CollectI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
by (rtac exI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
by (rtac refl 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
qed "idI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
val major::prems = goalw Trancl.thy [id_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    33
    "[| p: id;  !!x.[| p = (x,x) |] ==> P  \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
\    |] ==>  P";  
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    35
by (rtac (major RS CollectE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
by (etac exE 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
by (eresolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
qed "idE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    40
goalw Trancl.thy [id_def] "(a,b):id = (a=b)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
by(fast_tac prod_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
qed "pair_in_id_conv";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
(** Composition of two relations **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
val prems = goalw Trancl.thy [comp_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    47
    "[| (a,b):s; (b,c):r |] ==> (a,c) : r O s";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
by (fast_tac (set_cs addIs prems) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
qed "compI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    51
(*proof requires higher-level assumptions or a delaying of hyp_subst_tac*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
val prems = goalw Trancl.thy [comp_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
    "[| xz : r O s;  \
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    54
\       !!x y z. [| xz = (x,z);  (x,y):s;  (y,z):r |] ==> P \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
\    |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
by (cut_facts_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1 ORELSE ares_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
qed "compE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
val prems = goal Trancl.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    61
    "[| (a,c) : r O s;  \
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    62
\       !!y. [| (a,y):s;  (y,c):r |] ==> P \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
\    |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    64
by (rtac compE 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [Pair_inject,ssubst] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
qed "compEpair";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
val comp_cs = prod_cs addIs [compI, idI] addSEs [compE, idE];
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
goal Trancl.thy "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
by (fast_tac comp_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    72
qed "comp_mono";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
goal Trancl.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    75
    "!!r s. [| s <= Sigma A (%x.B);  r <= Sigma B (%x.C) |] ==> \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
\           (r O s) <= Sigma A (%x.C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
by (fast_tac comp_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
qed "comp_subset_Sigma";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
(** The relation rtrancl **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    82
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
goal Trancl.thy "mono(%s. id Un (r O s))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
by (rtac monoI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
by (REPEAT (ares_tac [monoI, subset_refl, comp_mono, Un_mono] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
qed "rtrancl_fun_mono";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
val rtrancl_unfold = rtrancl_fun_mono RS (rtrancl_def RS def_lfp_Tarski);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    89
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    90
(*Reflexivity of rtrancl*)
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    91
goal Trancl.thy "(a,a) : r^*";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    92
by (stac rtrancl_unfold 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    93
by (fast_tac comp_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
qed "rtrancl_refl";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
(*Closure under composition with r*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    97
val prems = goal Trancl.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    98
    "[| (a,b) : r^*;  (b,c) : r |] ==> (a,c) : r^*";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
by (stac rtrancl_unfold 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   100
by (fast_tac (comp_cs addIs prems) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   101
qed "rtrancl_into_rtrancl";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   103
(*rtrancl of r contains r*)
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   104
val [prem] = goal Trancl.thy "[| (a,b) : r |] ==> (a,b) : r^*";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   105
by (rtac (rtrancl_refl RS rtrancl_into_rtrancl) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   106
by (rtac prem 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
qed "r_into_rtrancl";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
(*monotonicity of rtrancl*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   110
goalw Trancl.thy [rtrancl_def] "!!r s. r <= s ==> r^* <= s^*";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   111
by(REPEAT(ares_tac [lfp_mono,Un_mono,comp_mono,subset_refl] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   112
qed "rtrancl_mono";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   113
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   114
(** standard induction rule **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   115
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   116
val major::prems = goal Trancl.thy 
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   117
  "[| (a,b) : r^*; \
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   118
\     !!x. P((x,x)); \
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   119
\     !!x y z.[| P((x,y)); (x,y): r^*; (y,z): r |]  ==>  P((x,z)) |] \
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   120
\  ==>  P((a,b))";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   121
by (rtac ([rtrancl_def, rtrancl_fun_mono, major] MRS def_induct) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   122
by (fast_tac (comp_cs addIs prems) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   123
qed "rtrancl_full_induct";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   124
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   125
(*nice induction rule*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   126
val major::prems = goal Trancl.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   127
    "[| (a::'a,b) : r^*;    \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   128
\       P(a); \
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   129
\	!!y z.[| (a,y) : r^*;  (y,z) : r;  P(y) |] ==> P(z) |]  \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   130
\     ==> P(b)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   131
(*by induction on this formula*)
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   132
by (subgoal_tac "! y. (a::'a,b) = (a,y) --> P(y)" 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   133
(*now solve first subgoal: this formula is sufficient*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   134
by (fast_tac HOL_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   135
(*now do the induction*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   136
by (resolve_tac [major RS rtrancl_full_induct] 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   137
by (fast_tac (comp_cs addIs prems) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   138
by (fast_tac (comp_cs addIs prems) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   139
qed "rtrancl_induct";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   140
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   141
(*transitivity of transitive closure!! -- by induction.*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   142
goal Trancl.thy "trans(r^*)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   143
by (rtac transI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   144
by (res_inst_tac [("b","z")] rtrancl_induct 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   145
by (DEPTH_SOLVE (eresolve_tac [asm_rl, rtrancl_into_rtrancl] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   146
qed "trans_rtrancl";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   147
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   148
(*elimination of rtrancl -- by induction on a special formula*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   149
val major::prems = goal Trancl.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   150
    "[| (a::'a,b) : r^*;  (a = b) ==> P; 	\
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   151
\	!!y.[| (a,y) : r^*; (y,b) : r |] ==> P 	\
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   152
\    |] ==> P";
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   153
by (subgoal_tac "(a::'a) = b  | (? y. (a,y) : r^* & (y,b) : r)" 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   154
by (rtac (major RS rtrancl_induct) 2);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   155
by (fast_tac (set_cs addIs prems) 2);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   156
by (fast_tac (set_cs addIs prems) 2);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   157
by (REPEAT (eresolve_tac ([asm_rl,exE,disjE,conjE]@prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   158
qed "rtranclE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   159
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   160
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   161
(**** The relation trancl ****)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   162
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   163
(** Conversions between trancl and rtrancl **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   164
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   165
val [major] = goalw Trancl.thy [trancl_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   166
    "(a,b) : r^+ ==> (a,b) : r^*";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   167
by (resolve_tac [major RS compEpair] 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   168
by (REPEAT (ares_tac [rtrancl_into_rtrancl] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   169
qed "trancl_into_rtrancl";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   170
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   171
(*r^+ contains r*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   172
val [prem] = goalw Trancl.thy [trancl_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   173
   "[| (a,b) : r |] ==> (a,b) : r^+";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   174
by (REPEAT (ares_tac [prem,compI,rtrancl_refl] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   175
qed "r_into_trancl";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   176
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   177
(*intro rule by definition: from rtrancl and r*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   178
val prems = goalw Trancl.thy [trancl_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   179
    "[| (a,b) : r^*;  (b,c) : r |]   ==>  (a,c) : r^+";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   180
by (REPEAT (resolve_tac ([compI]@prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   181
qed "rtrancl_into_trancl1";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   182
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   183
(*intro rule from r and rtrancl*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   184
val prems = goal Trancl.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   185
    "[| (a,b) : r;  (b,c) : r^* |]   ==>  (a,c) : r^+";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   186
by (resolve_tac (prems RL [rtranclE]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   187
by (etac subst 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   188
by (resolve_tac (prems RL [r_into_trancl]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   189
by (rtac (trans_rtrancl RS transD RS rtrancl_into_trancl1) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   190
by (REPEAT (ares_tac (prems@[r_into_rtrancl]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   191
qed "rtrancl_into_trancl2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   192
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   193
(*elimination of r^+ -- NOT an induction rule*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   194
val major::prems = goal Trancl.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   195
    "[| (a::'a,b) : r^+;  \
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   196
\       (a,b) : r ==> P; \
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   197
\	!!y.[| (a,y) : r^+;  (y,b) : r |] ==> P  \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   198
\    |] ==> P";
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   199
by (subgoal_tac "(a::'a,b) : r | (? y. (a,y) : r^+  &  (y,b) : r)" 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   200
by (REPEAT (eresolve_tac ([asm_rl,disjE,exE,conjE]@prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   201
by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   202
by (etac rtranclE 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   203
by (fast_tac comp_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   204
by (fast_tac (comp_cs addSIs [rtrancl_into_trancl1]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   205
qed "tranclE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   206
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   207
(*Transitivity of r^+.
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   208
  Proved by unfolding since it uses transitivity of rtrancl. *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   209
goalw Trancl.thy [trancl_def] "trans(r^+)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   210
by (rtac transI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   211
by (REPEAT (etac compEpair 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   212
by (rtac (rtrancl_into_rtrancl RS (trans_rtrancl RS transD RS compI)) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   213
by (REPEAT (assume_tac 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   214
qed "trans_trancl";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   215
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   216
val prems = goal Trancl.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   217
    "[| (a,b) : r;  (b,c) : r^+ |]   ==>  (a,c) : r^+";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   218
by (rtac (r_into_trancl RS (trans_trancl RS transD)) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   219
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   220
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   221
qed "trancl_into_trancl2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   222
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   223
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   224
val major::prems = goal Trancl.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
   225
    "[| (a,b) : r^*;  r <= Sigma A (%x.A) |] ==> a=b | a:A";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   226
by (cut_facts_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   227
by (rtac (major RS rtrancl_induct) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   228
by (rtac (refl RS disjI1) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   229
by (fast_tac (comp_cs addSEs [SigmaE2]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   230
qed "trancl_subset_Sigma_lemma";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   231
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   232
goalw Trancl.thy [trancl_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   233
    "!!r. r <= Sigma A (%x.A) ==> trancl(r) <= Sigma A (%x.A)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   234
by (fast_tac (comp_cs addSDs [trancl_subset_Sigma_lemma]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   235
qed "trancl_subset_Sigma";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   236
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   237
val prod_ss = prod_ss addsimps [pair_in_id_conv];