src/HOL/Cardinals/Fun_More_FP.thy
author traytel
Mon, 25 Nov 2013 13:48:00 +0100
changeset 55954 1502a1f707d9
parent 55951 9387251b6a46
child 56361 0d5e831175de
permissions -rw-r--r--
eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
     1 (*  Title:      HOL/Cardinals/Fun_More_FP.thy
     2     Author:     Andrei Popescu, TU Muenchen
     3     Copyright   2012
     4 
     5 More on injections, bijections and inverses (FP).
     6 *)
     7 
     8 header {* More on Injections, Bijections and Inverses (FP) *}
     9 
    10 theory Fun_More_FP
    11 imports Hilbert_Choice
    12 begin
    13 
    14 
    15 text {* This section proves more facts (additional to those in @{text "Fun.thy"},
    16 @{text "Hilbert_Choice.thy"}, and @{text "Finite_Set.thy"}),
    17 mainly concerning injections, bijections, inverses and (numeric) cardinals of
    18 finite sets. *}
    19 
    20 
    21 subsection {* Purely functional properties  *}
    22 
    23 
    24 (*2*)lemma bij_betw_id_iff:
    25 "(A = B) = (bij_betw id A B)"
    26 by(simp add: bij_betw_def)
    27 
    28 
    29 (*2*)lemma bij_betw_byWitness:
    30 assumes LEFT: "\<forall>a \<in> A. f'(f a) = a" and
    31         RIGHT: "\<forall>a' \<in> A'. f(f' a') = a'" and
    32         IM1: "f ` A \<le> A'" and IM2: "f' ` A' \<le> A"
    33 shows "bij_betw f A A'"
    34 using assms
    35 proof(unfold bij_betw_def inj_on_def, safe)
    36   fix a b assume *: "a \<in> A" "b \<in> A" and **: "f a = f b"
    37   have "a = f'(f a) \<and> b = f'(f b)" using * LEFT by simp
    38   with ** show "a = b" by simp
    39 next
    40   fix a' assume *: "a' \<in> A'"
    41   hence "f' a' \<in> A" using IM2 by blast
    42   moreover
    43   have "a' = f(f' a')" using * RIGHT by simp
    44   ultimately show "a' \<in> f ` A" by blast
    45 qed
    46 
    47 
    48 (*3*)corollary notIn_Un_bij_betw:
    49 assumes NIN: "b \<notin> A" and NIN': "f b \<notin> A'" and
    50        BIJ: "bij_betw f A A'"
    51 shows "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
    52 proof-
    53   have "bij_betw f {b} {f b}"
    54   unfolding bij_betw_def inj_on_def by simp
    55   with assms show ?thesis
    56   using bij_betw_combine[of f A A' "{b}" "{f b}"] by blast
    57 qed
    58 
    59 
    60 (*1*)lemma notIn_Un_bij_betw3:
    61 assumes NIN: "b \<notin> A" and NIN': "f b \<notin> A'"
    62 shows "bij_betw f A A' = bij_betw f (A \<union> {b}) (A' \<union> {f b})"
    63 proof
    64   assume "bij_betw f A A'"
    65   thus "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
    66   using assms notIn_Un_bij_betw[of b A f A'] by blast
    67 next
    68   assume *: "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
    69   have "f ` A = A'"
    70   proof(auto)
    71     fix a assume **: "a \<in> A"
    72     hence "f a \<in> A' \<union> {f b}" using * unfolding bij_betw_def by blast
    73     moreover
    74     {assume "f a = f b"
    75      hence "a = b" using * ** unfolding bij_betw_def inj_on_def by blast
    76      with NIN ** have False by blast
    77     }
    78     ultimately show "f a \<in> A'" by blast
    79   next
    80     fix a' assume **: "a' \<in> A'"
    81     hence "a' \<in> f`(A \<union> {b})"
    82     using * by (auto simp add: bij_betw_def)
    83     then obtain a where 1: "a \<in> A \<union> {b} \<and> f a = a'" by blast
    84     moreover
    85     {assume "a = b" with 1 ** NIN' have False by blast
    86     }
    87     ultimately have "a \<in> A" by blast
    88     with 1 show "a' \<in> f ` A" by blast
    89   qed
    90   thus "bij_betw f A A'" using * bij_betw_subset[of f "A \<union> {b}" _ A] by blast
    91 qed
    92 
    93 
    94 subsection {* Properties involving finite and infinite sets *}
    95 
    96 
    97 (*3*)lemma inj_on_finite:
    98 assumes "inj_on f A" "f ` A \<le> B" "finite B"
    99 shows "finite A"
   100 using assms by (metis finite_imageD finite_subset)
   101 
   102 
   103 (*3*)lemma infinite_imp_bij_betw:
   104 assumes INF: "\<not> finite A"
   105 shows "\<exists>h. bij_betw h A (A - {a})"
   106 proof(cases "a \<in> A")
   107   assume Case1: "a \<notin> A"  hence "A - {a} = A" by blast
   108   thus ?thesis using bij_betw_id[of A] by auto
   109 next
   110   assume Case2: "a \<in> A"
   111 find_theorems "\<not> finite _"
   112   have "\<not> finite (A - {a})" using INF by auto
   113   with infinite_iff_countable_subset[of "A - {a}"] obtain f::"nat \<Rightarrow> 'a"
   114   where 1: "inj f" and 2: "f ` UNIV \<le> A - {a}" by blast
   115   obtain g where g_def: "g = (\<lambda> n. if n = 0 then a else f (Suc n))" by blast
   116   obtain A' where A'_def: "A' = g ` UNIV" by blast
   117   have temp: "\<forall>y. f y \<noteq> a" using 2 by blast
   118   have 3: "inj_on g UNIV \<and> g ` UNIV \<le> A \<and> a \<in> g ` UNIV"
   119   proof(auto simp add: Case2 g_def, unfold inj_on_def, intro ballI impI,
   120         case_tac "x = 0", auto simp add: 2)
   121     fix y  assume "a = (if y = 0 then a else f (Suc y))"
   122     thus "y = 0" using temp by (case_tac "y = 0", auto)
   123   next
   124     fix x y
   125     assume "f (Suc x) = (if y = 0 then a else f (Suc y))"
   126     thus "x = y" using 1 temp unfolding inj_on_def by (case_tac "y = 0", auto)
   127   next
   128     fix n show "f (Suc n) \<in> A" using 2 by blast
   129   qed
   130   hence 4: "bij_betw g UNIV A' \<and> a \<in> A' \<and> A' \<le> A"
   131   using inj_on_imp_bij_betw[of g] unfolding A'_def by auto
   132   hence 5: "bij_betw (inv g) A' UNIV"
   133   by (auto simp add: bij_betw_inv_into)
   134   (*  *)
   135   obtain n where "g n = a" using 3 by auto
   136   hence 6: "bij_betw g (UNIV - {n}) (A' - {a})"
   137   using 3 4 unfolding A'_def
   138   by clarify (rule bij_betw_subset, auto simp: image_set_diff)
   139   (*  *)
   140   obtain v where v_def: "v = (\<lambda> m. if m < n then m else Suc m)" by blast
   141   have 7: "bij_betw v UNIV (UNIV - {n})"
   142   proof(unfold bij_betw_def inj_on_def, intro conjI, clarify)
   143     fix m1 m2 assume "v m1 = v m2"
   144     thus "m1 = m2"
   145     by(case_tac "m1 < n", case_tac "m2 < n",
   146        auto simp add: inj_on_def v_def, case_tac "m2 < n", auto)
   147   next
   148     show "v ` UNIV = UNIV - {n}"
   149     proof(auto simp add: v_def)
   150       fix m assume *: "m \<noteq> n" and **: "m \<notin> Suc ` {m'. \<not> m' < n}"
   151       {assume "n \<le> m" with * have 71: "Suc n \<le> m" by auto
   152        then obtain m' where 72: "m = Suc m'" using Suc_le_D by auto
   153        with 71 have "n \<le> m'" by auto
   154        with 72 ** have False by auto
   155       }
   156       thus "m < n" by force
   157     qed
   158   qed
   159   (*  *)
   160   obtain h' where h'_def: "h' = g o v o (inv g)" by blast
   161   hence 8: "bij_betw h' A' (A' - {a})" using 5 7 6
   162   by (auto simp add: bij_betw_trans)
   163   (*  *)
   164   obtain h where h_def: "h = (\<lambda> b. if b \<in> A' then h' b else b)" by blast
   165   have "\<forall>b \<in> A'. h b = h' b" unfolding h_def by auto
   166   hence "bij_betw h  A' (A' - {a})" using 8 bij_betw_cong[of A' h] by auto
   167   moreover
   168   {have "\<forall>b \<in> A - A'. h b = b" unfolding h_def by auto
   169    hence "bij_betw h  (A - A') (A - A')"
   170    using bij_betw_cong[of "A - A'" h id] bij_betw_id[of "A - A'"] by auto
   171   }
   172   moreover
   173   have "(A' Int (A - A') = {} \<and> A' \<union> (A - A') = A) \<and>
   174         ((A' - {a}) Int (A - A') = {} \<and> (A' - {a}) \<union> (A - A') = A - {a})"
   175   using 4 by blast
   176   ultimately have "bij_betw h A (A - {a})"
   177   using bij_betw_combine[of h A' "A' - {a}" "A - A'" "A - A'"] by simp
   178   thus ?thesis by blast
   179 qed
   180 
   181 
   182 (*3*)lemma infinite_imp_bij_betw2:
   183 assumes INF: "\<not> finite A"
   184 shows "\<exists>h. bij_betw h A (A \<union> {a})"
   185 proof(cases "a \<in> A")
   186   assume Case1: "a \<in> A"  hence "A \<union> {a} = A" by blast
   187   thus ?thesis using bij_betw_id[of A] by auto
   188 next
   189   let ?A' = "A \<union> {a}"
   190   assume Case2: "a \<notin> A" hence "A = ?A' - {a}" by blast
   191   moreover have "\<not> finite ?A'" using INF by auto
   192   ultimately obtain f where "bij_betw f ?A' A"
   193   using infinite_imp_bij_betw[of ?A' a] by auto
   194   hence "bij_betw(inv_into ?A' f) A ?A'" using bij_betw_inv_into by blast
   195   thus ?thesis by auto
   196 qed
   197 
   198 
   199 subsection {* Properties involving Hilbert choice *}
   200 
   201 
   202 (*2*)lemma bij_betw_inv_into_left:
   203 assumes BIJ: "bij_betw f A A'" and IN: "a \<in> A"
   204 shows "(inv_into A f) (f a) = a"
   205 using assms unfolding bij_betw_def
   206 by clarify (rule inv_into_f_f)
   207 
   208 (*2*)lemma bij_betw_inv_into_right:
   209 assumes "bij_betw f A A'" "a' \<in> A'"
   210 shows "f(inv_into A f a') = a'"
   211 using assms unfolding bij_betw_def using f_inv_into_f by force
   212 
   213 
   214 (*1*)lemma bij_betw_inv_into_subset:
   215 assumes BIJ: "bij_betw f A A'" and
   216         SUB: "B \<le> A" and IM: "f ` B = B'"
   217 shows "bij_betw (inv_into A f) B' B"
   218 using assms unfolding bij_betw_def
   219 by (auto intro: inj_on_inv_into)
   220 
   221 
   222 
   223 end