src/HOL/Hilbert_Choice_lemmas.ML
changeset 11453 1b15f655da2c
child 11454 7514e5e21cb8
equal deleted inserted replaced
11452:f3fbbaeb4fb8 11453:1b15f655da2c
       
     1 (*  Title:      HOL/Hilbert_Choice_lemmas
       
     2     ID:         $Id$
       
     3     Author:     Lawrence C Paulson
       
     4     Copyright   2001  University of Cambridge
       
     5 
       
     6 Lemmas for Hilbert's epsilon-operator and the Axiom of Choice
       
     7 *)
       
     8 
       
     9 
       
    10 (* ML bindings *)
       
    11 val someI = thm "someI";
       
    12 
       
    13 section "SOME: Hilbert's Epsilon-operator";
       
    14 
       
    15 (*Easier to apply than someI if witness ?a comes from an EX-formula*)
       
    16 Goal "EX x. P x ==> P (SOME x. P x)";
       
    17 by (etac exE 1);
       
    18 by (etac someI 1);
       
    19 qed "someI_ex";
       
    20 
       
    21 (*Easier to apply than someI: conclusion has only one occurrence of P*)
       
    22 val prems = Goal "[| P a;  !!x. P x ==> Q x |] ==> Q (SOME x. P x)";
       
    23 by (resolve_tac prems 1);
       
    24 by (rtac someI 1);
       
    25 by (resolve_tac prems 1) ;
       
    26 qed "someI2";
       
    27 
       
    28 (*Easier to apply than someI2 if witness ?a comes from an EX-formula*)
       
    29 val [major,minor] = Goal "[| EX a. P a; !!x. P x ==> Q x |] ==> Q (SOME x. P x)";
       
    30 by (rtac (major RS exE) 1);
       
    31 by (etac someI2 1 THEN etac minor 1);
       
    32 qed "someI2_ex";
       
    33 
       
    34 val prems = Goal "[| P a;  !!x. P x ==> x=a |] ==> (SOME x. P x) = a";
       
    35 by (rtac someI2 1);
       
    36 by (REPEAT (ares_tac prems 1)) ;
       
    37 qed "some_equality";
       
    38 AddIs [some_equality];
       
    39 
       
    40 Goal "[| EX!x. P x; P a |] ==> (SOME x. P x) = a";
       
    41 by (rtac some_equality 1);
       
    42 by  (atac 1);
       
    43 by (etac ex1E 1);
       
    44 by (etac all_dupE 1);
       
    45 by (dtac mp 1);
       
    46 by  (atac 1);
       
    47 by (etac ssubst 1);
       
    48 by (etac allE 1);
       
    49 by (etac mp 1);
       
    50 by (atac 1);
       
    51 qed "some1_equality";
       
    52 
       
    53 Goal "P (SOME x. P x) =  (EX x. P x)";
       
    54 by (rtac iffI 1);
       
    55 by (etac exI 1);
       
    56 by (etac exE 1);
       
    57 by (etac someI 1);
       
    58 qed "some_eq_ex";
       
    59 
       
    60 Goal "(SOME y. y=x) = x";
       
    61 by (rtac some_equality 1);
       
    62 by (rtac refl 1);
       
    63 by (atac 1);
       
    64 qed "some_eq_trivial";
       
    65 
       
    66 Goal "(SOME y. x=y) = x";
       
    67 by (rtac some_equality 1);
       
    68 by (rtac refl 1);
       
    69 by (etac sym 1);
       
    70 qed "some_sym_eq_trivial";
       
    71 
       
    72 
       
    73 AddXEs [someI_ex];
       
    74 AddIs [some_equality];
       
    75 
       
    76 Addsimps [some_eq_trivial, some_sym_eq_trivial];
       
    77 
       
    78 
       
    79 (** "Axiom" of Choice, proved using the description operator **)
       
    80 
       
    81 (*"choice" is now proved in Tools/meson.ML*)
       
    82 
       
    83 Goal "ALL x:S. EX y. Q x y ==> EX f. ALL x:S. Q x (f x)";
       
    84 by (fast_tac (claset() addEs [someI]) 1);
       
    85 qed "bchoice";
       
    86 
       
    87 
       
    88 (**** Function Inverse ****)
       
    89 
       
    90 val inv_def = thm "inv_def";
       
    91 val Inv_def = thm "Inv_def";
       
    92 
       
    93 
       
    94 Goal "inv id = id";
       
    95 by (simp_tac (simpset() addsimps [inv_def,id_def]) 1);
       
    96 qed "inv_id";
       
    97 Addsimps [inv_id];
       
    98 
       
    99 (*A one-to-one function has an inverse.*)
       
   100 Goalw [inv_def] "inj(f) ==> inv f (f x) = x";
       
   101 by (asm_simp_tac (simpset() addsimps [inj_eq]) 1); 
       
   102 qed "inv_f_f";
       
   103 Addsimps [inv_f_f];
       
   104 
       
   105 Goal "[| inj(f);  f x = y |] ==> inv f y = x";
       
   106 by (etac subst 1);
       
   107 by (etac inv_f_f 1);
       
   108 qed "inv_f_eq";
       
   109 
       
   110 Goal "[| inj f; ALL x. f(g x) = x |] ==> inv f = g";
       
   111 by (blast_tac (claset() addIs [ext, inv_f_eq]) 1); 
       
   112 qed "inj_imp_inv_eq";
       
   113 
       
   114 (* Useful??? *)
       
   115 val [oneone,minor] = Goal
       
   116     "[| inj(f); !!y. y: range(f) ==> P(inv f y) |] ==> P(x)";
       
   117 by (res_inst_tac [("t", "x")] (oneone RS (inv_f_f RS subst)) 1);
       
   118 by (rtac (rangeI RS minor) 1);
       
   119 qed "inj_transfer";
       
   120 
       
   121 Goal "(inj f) = (inv f o f = id)";
       
   122 by (asm_simp_tac (simpset() addsimps [o_def, expand_fun_eq]) 1);
       
   123 by (blast_tac (claset() addIs [inj_inverseI, inv_f_f]) 1);
       
   124 qed "inj_iff";
       
   125 
       
   126 Goal "inj f ==> surj (inv f)";
       
   127 by (blast_tac (claset() addIs [surjI, inv_f_f]) 1);
       
   128 qed "inj_imp_surj_inv";
       
   129 
       
   130 Goalw [inv_def] "y : range(f) ==> f(inv f y) = y";
       
   131 by (fast_tac (claset() addIs [someI]) 1);
       
   132 qed "f_inv_f";
       
   133 
       
   134 Goal "surj f ==> f(inv f y) = y";
       
   135 by (asm_simp_tac (simpset() addsimps [f_inv_f, surj_range]) 1);
       
   136 qed "surj_f_inv_f";
       
   137 
       
   138 Goal "[| inv f x = inv f y;  x: range(f);  y: range(f) |] ==> x=y";
       
   139 by (rtac (arg_cong RS box_equals) 1);
       
   140 by (REPEAT (ares_tac [f_inv_f] 1));
       
   141 qed "inv_injective";
       
   142 
       
   143 Goal "A <= range(f) ==> inj_on (inv f) A";
       
   144 by (fast_tac (claset() addIs [inj_onI] 
       
   145                        addEs [inv_injective, injD]) 1);
       
   146 qed "inj_on_inv";
       
   147 
       
   148 Goal "surj f ==> inj (inv f)";
       
   149 by (asm_simp_tac (simpset() addsimps [inj_on_inv, surj_range]) 1);
       
   150 qed "surj_imp_inj_inv";
       
   151 
       
   152 Goal "(surj f) = (f o inv f = id)";
       
   153 by (asm_simp_tac (simpset() addsimps [o_def, expand_fun_eq]) 1);
       
   154 by (blast_tac (claset() addIs [surjI, surj_f_inv_f]) 1);
       
   155 qed "surj_iff";
       
   156 
       
   157 Goal "[| surj f; ALL x. g(f x) = x |] ==> inv f = g";
       
   158 by (rtac ext 1);
       
   159 by (dres_inst_tac [("x","inv f x")] spec 1); 
       
   160 by (asm_full_simp_tac (simpset() addsimps [surj_f_inv_f]) 1); 
       
   161 qed "surj_imp_inv_eq";
       
   162 
       
   163 Goalw [bij_def] "bij f ==> bij (inv f)";
       
   164 by (asm_simp_tac (simpset() addsimps [inj_imp_surj_inv, surj_imp_inj_inv]) 1);
       
   165 qed "bij_imp_bij_inv";
       
   166 
       
   167 val prems = 
       
   168 Goalw [inv_def] "[| !! x. g (f x) = x;  !! y. f (g y) = y |] ==> inv f = g";
       
   169 by (rtac ext 1);
       
   170 by (auto_tac (claset(), simpset() addsimps prems));
       
   171 qed "inv_equality";
       
   172 
       
   173 Goalw [bij_def] "bij f ==> inv (inv f) = f";
       
   174 by (rtac inv_equality 1);
       
   175 by (auto_tac (claset(), simpset() addsimps [surj_f_inv_f]));
       
   176 qed "inv_inv_eq";
       
   177 
       
   178 (** bij(inv f) implies little about f.  Consider f::bool=>bool such that
       
   179     f(True)=f(False)=True.  Then it's consistent with axiom someI that
       
   180     inv(f) could be any function at all, including the identity function.
       
   181     If inv(f)=id then inv(f) is a bijection, but inj(f), surj(f) and
       
   182     inv(inv(f))=f all fail.
       
   183 **)
       
   184 
       
   185 Goalw [bij_def] "[| bij f; bij g |] ==> inv (f o g) = inv g o inv f";
       
   186 by (rtac (inv_equality) 1);
       
   187 by (auto_tac (claset(), simpset() addsimps [surj_f_inv_f]));
       
   188 qed "o_inv_distrib";
       
   189 
       
   190 
       
   191 Goal "surj f ==> f ` (inv f ` A) = A";
       
   192 by (asm_simp_tac (simpset() addsimps [image_eq_UN, surj_f_inv_f]) 1);
       
   193 qed "image_surj_f_inv_f";
       
   194 
       
   195 Goal "inj f ==> (inv f) ` (f ` A) = A";
       
   196 by (asm_simp_tac (simpset() addsimps [image_eq_UN]) 1);
       
   197 qed "image_inv_f_f";
       
   198 
       
   199 Goalw [image_def] "inj(f) ==> inv(f)`(f`X) = X";
       
   200 by Auto_tac;
       
   201 qed "inv_image_comp";
       
   202 
       
   203 Goal "bij f ==> f ` Collect P = {y. P (inv f y)}";
       
   204 by Auto_tac;
       
   205 by (force_tac (claset(), simpset() addsimps [bij_is_inj]) 1);
       
   206 by (blast_tac (claset() addIs [bij_is_surj RS surj_f_inv_f RS sym]) 1);
       
   207 qed "bij_image_Collect_eq";
       
   208 
       
   209 Goal "bij f ==> f -` A = inv f ` A";
       
   210 by Safe_tac;
       
   211 by (asm_simp_tac (simpset() addsimps [bij_is_surj RS surj_f_inv_f]) 2);
       
   212 by (blast_tac (claset() addIs [bij_is_inj RS inv_f_f RS sym]) 1);
       
   213 qed "bij_vimage_eq_inv_image";
       
   214 
       
   215 (*** Inverse ***)
       
   216 
       
   217 Goalw [Inv_def] "f ` A = B ==> (lam x: B. (Inv A f) x) : B funcset A";
       
   218 by (fast_tac (claset() addIs [restrict_in_funcset, someI2]) 1);
       
   219 qed "Inv_funcset";
       
   220 
       
   221 Goal "[| inj_on f A;  x : A |] ==> Inv A f (f x) = x";
       
   222 by (asm_full_simp_tac (simpset() addsimps [Inv_def, inj_on_def]) 1);
       
   223 by (blast_tac (claset() addIs [someI2]) 1); 
       
   224 qed "Inv_f_f";
       
   225 
       
   226 Goal "y : f`A  ==> f (Inv A f y) = y";
       
   227 by (asm_simp_tac (simpset() addsimps [Inv_def]) 1);
       
   228 by (fast_tac (claset() addIs [someI2]) 1);
       
   229 qed "f_Inv_f";
       
   230 
       
   231 Goal "[| Inv A f x = Inv A f y;  x : f`A;  y : f`A |] ==> x=y";
       
   232 by (rtac (arg_cong RS box_equals) 1);
       
   233 by (REPEAT (ares_tac [f_Inv_f] 1));
       
   234 qed "Inv_injective";
       
   235 
       
   236 Goal "B <= f`A ==> inj_on (Inv A f) B";
       
   237 by (rtac inj_onI 1);
       
   238 by (blast_tac (claset() addIs [inj_onI] addDs [Inv_injective, injD]) 1);
       
   239 qed "inj_on_Inv";
       
   240 
       
   241 Goal "[| inj_on f A;  f ` A = B |] \
       
   242 \     ==> compose A (lam y:B. (Inv A f) y) f = (lam x: A. x)";
       
   243 by (asm_simp_tac (simpset() addsimps [compose_def]) 1);
       
   244 by (rtac restrict_ext 1); 
       
   245 by Auto_tac; 
       
   246 by (etac subst 1); 
       
   247 by (asm_full_simp_tac (simpset() addsimps [Inv_f_f]) 1);
       
   248 qed "compose_Inv_id";
       
   249 
       
   250 
       
   251 (**** split ****)
       
   252 
       
   253 (*Can't be added to simpset: loops!*)
       
   254 Goal "(SOME x. P x) = (SOME (a,b). P(a,b))";
       
   255 by (simp_tac (simpset() addsimps [split_Pair_apply]) 1);
       
   256 qed "split_paired_Eps";
       
   257 
       
   258 Goalw [split_def] "Eps (split P) = (SOME xy. P (fst xy) (snd xy))";
       
   259 by (rtac refl 1);
       
   260 qed "Eps_split";
       
   261 
       
   262 Goal "(@(x',y'). x = x' & y = y') = (x,y)";
       
   263 by (Blast_tac 1);
       
   264 qed "Eps_split_eq";
       
   265 Addsimps [Eps_split_eq];