author | lcp |
Fri, 19 Aug 1994 11:02:45 +0200 | |
changeset 115 | 0ec63df3ae04 |
child 128 | 89669c58e506 |
permissions | -rw-r--r-- |
115
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
1 |
(* Title: HOL/Fun |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
2 |
ID: $Id$ |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
3 |
Author: Tobias Nipkow, Cambridge University Computer Laboratory |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
4 |
Copyright 1993 University of Cambridge |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
5 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
6 |
Lemmas about functions. |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
7 |
*) |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
8 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
9 |
goal Fun.thy "(f = g) = (!x. f(x)=g(x))"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
10 |
by (rtac iffI 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
11 |
by(asm_simp_tac HOL_ss 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
12 |
by(rtac ext 1 THEN asm_simp_tac HOL_ss 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
13 |
val expand_fun_eq = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
14 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
15 |
val prems = goal Fun.thy |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
16 |
"[| f(x)=u; !!x. P(x) ==> g(f(x)) = x; P(x) |] ==> x=g(u)"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
17 |
by (rtac (arg_cong RS box_equals) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
18 |
by (REPEAT (resolve_tac (prems@[refl]) 1)); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
19 |
val apply_inverse = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
20 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
21 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
22 |
(*** Range of a function ***) |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
23 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
24 |
(*Frequently b does not have the syntactic form of f(x).*) |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
25 |
val [prem] = goalw Fun.thy [range_def] "b=f(x) ==> b : range(f)"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
26 |
by (EVERY1 [rtac CollectI, rtac exI, rtac prem]); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
27 |
val range_eqI = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
28 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
29 |
val rangeI = refl RS range_eqI; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
30 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
31 |
val [major,minor] = goalw Fun.thy [range_def] |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
32 |
"[| b : range(%x.f(x)); !!x. b=f(x) ==> P |] ==> P"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
33 |
by (rtac (major RS CollectD RS exE) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
34 |
by (etac minor 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
35 |
val rangeE = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
36 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
37 |
(*** Image of a set under a function ***) |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
38 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
39 |
val prems = goalw Fun.thy [image_def] "[| b=f(x); x:A |] ==> b : f``A"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
40 |
by (REPEAT (resolve_tac (prems @ [CollectI,bexI,prem]) 1)); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
41 |
val image_eqI = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
42 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
43 |
val imageI = refl RS image_eqI; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
44 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
45 |
(*The eta-expansion gives variable-name preservation.*) |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
46 |
val major::prems = goalw Fun.thy [image_def] |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
47 |
"[| b : (%x.f(x))``A; !!x.[| b=f(x); x:A |] ==> P |] ==> P"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
48 |
by (rtac (major RS CollectD RS bexE) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
49 |
by (REPEAT (ares_tac prems 1)); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
50 |
val imageE = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
51 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
52 |
goalw Fun.thy [o_def] "(f o g)``r = f``(g``r)"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
53 |
by (rtac set_ext 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
54 |
by (fast_tac (HOL_cs addIs [imageI] addSEs [imageE]) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
55 |
val image_compose = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
56 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
57 |
goal Fun.thy "f``(A Un B) = f``A Un f``B"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
58 |
by (rtac set_ext 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
59 |
by (fast_tac (HOL_cs addIs [imageI,UnCI] addSEs [imageE,UnE]) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
60 |
val image_Un = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
61 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
62 |
(*** inj(f): f is a one-to-one function ***) |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
63 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
64 |
val prems = goalw Fun.thy [inj_def] |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
65 |
"[| !! x y. f(x) = f(y) ==> x=y |] ==> inj(f)"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
66 |
by (fast_tac (HOL_cs addIs prems) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
67 |
val injI = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
68 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
69 |
val [major] = goal Fun.thy "(!!x. g(f(x)) = x) ==> inj(f)"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
70 |
by (rtac injI 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
71 |
by (etac (arg_cong RS box_equals) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
72 |
by (rtac major 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
73 |
by (rtac major 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
74 |
val inj_inverseI = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
75 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
76 |
val [major,minor] = goalw Fun.thy [inj_def] |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
77 |
"[| inj(f); f(x) = f(y) |] ==> x=y"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
78 |
by (rtac (major RS spec RS spec RS mp) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
79 |
by (rtac minor 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
80 |
val injD = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
81 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
82 |
(*Useful with the simplifier*) |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
83 |
val [major] = goal Fun.thy "inj(f) ==> (f(x) = f(y)) = (x=y)"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
84 |
by (rtac iffI 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
85 |
by (etac (major RS injD) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
86 |
by (etac arg_cong 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
87 |
val inj_eq = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
88 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
89 |
val [major] = goal Fun.thy "inj(f) ==> (@x.f(x)=f(y)) = y"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
90 |
by (rtac (major RS injD) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
91 |
by (rtac selectI 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
92 |
by (rtac refl 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
93 |
val inj_select = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
94 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
95 |
(*A one-to-one function has an inverse (given using select).*) |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
96 |
val [major] = goalw Fun.thy [Inv_def] "inj(f) ==> Inv(f,f(x)) = x"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
97 |
by (EVERY1 [rtac (major RS inj_select)]); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
98 |
val Inv_f_f = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
99 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
100 |
(* Useful??? *) |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
101 |
val [oneone,minor] = goal Fun.thy |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
102 |
"[| inj(f); !!y. y: range(f) ==> P(Inv(f,y)) |] ==> P(x)"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
103 |
by (res_inst_tac [("t", "x")] (oneone RS (Inv_f_f RS subst)) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
104 |
by (rtac (rangeI RS minor) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
105 |
val inj_transfer = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
106 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
107 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
108 |
(*** inj_onto(f,A): f is one-to-one over A ***) |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
109 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
110 |
val prems = goalw Fun.thy [inj_onto_def] |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
111 |
"(!! x y. [| f(x) = f(y); x:A; y:A |] ==> x=y) ==> inj_onto(f,A)"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
112 |
by (fast_tac (HOL_cs addIs prems addSIs [ballI]) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
113 |
val inj_ontoI = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
114 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
115 |
val [major] = goal Fun.thy |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
116 |
"(!!x. x:A ==> g(f(x)) = x) ==> inj_onto(f,A)"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
117 |
by (rtac inj_ontoI 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
118 |
by (etac (apply_inverse RS trans) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
119 |
by (REPEAT (eresolve_tac [asm_rl,major] 1)); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
120 |
val inj_onto_inverseI = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
121 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
122 |
val major::prems = goalw Fun.thy [inj_onto_def] |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
123 |
"[| inj_onto(f,A); f(x)=f(y); x:A; y:A |] ==> x=y"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
124 |
by (rtac (major RS bspec RS bspec RS mp) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
125 |
by (REPEAT (resolve_tac prems 1)); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
126 |
val inj_ontoD = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
127 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
128 |
val major::prems = goal Fun.thy |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
129 |
"[| inj_onto(f,A); ~x=y; x:A; y:A |] ==> ~ f(x)=f(y)"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
130 |
by (rtac contrapos 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
131 |
by (etac (major RS inj_ontoD) 2); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
132 |
by (REPEAT (resolve_tac prems 1)); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
133 |
val inj_onto_contraD = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
134 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
135 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
136 |
(*** Lemmas about inj ***) |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
137 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
138 |
val prems = goalw Fun.thy [o_def] |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
139 |
"[| inj(f); inj_onto(g,range(f)) |] ==> inj(g o f)"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
140 |
by (cut_facts_tac prems 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
141 |
by (fast_tac (HOL_cs addIs [injI,rangeI] |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
142 |
addEs [injD,inj_ontoD]) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
143 |
val comp_inj = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
144 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
145 |
val [prem] = goal Fun.thy "inj(f) ==> inj_onto(f,A)"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
146 |
by (fast_tac (HOL_cs addIs [prem RS injD, inj_ontoI]) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
147 |
val inj_imp = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
148 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
149 |
val [prem] = goalw Fun.thy [Inv_def] "y : range(f) ==> f(Inv(f,y)) = y"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
150 |
by (EVERY1 [rtac (prem RS rangeE), rtac selectI, etac sym]); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
151 |
val f_Inv_f = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
152 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
153 |
val prems = goal Fun.thy |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
154 |
"[| Inv(f,x)=Inv(f,y); x: range(f); y: range(f) |] ==> x=y"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
155 |
by (rtac (arg_cong RS box_equals) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
156 |
by (REPEAT (resolve_tac (prems @ [f_Inv_f]) 1)); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
157 |
val Inv_injective = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
158 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
159 |
val prems = goal Fun.thy |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
160 |
"[| inj(f); A<=range(f) |] ==> inj_onto(Inv(f), A)"; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
161 |
by (cut_facts_tac prems 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
162 |
by (fast_tac (HOL_cs addIs [inj_ontoI] |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
163 |
addEs [Inv_injective,injD,subsetD]) 1); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
164 |
val inj_onto_Inv = result(); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
165 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
166 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
167 |
(*** Set reasoning tools ***) |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
168 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
169 |
val set_cs = HOL_cs |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
170 |
addSIs [ballI, subsetI, InterI, INT_I, INT1_I, CollectI, |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
171 |
ComplI, IntI, DiffI, UnCI, insertCI] |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
172 |
addIs [bexI, UnionI, UN_I, UN1_I, imageI, rangeI] |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
173 |
addSEs [bexE, UnionE, UN_E, UN1_E, DiffE, |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
174 |
CollectE, ComplE, IntE, UnE, insertE, imageE, rangeE, emptyE] |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
175 |
addEs [ballE, InterD, InterE, INT_D, INT_E, make_elim INT1_D, |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
176 |
subsetD, subsetCE]; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
177 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
178 |
fun cfast_tac prems = cut_facts_tac prems THEN' fast_tac set_cs; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
179 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
180 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
181 |
fun prover s = prove_goal Fun.thy s (fn _=>[fast_tac set_cs 1]); |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
182 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
183 |
val mem_simps = map prover |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
184 |
[ "(a : A Un B) = (a:A | a:B)", |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
185 |
"(a : A Int B) = (a:A & a:B)", |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
186 |
"(a : Compl(B)) = (~a:B)", |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
187 |
"(a : A-B) = (a:A & ~a:B)", |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
188 |
"(a : {b}) = (a=b)", |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
189 |
"(a : {x.P(x)}) = P(a)" ]; |
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
190 |
|
0ec63df3ae04
HOL/fun.ML: renamed Fun.ML to avoid problems with MLs "fun" keyword
lcp
parents:
diff
changeset
|
191 |
val set_ss = HOL_ss addsimps mem_simps; |