src/ZF/Perm.ML
author paulson
Wed, 25 Nov 1998 15:54:41 +0100
changeset 5971 c5a7a7685826
parent 5529 4a54acae6a15
child 6053 8a1059aa01f0
permissions -rw-r--r--
simplified ensures_UNIV
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
     1
(*  Title:      ZF/Perm.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
735
f99621230c2e moved version of Cantors theorem to ZF/Perm.ML
lcp
parents: 693
diff changeset
     6
The theory underlying permutation groups
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
  -- Composition of relations, the identity relation
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
  -- Injections, surjections, bijections
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
  -- Lemmas for the Schroeder-Bernstein Theorem
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
open Perm;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
(** Surjective function space **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    16
Goalw [surj_def] "f: surj(A,B) ==> f: A->B";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
by (etac CollectD1 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
    18
qed "surj_is_fun";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    20
Goalw [surj_def] "f : Pi(A,B) ==> f: surj(A,range(f))";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    21
by (blast_tac (claset() addIs [apply_equality, range_of_fun, domain_type]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
    22
qed "fun_is_surj";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    24
Goalw [surj_def] "f: surj(A,B) ==> range(f)=B";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    25
by (best_tac (claset() addIs [apply_Pair] addEs [range_type]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
    26
qed "surj_range";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    27
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
    28
(** A function with a right inverse is a surjection **)
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
    29
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
    30
val prems = goalw Perm.thy [surj_def]
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
    31
    "[| f: A->B;  !!y. y:B ==> d(y): A;  !!y. y:B ==> f`d(y) = y \
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
    32
\    |] ==> f: surj(A,B)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    33
by (blast_tac (claset() addIs prems) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
    34
qed "f_imp_surjective";
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
    35
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5284
diff changeset
    36
val prems = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
    37
    "[| !!x. x:A ==> c(x): B;           \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
    38
\       !!y. y:B ==> d(y): A;           \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
    39
\       !!y. y:B ==> c(d(y)) = y        \
3840
e0baea4d485a fixed dots;
wenzelm
parents: 3207
diff changeset
    40
\    |] ==> (lam x:A. c(x)) : surj(A,B)";
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
    41
by (res_inst_tac [("d", "d")] f_imp_surjective 1);
5529
4a54acae6a15 tidying
paulson
parents: 5525
diff changeset
    42
by (ALLGOALS (asm_simp_tac (simpset() addsimps [lam_type]@prems) ));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
    43
qed "lam_surjective";
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
    44
735
f99621230c2e moved version of Cantors theorem to ZF/Perm.ML
lcp
parents: 693
diff changeset
    45
(*Cantor's theorem revisited*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
    46
Goalw [surj_def] "f ~: surj(A,Pow(A))";
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
    47
by Safe_tac;
735
f99621230c2e moved version of Cantors theorem to ZF/Perm.ML
lcp
parents: 693
diff changeset
    48
by (cut_facts_tac [cantor] 1);
f99621230c2e moved version of Cantors theorem to ZF/Perm.ML
lcp
parents: 693
diff changeset
    49
by (fast_tac subset_cs 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
    50
qed "cantor_surj";
735
f99621230c2e moved version of Cantors theorem to ZF/Perm.ML
lcp
parents: 693
diff changeset
    51
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    52
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
(** Injective function space **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    55
Goalw [inj_def] "f: inj(A,B) ==> f: A->B";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    56
by (etac CollectD1 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
    57
qed "inj_is_fun";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    58
1787
9246e236a57f Tidied some proofs
paulson
parents: 1782
diff changeset
    59
(*Good for dealing with sets of pairs, but a bit ugly in use [used in AC]*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
    60
Goalw [inj_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
    61
    "[| <a,b>:f;  <c,b>:f;  f: inj(A,B) |] ==> a=c";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    62
by (REPEAT (eresolve_tac [asm_rl, Pair_mem_PiE, CollectE] 1));
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
    63
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
    64
qed "inj_equality";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    66
Goalw [inj_def] "[| f:inj(A,B);  a:A;  b:A;  f`a=f`b |] ==> a=b";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
    67
by (Blast_tac 1);
826
190974c664a3 inj_apply_equality: new
lcp
parents: 791
diff changeset
    68
val inj_apply_equality = result();
190974c664a3 inj_apply_equality: new
lcp
parents: 791
diff changeset
    69
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 437
diff changeset
    70
(** A function with a left inverse is an injection **)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 437
diff changeset
    71
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    72
Goal "[| f: A->B;  ALL x:A. d(f`x)=x |] ==> f: inj(A,B)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    73
by (asm_simp_tac (simpset() addsimps [inj_def]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    74
by (deepen_tac (claset() addEs [subst_context RS box_equals]) 0 1);
1787
9246e236a57f Tidied some proofs
paulson
parents: 1782
diff changeset
    75
bind_thm ("f_imp_injective", ballI RSN (2,result()));
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 437
diff changeset
    76
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5284
diff changeset
    77
val prems = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
    78
    "[| !!x. x:A ==> c(x): B;           \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
    79
\       !!x. x:A ==> d(c(x)) = x        \
3840
e0baea4d485a fixed dots;
wenzelm
parents: 3207
diff changeset
    80
\    |] ==> (lam x:A. c(x)) : inj(A,B)";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 437
diff changeset
    81
by (res_inst_tac [("d", "d")] f_imp_injective 1);
5529
4a54acae6a15 tidying
paulson
parents: 5525
diff changeset
    82
by (ALLGOALS (asm_simp_tac (simpset() addsimps [lam_type]@prems)));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
    83
qed "lam_injective";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 437
diff changeset
    84
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 437
diff changeset
    85
(** Bijections **)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    86
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    87
Goalw [bij_def] "f: bij(A,B) ==> f: inj(A,B)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    88
by (etac IntD1 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
    89
qed "bij_is_inj";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    90
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    91
Goalw [bij_def] "f: bij(A,B) ==> f: surj(A,B)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
by (etac IntD2 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
    93
qed "bij_is_surj";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    94
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
(* f: bij(A,B) ==> f: A->B *)
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    96
bind_thm ("bij_is_fun", (bij_is_inj RS inj_is_fun));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    97
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
    98
val prems = goalw Perm.thy [bij_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
    99
    "[| !!x. x:A ==> c(x): B;           \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
   100
\       !!y. y:B ==> d(y): A;           \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
   101
\       !!x. x:A ==> d(c(x)) = x;       \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
   102
\       !!y. y:B ==> c(d(y)) = y        \
3840
e0baea4d485a fixed dots;
wenzelm
parents: 3207
diff changeset
   103
\    |] ==> (lam x:A. c(x)) : bij(A,B)";
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   104
by (REPEAT (ares_tac (prems @ [IntI, lam_injective, lam_surjective]) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   105
qed "lam_bijective";
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   106
5284
c77e9dd9bafc Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents: 5268
diff changeset
   107
Goal "(ALL y : x. EX! y'. f(y') = f(y))  \
c77e9dd9bafc Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents: 5268
diff changeset
   108
\     ==> (lam z:{f(y). y:x}. THE y. f(y) = z) : bij({f(y). y:x}, x)";
c77e9dd9bafc Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents: 5268
diff changeset
   109
by (res_inst_tac [("d","f")] lam_bijective 1);
c77e9dd9bafc Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents: 5268
diff changeset
   110
by (auto_tac (claset(),
c77e9dd9bafc Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents: 5268
diff changeset
   111
	      simpset() addsimps [the_equality2]));
c77e9dd9bafc Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents: 5268
diff changeset
   112
qed "RepFun_bijective";
c77e9dd9bafc Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents: 5268
diff changeset
   113
6
8ce8c4d13d4d Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents: 0
diff changeset
   114
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   115
(** Identity function **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   116
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   117
val [prem] = goalw Perm.thy [id_def] "a:A ==> <a,a> : id(A)";  
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   118
by (rtac (prem RS lamI) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   119
qed "idI";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   120
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   121
val major::prems = goalw Perm.thy [id_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   122
    "[| p: id(A);  !!x.[| x:A; p=<x,x> |] ==> P  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   123
\    |] ==>  P";  
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   124
by (rtac (major RS lamE) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   125
by (REPEAT (ares_tac prems 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   126
qed "idE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   127
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   128
Goalw [id_def] "id(A) : A->A";  
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   129
by (rtac lam_type 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   130
by (assume_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   131
qed "id_type";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   132
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   133
Goalw [id_def] "x:A ==> id(A)`x = x";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   134
by (Asm_simp_tac 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   135
qed "id_conv";
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   136
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   137
Addsimps [id_conv];
826
190974c664a3 inj_apply_equality: new
lcp
parents: 791
diff changeset
   138
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   139
val [prem] = goalw Perm.thy [id_def] "A<=B ==> id(A) <= id(B)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   140
by (rtac (prem RS lam_mono) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   141
qed "id_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   142
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5137
diff changeset
   143
Goalw [inj_def,id_def] "A<=B ==> id(A): inj(A,B)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   144
by (REPEAT (ares_tac [CollectI,lam_type] 1));
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   145
by (etac subsetD 1 THEN assume_tac 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   146
by (Simp_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   147
qed "id_subset_inj";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   148
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   149
val id_inj = subset_refl RS id_subset_inj;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   150
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   151
Goalw [id_def,surj_def] "id(A): surj(A,A)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   152
by (blast_tac (claset() addIs [lam_type, beta]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   153
qed "id_surj";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   154
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   155
Goalw [bij_def] "id(A): bij(A,A)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   156
by (blast_tac (claset() addIs [id_inj, id_surj]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   157
qed "id_bij";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   158
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   159
Goalw [id_def] "A <= B <-> id(A) : A->B";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   160
by (fast_tac (claset() addSIs [lam_type] addDs [apply_type] 
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   161
                      addss (simpset())) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   162
qed "subset_iff_id";
517
a9f93400f307 for infinite datatypes with arbitrary index sets
lcp
parents: 502
diff changeset
   163
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   164
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   165
(*** Converse of a function ***)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   166
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   167
Goalw [inj_def] "f: inj(A,B) ==> converse(f) : range(f)->A";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   168
by (asm_simp_tac (simpset() addsimps [Pi_iff, function_def]) 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1791
diff changeset
   169
by (etac CollectE 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   170
by (asm_simp_tac (simpset() addsimps [apply_iff]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   171
by (blast_tac (claset() addDs [fun_is_rel]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   172
qed "inj_converse_fun";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   173
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   174
(** Equations for converse(f) **)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   175
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   176
(*The premises are equivalent to saying that f is injective...*) 
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5242
diff changeset
   177
Goal "[| f: A->B;  converse(f): C->A;  a: A |] ==> converse(f)`(f`a) = a";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   178
by (blast_tac (claset() addIs [apply_Pair, apply_equality, converseI]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   179
qed "left_inverse_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   180
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5242
diff changeset
   181
Goal "[| f: inj(A,B);  a: A |] ==> converse(f)`(f`a) = a";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   182
by (blast_tac (claset() addIs [left_inverse_lemma, inj_converse_fun,
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   183
			      inj_is_fun]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   184
qed "left_inverse";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   185
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   186
val left_inverse_bij = bij_is_inj RS left_inverse;
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   187
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5284
diff changeset
   188
Goal "[| f: A->B;  converse(f): C->A;  b: C |] ==> f`(converse(f)`b) = b";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   189
by (rtac (apply_Pair RS (converseD RS apply_equality)) 1);
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5284
diff changeset
   190
by (REPEAT (assume_tac 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   191
qed "right_inverse_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   192
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   193
(*Should the premises be f:surj(A,B), b:B for symmetry with left_inverse?
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   194
  No: they would not imply that converse(f) was a function! *)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   195
Goal "[| f: inj(A,B);  b: range(f) |] ==> f`(converse(f)`b) = b";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   196
by (rtac right_inverse_lemma 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   197
by (REPEAT (ares_tac [inj_converse_fun,inj_is_fun] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   198
qed "right_inverse";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   199
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   200
(*Cannot add [left_inverse, right_inverse] to default simpset: there are too
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   201
  many ways of expressing sufficient conditions.*)
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   202
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   203
Goal "[| f: bij(A,B);  b: B |] ==> f`(converse(f)`b) = b";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   204
by (fast_tac (claset() addss
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   205
	      (simpset() addsimps [bij_def, right_inverse, surj_range])) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   206
qed "right_inverse_bij";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   207
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   208
(** Converses of injections, surjections, bijections **)
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   209
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   210
Goal "f: inj(A,B) ==> converse(f): inj(range(f), A)";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
   211
by (rtac f_imp_injective 1);
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
   212
by (etac inj_converse_fun 1);
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
   213
by (rtac right_inverse 1);
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   214
by (REPEAT (assume_tac 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   215
qed "inj_converse_inj";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   216
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   217
Goal "f: inj(A,B) ==> converse(f): surj(range(f), A)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   218
by (blast_tac (claset() addIs [f_imp_surjective, inj_converse_fun, left_inverse,
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   219
			      inj_is_fun, range_of_fun RS apply_type]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   220
qed "inj_converse_surj";
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   221
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   222
Goalw [bij_def] "f: bij(A,B) ==> converse(f): bij(B,A)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   223
by (fast_tac (claset() addEs [surj_range RS subst, inj_converse_inj,
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   224
			      inj_converse_surj]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   225
qed "bij_converse_bij";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   226
(*Adding this as an SI seems to cause looping*)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   227
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   228
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   229
(** Composition of two relations **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   230
791
354a56e923ff updated comment;
lcp
parents: 782
diff changeset
   231
(*The inductive definition package could derive these theorems for (r O s)*)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   232
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   233
Goalw [comp_def] "[| <a,b>:s; <b,c>:r |] ==> <a,c> : r O s";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   234
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   235
qed "compI";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   236
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   237
val prems = goalw Perm.thy [comp_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   238
    "[| xz : r O s;  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   239
\       !!x y z. [| xz=<x,z>;  <x,y>:s;  <y,z>:r |] ==> P \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   240
\    |] ==> P";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   241
by (cut_facts_tac prems 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   242
by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1 ORELSE ares_tac prems 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   243
qed "compE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   244
2688
889a1cbd1aca rule_by_tactic no longer standardizes its result
paulson
parents: 2637
diff changeset
   245
bind_thm ("compEpair", 
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   246
    rule_by_tactic (REPEAT_FIRST (etac Pair_inject ORELSE' bound_hyp_subst_tac)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
   247
                    THEN prune_params_tac)
2688
889a1cbd1aca rule_by_tactic no longer standardizes its result
paulson
parents: 2637
diff changeset
   248
        (read_instantiate [("xz","<a,c>")] compE));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   249
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   250
AddSIs [idI];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   251
AddIs  [compI];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   252
AddSEs [compE,idE];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   253
5202
084ceb3844f5 A few new lemmas by Mark Staples
paulson
parents: 5147
diff changeset
   254
Goal "converse(R O S) = converse(S) O converse(R)";
084ceb3844f5 A few new lemmas by Mark Staples
paulson
parents: 5147
diff changeset
   255
by (Blast_tac 1);
084ceb3844f5 A few new lemmas by Mark Staples
paulson
parents: 5147
diff changeset
   256
qed "converse_comp";
084ceb3844f5 A few new lemmas by Mark Staples
paulson
parents: 5147
diff changeset
   257
084ceb3844f5 A few new lemmas by Mark Staples
paulson
parents: 5147
diff changeset
   258
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   259
(** Domain and Range -- see Suppes, section 3.1 **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   260
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   261
(*Boyer et al., Set Theory in First-Order Logic, JAR 2 (1986), 287-327*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   262
Goal "range(r O s) <= range(r)";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   263
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   264
qed "range_comp";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   265
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   266
Goal "domain(r) <= range(s) ==> range(r O s) = range(r)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   267
by (rtac (range_comp RS equalityI) 1);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   268
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   269
qed "range_comp_eq";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   270
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   271
Goal "domain(r O s) <= domain(s)";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   272
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   273
qed "domain_comp";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   274
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   275
Goal "range(s) <= domain(r) ==> domain(r O s) = domain(s)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   276
by (rtac (domain_comp RS equalityI) 1);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   277
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   278
qed "domain_comp_eq";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   279
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   280
Goal "(r O s)``A = r``(s``A)";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   281
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   282
qed "image_comp";
218
be834b9a0c72 ZF/perm/image_comp: new
lcp
parents: 37
diff changeset
   283
be834b9a0c72 ZF/perm/image_comp: new
lcp
parents: 37
diff changeset
   284
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   285
(** Other results **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   286
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   287
Goal "[| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   288
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   289
qed "comp_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   290
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   291
(*composition preserves relations*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   292
Goal "[| s<=A*B;  r<=B*C |] ==> (r O s) <= A*C";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   293
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   294
qed "comp_rel";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   295
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   296
(*associative law for composition*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   297
Goal "(r O s) O t = r O (s O t)";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   298
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   299
qed "comp_assoc";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   300
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   301
(*left identity of composition; provable inclusions are
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   302
        id(A) O r <= r       
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   303
  and   [| r<=A*B; B<=C |] ==> r <= id(C) O r *)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   304
Goal "r<=A*B ==> id(B) O r = r";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   305
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   306
qed "left_comp_id";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   307
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   308
(*right identity of composition; provable inclusions are
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   309
        r O id(A) <= r
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   310
  and   [| r<=A*B; A<=C |] ==> r <= r O id(C) *)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   311
Goal "r<=A*B ==> r O id(A) = r";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   312
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   313
qed "right_comp_id";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   314
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   315
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   316
(** Composition preserves functions, injections, and surjections **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   317
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   318
Goalw [function_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   319
    "[| function(g);  function(f) |] ==> function(f O g)";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   320
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   321
qed "comp_function";
693
b89939545725 ZF/Perm/inj_converse_fun: tidied; removed uses of PiI/E
lcp
parents: 517
diff changeset
   322
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   323
Goal "[| g: A->B;  f: B->C |] ==> (f O g) : A->C";
1787
9246e236a57f Tidied some proofs
paulson
parents: 1782
diff changeset
   324
by (asm_full_simp_tac
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   325
    (simpset() addsimps [Pi_def, comp_function, Pow_iff, comp_rel]
1787
9246e236a57f Tidied some proofs
paulson
parents: 1782
diff changeset
   326
           setloop etac conjE) 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1791
diff changeset
   327
by (stac (range_rel_subset RS domain_comp_eq) 1 THEN assume_tac 2);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   328
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   329
qed "comp_fun";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   330
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   331
Goal "[| g: A->B;  f: B->C;  a:A |] ==> (f O g)`a = f`(g`a)";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   332
by (REPEAT (ares_tac [comp_fun,apply_equality,compI,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
   333
                      apply_Pair,apply_type] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   334
qed "comp_fun_apply";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   335
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   336
Addsimps [comp_fun_apply];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   337
862
ce99db6728ba Proved comp_lam.
lcp
parents: 826
diff changeset
   338
(*Simplifies compositions of lambda-abstractions*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5284
diff changeset
   339
val [prem] = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
   340
    "[| !!x. x:A ==> b(x): B    \
3840
e0baea4d485a fixed dots;
wenzelm
parents: 3207
diff changeset
   341
\    |] ==> (lam y:B. c(y)) O (lam x:A. b(x)) = (lam x:A. c(b(x)))";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
   342
by (rtac fun_extension 1);
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
   343
by (rtac comp_fun 1);
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
   344
by (rtac lam_funtype 2);
862
ce99db6728ba Proved comp_lam.
lcp
parents: 826
diff changeset
   345
by (typechk_tac (prem::ZF_typechecks));
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   346
by (asm_simp_tac (simpset() 
2637
e9b203f854ae reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents: 2493
diff changeset
   347
             setSolver type_auto_tac [lam_type, lam_funtype, prem]) 1);
862
ce99db6728ba Proved comp_lam.
lcp
parents: 826
diff changeset
   348
qed "comp_lam";
ce99db6728ba Proved comp_lam.
lcp
parents: 826
diff changeset
   349
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   350
Goal "[| g: inj(A,B);  f: inj(B,C) |] ==> (f O g) : inj(A,C)";
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   351
by (res_inst_tac [("d", "%y. converse(g) ` (converse(f) ` y)")]
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   352
    f_imp_injective 1);
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   353
by (REPEAT (ares_tac [comp_fun, inj_is_fun] 1));
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   354
by (asm_simp_tac (simpset()  addsimps [left_inverse] 
2637
e9b203f854ae reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents: 2493
diff changeset
   355
                        setSolver type_auto_tac [inj_is_fun, apply_type]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   356
qed "comp_inj";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   357
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   358
Goalw [surj_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   359
    "[| g: surj(A,B);  f: surj(B,C) |] ==> (f O g) : surj(A,C)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   360
by (blast_tac (claset() addSIs [comp_fun,comp_fun_apply]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   361
qed "comp_surj";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   362
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   363
Goalw [bij_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   364
    "[| g: bij(A,B);  f: bij(B,C) |] ==> (f O g) : bij(A,C)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   365
by (blast_tac (claset() addIs [comp_inj,comp_surj]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   366
qed "comp_bij";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   367
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   368
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   369
(** Dual properties of inj and surj -- useful for proofs from
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   370
    D Pastre.  Automatic theorem proving in set theory. 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   371
    Artificial Intelligence, 10:1--27, 1978. **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   372
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   373
Goalw [inj_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   374
    "[| (f O g): inj(A,C);  g: A->B;  f: B->C |] ==> g: inj(A,B)";
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   375
by Safe_tac;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   376
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1));
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   377
by (asm_simp_tac (simpset() ) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   378
qed "comp_mem_injD1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   379
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   380
Goalw [inj_def,surj_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   381
    "[| (f O g): inj(A,C);  g: surj(A,B);  f: B->C |] ==> f: inj(B,C)";
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   382
by Safe_tac;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   383
by (res_inst_tac [("x1", "x")] (bspec RS bexE) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   384
by (eres_inst_tac [("x1", "w")] (bspec RS bexE) 3);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   385
by (REPEAT (assume_tac 1));
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   386
by Safe_tac;
6
8ce8c4d13d4d Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents: 0
diff changeset
   387
by (res_inst_tac [("t", "op `(g)")] subst_context 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   388
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1));
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   389
by (asm_simp_tac (simpset() ) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   390
qed "comp_mem_injD2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   391
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   392
Goalw [surj_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   393
    "[| (f O g): surj(A,C);  g: A->B;  f: B->C |] ==> f: surj(B,C)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   394
by (blast_tac (claset() addSIs [comp_fun_apply RS sym, apply_funtype]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   395
qed "comp_mem_surjD1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   396
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5242
diff changeset
   397
Goal "[| (f O g)`a = c;  g: A->B;  f: B->C;  a:A |] ==> f`(g`a) = c";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   398
by (REPEAT (ares_tac [comp_fun_apply RS sym RS trans] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   399
qed "comp_fun_applyD";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   400
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   401
Goalw [inj_def,surj_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   402
    "[| (f O g): surj(A,C);  g: A->B;  f: inj(B,C) |] ==> g: surj(A,B)";
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   403
by Safe_tac;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   404
by (eres_inst_tac [("x1", "f`y")] (bspec RS bexE) 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   405
by (REPEAT (ares_tac [apply_type] 1 ORELSE dtac comp_fun_applyD 1));
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   406
by (blast_tac (claset() addSIs [apply_funtype]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   407
qed "comp_mem_surjD2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   408
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   409
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   410
(** inverses of composition **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   411
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   412
(*left inverse of composition; one inclusion is
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   413
        f: A->B ==> id(A) <= converse(f) O f *)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   414
Goalw [inj_def] "f: inj(A,B) ==> converse(f) O f = id(A)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   415
by (fast_tac (claset() addIs [apply_Pair] 
1787
9246e236a57f Tidied some proofs
paulson
parents: 1782
diff changeset
   416
                      addEs [domain_type]
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   417
               addss (simpset() addsimps [apply_iff])) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   418
qed "left_comp_inverse";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   419
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   420
(*right inverse of composition; one inclusion is
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
   421
                f: A->B ==> f O converse(f) <= id(B) 
735
f99621230c2e moved version of Cantors theorem to ZF/Perm.ML
lcp
parents: 693
diff changeset
   422
*)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   423
val [prem] = goalw Perm.thy [surj_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   424
    "f: surj(A,B) ==> f O converse(f) = id(B)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   425
val appfD = (prem RS CollectD1) RSN (3,apply_equality2);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   426
by (cut_facts_tac [prem] 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   427
by (rtac equalityI 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   428
by (best_tac (claset() addEs [domain_type, range_type, make_elim appfD]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   429
by (blast_tac (claset() addIs [apply_Pair]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   430
qed "right_comp_inverse";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   431
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   432
(** Proving that a function is a bijection **)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   433
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   434
Goalw [id_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   435
    "[| f: A->B;  g: B->A |] ==> \
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   436
\             f O g = id(B) <-> (ALL y:B. f`(g`y)=y)";
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   437
by Safe_tac;
3840
e0baea4d485a fixed dots;
wenzelm
parents: 3207
diff changeset
   438
by (dres_inst_tac [("t", "%h. h`y ")] subst_context 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   439
by (Asm_full_simp_tac 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   440
by (rtac fun_extension 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   441
by (REPEAT (ares_tac [comp_fun, lam_type] 1));
4477
b3e5857d8d99 New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents: 4152
diff changeset
   442
by Auto_tac;
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   443
qed "comp_eq_id_iff";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   444
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   445
Goalw [bij_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   446
    "[| f: A->B;  g: B->A;  f O g = id(B);  g O f = id(A) \
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   447
\             |] ==> f : bij(A,B)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   448
by (asm_full_simp_tac (simpset() addsimps [comp_eq_id_iff]) 1);
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   449
by (REPEAT (ares_tac [conjI, f_imp_injective, f_imp_surjective] 1
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   450
       ORELSE eresolve_tac [bspec, apply_type] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   451
qed "fg_imp_bijective";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   452
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   453
Goal "[| f: A->A;  f O f = id(A) |] ==> f : bij(A,A)";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   454
by (REPEAT (ares_tac [fg_imp_bijective] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   455
qed "nilpotent_imp_bijective";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 218
diff changeset
   456
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   457
Goal "[| converse(f): B->A;  f: A->B |] ==> f : bij(A,B)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   458
by (asm_simp_tac (simpset() addsimps [fg_imp_bijective, comp_eq_id_iff, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1248
diff changeset
   459
                                  left_inverse_lemma, right_inverse_lemma]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   460
qed "invertible_imp_bijective";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   461
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   462
(** Unions of functions -- cf similar theorems on func.ML **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   463
1709
220dd588bfc9 New lemma inspired by KG
paulson
parents: 1610
diff changeset
   464
(*Theorem by KG, proof by LCP*)
5466
2ea5d254e44e eliminated equals0E
paulson
parents: 5321
diff changeset
   465
Goal "[| f: inj(A,B);  g: inj(C,D);  B Int D = 0 |] \
2ea5d254e44e eliminated equals0E
paulson
parents: 5321
diff changeset
   466
\     ==> (lam a: A Un C. if(a:A, f`a, g`a)) : inj(A Un C, B Un D)";
1709
220dd588bfc9 New lemma inspired by KG
paulson
parents: 1610
diff changeset
   467
by (res_inst_tac [("d","%z. if(z:B, converse(f)`z, converse(g)`z)")]
220dd588bfc9 New lemma inspired by KG
paulson
parents: 1610
diff changeset
   468
        lam_injective 1);
220dd588bfc9 New lemma inspired by KG
paulson
parents: 1610
diff changeset
   469
by (ALLGOALS 
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   470
    (asm_simp_tac (simpset() addsimps [inj_is_fun RS apply_type, left_inverse] 
5116
8eb343ab5748 Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents: 5067
diff changeset
   471
                         setloop (split_tac [split_if] ORELSE' etac UnE))));
5466
2ea5d254e44e eliminated equals0E
paulson
parents: 5321
diff changeset
   472
by (blast_tac (claset() addIs [inj_is_fun RS apply_type]) 1);
1709
220dd588bfc9 New lemma inspired by KG
paulson
parents: 1610
diff changeset
   473
qed "inj_disjoint_Un";
1610
60ab5844fe81 Updated comments
paulson
parents: 1461
diff changeset
   474
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   475
Goalw [surj_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   476
    "[| f: surj(A,B);  g: surj(C,D);  A Int C = 0 |] ==> \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   477
\           (f Un g) : surj(A Un C, B Un D)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   478
by (blast_tac (claset() addIs [fun_disjoint_apply1, fun_disjoint_apply2,
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   479
			      fun_disjoint_Un, trans]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   480
qed "surj_disjoint_Un";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   481
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   482
(*A simple, high-level proof; the version for injections follows from it,
502
77e36960fd9e ZF/Perm.ML/inj_converse_inj, comp_inj: simpler proofs using f_imp_injective
lcp
parents: 484
diff changeset
   483
  using  f:inj(A,B) <-> f:bij(A,range(f))  *)
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5242
diff changeset
   484
Goal "[| f: bij(A,B);  g: bij(C,D);  A Int C = 0;  B Int D = 0 |] ==> \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   485
\           (f Un g) : bij(A Un C, B Un D)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   486
by (rtac invertible_imp_bijective 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1791
diff changeset
   487
by (stac converse_Un 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   488
by (REPEAT (ares_tac [fun_disjoint_Un, bij_is_fun, bij_converse_bij] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   489
qed "bij_disjoint_Un";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   490
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   491
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   492
(** Restrictions as surjections and bijections *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   493
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   494
val prems = goalw Perm.thy [surj_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   495
    "f: Pi(A,B) ==> f: surj(A, f``A)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   496
val rls = apply_equality :: (prems RL [apply_Pair,Pi_type]);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   497
by (fast_tac (claset() addIs rls) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   498
qed "surj_image";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   499
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   500
Goal "[| f: Pi(C,B);  A<=C |] ==> restrict(f,A)``A = f``A";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   501
by (rtac equalityI 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   502
by (SELECT_GOAL (rewtac restrict_def) 2);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   503
by (REPEAT (eresolve_tac [imageE, apply_equality RS subst] 2
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   504
     ORELSE ares_tac [subsetI,lamI,imageI] 2));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   505
by (REPEAT (ares_tac [image_mono,restrict_subset,subset_refl] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   506
qed "restrict_image";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   507
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   508
Goalw [inj_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   509
    "[| f: inj(A,B);  C<=A |] ==> restrict(f,C): inj(C,B)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   510
by (safe_tac (claset() addSEs [restrict_type2]));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   511
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp, subsetD,
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   512
                          box_equals, restrict] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   513
qed "restrict_inj";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   514
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5284
diff changeset
   515
Goal "[| f: Pi(A,B);  C<=A |] ==> restrict(f,C): surj(C, f``C)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   516
by (rtac (restrict_image RS subst) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   517
by (rtac (restrict_type2 RS surj_image) 3);
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5284
diff changeset
   518
by (REPEAT (assume_tac 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   519
qed "restrict_surj";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   520
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   521
Goalw [inj_def,bij_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   522
    "[| f: inj(A,B);  C<=A |] ==> restrict(f,C): bij(C, f``C)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   523
by (blast_tac (claset() addSIs [restrict, restrict_surj]
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   524
		       addIs [box_equals, surj_is_fun]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   525
qed "restrict_bij";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   526
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   527
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   528
(*** Lemmas for Ramsey's Theorem ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   529
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   530
Goalw [inj_def] "[| f: inj(A,B);  B<=D |] ==> f: inj(A,D)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   531
by (blast_tac (claset() addIs [fun_weaken_type]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   532
qed "inj_weaken_type";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   533
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   534
val [major] = goal Perm.thy  
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   535
    "[| f: inj(succ(m), A) |] ==> restrict(f,m) : inj(m, A-{f`m})";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   536
by (rtac (major RS restrict_bij RS bij_is_inj RS inj_weaken_type) 1);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2688
diff changeset
   537
by (Blast_tac 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   538
by (cut_facts_tac [major] 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   539
by (rewtac inj_def);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   540
by (fast_tac (claset() addEs [range_type, mem_irrefl] 
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   541
	              addDs [apply_equality]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   542
qed "inj_succ_restrict";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   543
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4477
diff changeset
   544
Goalw [inj_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   545
    "[| f: inj(A,B);  a~:A;  b~:B |]  ==> \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   546
\         cons(<a,b>,f) : inj(cons(a,A), cons(b,B))";
5525
896f8234b864 improved addbefore and addSbefore
oheimb
parents: 5466
diff changeset
   547
by (force_tac (claset() addIs [apply_type],
896f8234b864 improved addbefore and addSbefore
oheimb
parents: 5466
diff changeset
   548
               simpset() addsimps [fun_extend, fun_extend_apply2,
896f8234b864 improved addbefore and addSbefore
oheimb
parents: 5466
diff changeset
   549
						fun_extend_apply1]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 735
diff changeset
   550
qed "inj_extend";
1787
9246e236a57f Tidied some proofs
paulson
parents: 1782
diff changeset
   551