src/HOL/Fun.thy
author nipkow
Thu, 07 Jul 2011 21:53:53 +0200
changeset 44567 8e421a529a48
parent 43770 ec9eb1fbfcb8
child 44745 74f1f2dd8f52
permissions -rw-r--r--
added translation to fix critical pair between abbreviations for surj and ~=
     1 (*  Title:      HOL/Fun.thy
     2     Author:     Tobias Nipkow, Cambridge University Computer Laboratory
     3     Copyright   1994  University of Cambridge
     4 *)
     5 
     6 header {* Notions about functions *}
     7 
     8 theory Fun
     9 imports Complete_Lattice
    10 uses ("Tools/enriched_type.ML")
    11 begin
    12 
    13 text{*As a simplification rule, it replaces all function equalities by
    14   first-order equalities.*}
    15 lemma fun_eq_iff: "f = g \<longleftrightarrow> (\<forall>x. f x = g x)"
    16 apply (rule iffI)
    17 apply (simp (no_asm_simp))
    18 apply (rule ext)
    19 apply (simp (no_asm_simp))
    20 done
    21 
    22 lemma apply_inverse:
    23   "f x = u \<Longrightarrow> (\<And>x. P x \<Longrightarrow> g (f x) = x) \<Longrightarrow> P x \<Longrightarrow> x = g u"
    24   by auto
    25 
    26 
    27 subsection {* The Identity Function @{text id} *}
    28 
    29 definition
    30   id :: "'a \<Rightarrow> 'a"
    31 where
    32   "id = (\<lambda>x. x)"
    33 
    34 lemma id_apply [simp]: "id x = x"
    35   by (simp add: id_def)
    36 
    37 lemma image_ident [simp]: "(%x. x) ` Y = Y"
    38 by blast
    39 
    40 lemma image_id [simp]: "id ` Y = Y"
    41 by (simp add: id_def)
    42 
    43 lemma vimage_ident [simp]: "(%x. x) -` Y = Y"
    44 by blast
    45 
    46 lemma vimage_id [simp]: "id -` A = A"
    47 by (simp add: id_def)
    48 
    49 
    50 subsection {* The Composition Operator @{text "f \<circ> g"} *}
    51 
    52 definition
    53   comp :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "o" 55)
    54 where
    55   "f o g = (\<lambda>x. f (g x))"
    56 
    57 notation (xsymbols)
    58   comp  (infixl "\<circ>" 55)
    59 
    60 notation (HTML output)
    61   comp  (infixl "\<circ>" 55)
    62 
    63 text{*compatibility*}
    64 lemmas o_def = comp_def
    65 
    66 lemma o_apply [simp]: "(f o g) x = f (g x)"
    67 by (simp add: comp_def)
    68 
    69 lemma o_assoc: "f o (g o h) = f o g o h"
    70 by (simp add: comp_def)
    71 
    72 lemma id_o [simp]: "id o g = g"
    73 by (simp add: comp_def)
    74 
    75 lemma o_id [simp]: "f o id = f"
    76 by (simp add: comp_def)
    77 
    78 lemma o_eq_dest:
    79   "a o b = c o d \<Longrightarrow> a (b v) = c (d v)"
    80   by (simp only: o_def) (fact fun_cong)
    81 
    82 lemma o_eq_elim:
    83   "a o b = c o d \<Longrightarrow> ((\<And>v. a (b v) = c (d v)) \<Longrightarrow> R) \<Longrightarrow> R"
    84   by (erule meta_mp) (fact o_eq_dest) 
    85 
    86 lemma image_compose: "(f o g) ` r = f`(g`r)"
    87 by (simp add: comp_def, blast)
    88 
    89 lemma vimage_compose: "(g \<circ> f) -` x = f -` (g -` x)"
    90   by auto
    91 
    92 lemma UN_o: "UNION A (g o f) = UNION (f`A) g"
    93 by (unfold comp_def, blast)
    94 
    95 
    96 subsection {* The Forward Composition Operator @{text fcomp} *}
    97 
    98 definition
    99   fcomp :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "\<circ>>" 60)
   100 where
   101   "f \<circ>> g = (\<lambda>x. g (f x))"
   102 
   103 lemma fcomp_apply [simp]:  "(f \<circ>> g) x = g (f x)"
   104   by (simp add: fcomp_def)
   105 
   106 lemma fcomp_assoc: "(f \<circ>> g) \<circ>> h = f \<circ>> (g \<circ>> h)"
   107   by (simp add: fcomp_def)
   108 
   109 lemma id_fcomp [simp]: "id \<circ>> g = g"
   110   by (simp add: fcomp_def)
   111 
   112 lemma fcomp_id [simp]: "f \<circ>> id = f"
   113   by (simp add: fcomp_def)
   114 
   115 code_const fcomp
   116   (Eval infixl 1 "#>")
   117 
   118 no_notation fcomp (infixl "\<circ>>" 60)
   119 
   120 
   121 subsection {* Mapping functions *}
   122 
   123 definition map_fun :: "('c \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'c \<Rightarrow> 'd" where
   124   "map_fun f g h = g \<circ> h \<circ> f"
   125 
   126 lemma map_fun_apply [simp]:
   127   "map_fun f g h x = g (h (f x))"
   128   by (simp add: map_fun_def)
   129 
   130 
   131 subsection {* Injectivity and Bijectivity *}
   132 
   133 definition inj_on :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> bool" where -- "injective"
   134   "inj_on f A \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>A. f x = f y \<longrightarrow> x = y)"
   135 
   136 definition bij_betw :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> bool" where -- "bijective"
   137   "bij_betw f A B \<longleftrightarrow> inj_on f A \<and> f ` A = B"
   138 
   139 text{*A common special case: functions injective, surjective or bijective over
   140 the entire domain type.*}
   141 
   142 abbreviation
   143   "inj f \<equiv> inj_on f UNIV"
   144 
   145 abbreviation surj :: "('a \<Rightarrow> 'b) \<Rightarrow> bool" where -- "surjective"
   146   "surj f \<equiv> (range f = UNIV)"
   147 
   148 abbreviation
   149   "bij f \<equiv> bij_betw f UNIV UNIV"
   150 
   151 text{* The negated case: *}
   152 translations
   153 "\<not> CONST surj f" <= "CONST range f \<noteq> CONST UNIV"
   154 
   155 lemma injI:
   156   assumes "\<And>x y. f x = f y \<Longrightarrow> x = y"
   157   shows "inj f"
   158   using assms unfolding inj_on_def by auto
   159 
   160 theorem range_ex1_eq: "inj f \<Longrightarrow> b : range f = (EX! x. b = f x)"
   161   by (unfold inj_on_def, blast)
   162 
   163 lemma injD: "[| inj(f); f(x) = f(y) |] ==> x=y"
   164 by (simp add: inj_on_def)
   165 
   166 lemma inj_on_eq_iff: "inj_on f A ==> x:A ==> y:A ==> (f(x) = f(y)) = (x=y)"
   167 by (force simp add: inj_on_def)
   168 
   169 lemma inj_on_cong:
   170   "(\<And> a. a : A \<Longrightarrow> f a = g a) \<Longrightarrow> inj_on f A = inj_on g A"
   171 unfolding inj_on_def by auto
   172 
   173 lemma inj_on_strict_subset:
   174   "\<lbrakk> inj_on f B; A < B \<rbrakk> \<Longrightarrow> f`A < f`B"
   175 unfolding inj_on_def unfolding image_def by blast
   176 
   177 lemma inj_comp:
   178   "inj f \<Longrightarrow> inj g \<Longrightarrow> inj (f \<circ> g)"
   179   by (simp add: inj_on_def)
   180 
   181 lemma inj_fun: "inj f \<Longrightarrow> inj (\<lambda>x y. f x)"
   182   by (simp add: inj_on_def fun_eq_iff)
   183 
   184 lemma inj_eq: "inj f ==> (f(x) = f(y)) = (x=y)"
   185 by (simp add: inj_on_eq_iff)
   186 
   187 lemma inj_on_id[simp]: "inj_on id A"
   188   by (simp add: inj_on_def)
   189 
   190 lemma inj_on_id2[simp]: "inj_on (%x. x) A"
   191 by (simp add: inj_on_def)
   192 
   193 lemma inj_on_Int: "\<lbrakk>inj_on f A; inj_on f B\<rbrakk> \<Longrightarrow> inj_on f (A \<inter> B)"
   194 unfolding inj_on_def by blast
   195 
   196 lemma inj_on_INTER:
   197   "\<lbrakk>I \<noteq> {}; \<And> i. i \<in> I \<Longrightarrow> inj_on f (A i)\<rbrakk> \<Longrightarrow> inj_on f (\<Inter> i \<in> I. A i)"
   198 unfolding inj_on_def by blast
   199 
   200 lemma inj_on_Inter:
   201   "\<lbrakk>S \<noteq> {}; \<And> A. A \<in> S \<Longrightarrow> inj_on f A\<rbrakk> \<Longrightarrow> inj_on f (Inter S)"
   202 unfolding inj_on_def by blast
   203 
   204 lemma inj_on_UNION_chain:
   205   assumes CH: "\<And> i j. \<lbrakk>i \<in> I; j \<in> I\<rbrakk> \<Longrightarrow> A i \<le> A j \<or> A j \<le> A i" and
   206          INJ: "\<And> i. i \<in> I \<Longrightarrow> inj_on f (A i)"
   207   shows "inj_on f (\<Union> i \<in> I. A i)"
   208 proof(unfold inj_on_def UNION_def, auto)
   209   fix i j x y
   210   assume *: "i \<in> I" "j \<in> I" and **: "x \<in> A i" "y \<in> A j"
   211          and ***: "f x = f y"
   212   show "x = y"
   213   proof-
   214     {assume "A i \<le> A j"
   215      with ** have "x \<in> A j" by auto
   216      with INJ * ** *** have ?thesis
   217      by(auto simp add: inj_on_def)
   218     }
   219     moreover
   220     {assume "A j \<le> A i"
   221      with ** have "y \<in> A i" by auto
   222      with INJ * ** *** have ?thesis
   223      by(auto simp add: inj_on_def)
   224     }
   225     ultimately show ?thesis using  CH * by blast
   226   qed
   227 qed
   228 
   229 lemma surj_id: "surj id"
   230 by simp
   231 
   232 lemma bij_id[simp]: "bij id"
   233 by (simp add: bij_betw_def)
   234 
   235 lemma inj_onI:
   236     "(!! x y. [|  x:A;  y:A;  f(x) = f(y) |] ==> x=y) ==> inj_on f A"
   237 by (simp add: inj_on_def)
   238 
   239 lemma inj_on_inverseI: "(!!x. x:A ==> g(f(x)) = x) ==> inj_on f A"
   240 by (auto dest:  arg_cong [of concl: g] simp add: inj_on_def)
   241 
   242 lemma inj_onD: "[| inj_on f A;  f(x)=f(y);  x:A;  y:A |] ==> x=y"
   243 by (unfold inj_on_def, blast)
   244 
   245 lemma inj_on_iff: "[| inj_on f A;  x:A;  y:A |] ==> (f(x)=f(y)) = (x=y)"
   246 by (blast dest!: inj_onD)
   247 
   248 lemma comp_inj_on:
   249      "[| inj_on f A;  inj_on g (f`A) |] ==> inj_on (g o f) A"
   250 by (simp add: comp_def inj_on_def)
   251 
   252 lemma inj_on_imageI: "inj_on (g o f) A \<Longrightarrow> inj_on g (f ` A)"
   253 apply(simp add:inj_on_def image_def)
   254 apply blast
   255 done
   256 
   257 lemma inj_on_image_iff: "\<lbrakk> ALL x:A. ALL y:A. (g(f x) = g(f y)) = (g x = g y);
   258   inj_on f A \<rbrakk> \<Longrightarrow> inj_on g (f ` A) = inj_on g A"
   259 apply(unfold inj_on_def)
   260 apply blast
   261 done
   262 
   263 lemma inj_on_contraD: "[| inj_on f A;  ~x=y;  x:A;  y:A |] ==> ~ f(x)=f(y)"
   264 by (unfold inj_on_def, blast)
   265 
   266 lemma inj_singleton: "inj (%s. {s})"
   267 by (simp add: inj_on_def)
   268 
   269 lemma inj_on_empty[iff]: "inj_on f {}"
   270 by(simp add: inj_on_def)
   271 
   272 lemma subset_inj_on: "[| inj_on f B; A <= B |] ==> inj_on f A"
   273 by (unfold inj_on_def, blast)
   274 
   275 lemma inj_on_Un:
   276  "inj_on f (A Un B) =
   277   (inj_on f A & inj_on f B & f`(A-B) Int f`(B-A) = {})"
   278 apply(unfold inj_on_def)
   279 apply (blast intro:sym)
   280 done
   281 
   282 lemma inj_on_insert[iff]:
   283   "inj_on f (insert a A) = (inj_on f A & f a ~: f`(A-{a}))"
   284 apply(unfold inj_on_def)
   285 apply (blast intro:sym)
   286 done
   287 
   288 lemma inj_on_diff: "inj_on f A ==> inj_on f (A-B)"
   289 apply(unfold inj_on_def)
   290 apply (blast)
   291 done
   292 
   293 lemma comp_inj_on_iff:
   294   "inj_on f A \<Longrightarrow> inj_on f' (f ` A) \<longleftrightarrow> inj_on (f' o f) A"
   295 by(auto simp add: comp_inj_on inj_on_def)
   296 
   297 lemma inj_on_imageI2:
   298   "inj_on (f' o f) A \<Longrightarrow> inj_on f A"
   299 by(auto simp add: comp_inj_on inj_on_def)
   300 
   301 lemma surj_def: "surj f \<longleftrightarrow> (\<forall>y. \<exists>x. y = f x)"
   302   by auto
   303 
   304 lemma surjI: assumes *: "\<And> x. g (f x) = x" shows "surj g"
   305   using *[symmetric] by auto
   306 
   307 lemma surjD: "surj f \<Longrightarrow> \<exists>x. y = f x"
   308   by (simp add: surj_def)
   309 
   310 lemma surjE: "surj f \<Longrightarrow> (\<And>x. y = f x \<Longrightarrow> C) \<Longrightarrow> C"
   311   by (simp add: surj_def, blast)
   312 
   313 lemma comp_surj: "[| surj f;  surj g |] ==> surj (g o f)"
   314 apply (simp add: comp_def surj_def, clarify)
   315 apply (drule_tac x = y in spec, clarify)
   316 apply (drule_tac x = x in spec, blast)
   317 done
   318 
   319 lemma bij_betw_imp_surj: "bij_betw f A UNIV \<Longrightarrow> surj f"
   320   unfolding bij_betw_def by auto
   321 
   322 lemma bij_betw_empty1:
   323   assumes "bij_betw f {} A"
   324   shows "A = {}"
   325 using assms unfolding bij_betw_def by blast
   326 
   327 lemma bij_betw_empty2:
   328   assumes "bij_betw f A {}"
   329   shows "A = {}"
   330 using assms unfolding bij_betw_def by blast
   331 
   332 lemma inj_on_imp_bij_betw:
   333   "inj_on f A \<Longrightarrow> bij_betw f A (f ` A)"
   334 unfolding bij_betw_def by simp
   335 
   336 lemma bij_def: "bij f \<longleftrightarrow> inj f \<and> surj f"
   337   unfolding bij_betw_def ..
   338 
   339 lemma bijI: "[| inj f; surj f |] ==> bij f"
   340 by (simp add: bij_def)
   341 
   342 lemma bij_is_inj: "bij f ==> inj f"
   343 by (simp add: bij_def)
   344 
   345 lemma bij_is_surj: "bij f ==> surj f"
   346 by (simp add: bij_def)
   347 
   348 lemma bij_betw_imp_inj_on: "bij_betw f A B \<Longrightarrow> inj_on f A"
   349 by (simp add: bij_betw_def)
   350 
   351 lemma bij_betw_trans:
   352   "bij_betw f A B \<Longrightarrow> bij_betw g B C \<Longrightarrow> bij_betw (g o f) A C"
   353 by(auto simp add:bij_betw_def comp_inj_on)
   354 
   355 lemma bij_comp: "bij f \<Longrightarrow> bij g \<Longrightarrow> bij (g o f)"
   356   by (rule bij_betw_trans)
   357 
   358 lemma bij_betw_comp_iff:
   359   "bij_betw f A A' \<Longrightarrow> bij_betw f' A' A'' \<longleftrightarrow> bij_betw (f' o f) A A''"
   360 by(auto simp add: bij_betw_def inj_on_def)
   361 
   362 lemma bij_betw_comp_iff2:
   363   assumes BIJ: "bij_betw f' A' A''" and IM: "f ` A \<le> A'"
   364   shows "bij_betw f A A' \<longleftrightarrow> bij_betw (f' o f) A A''"
   365 using assms
   366 proof(auto simp add: bij_betw_comp_iff)
   367   assume *: "bij_betw (f' \<circ> f) A A''"
   368   thus "bij_betw f A A'"
   369   using IM
   370   proof(auto simp add: bij_betw_def)
   371     assume "inj_on (f' \<circ> f) A"
   372     thus "inj_on f A" using inj_on_imageI2 by blast
   373   next
   374     fix a' assume **: "a' \<in> A'"
   375     hence "f' a' \<in> A''" using BIJ unfolding bij_betw_def by auto
   376     then obtain a where 1: "a \<in> A \<and> f'(f a) = f' a'" using *
   377     unfolding bij_betw_def by force
   378     hence "f a \<in> A'" using IM by auto
   379     hence "f a = a'" using BIJ ** 1 unfolding bij_betw_def inj_on_def by auto
   380     thus "a' \<in> f ` A" using 1 by auto
   381   qed
   382 qed
   383 
   384 lemma bij_betw_inv: assumes "bij_betw f A B" shows "EX g. bij_betw g B A"
   385 proof -
   386   have i: "inj_on f A" and s: "f ` A = B"
   387     using assms by(auto simp:bij_betw_def)
   388   let ?P = "%b a. a:A \<and> f a = b" let ?g = "%b. The (?P b)"
   389   { fix a b assume P: "?P b a"
   390     hence ex1: "\<exists>a. ?P b a" using s unfolding image_def by blast
   391     hence uex1: "\<exists>!a. ?P b a" by(blast dest:inj_onD[OF i])
   392     hence " ?g b = a" using the1_equality[OF uex1, OF P] P by simp
   393   } note g = this
   394   have "inj_on ?g B"
   395   proof(rule inj_onI)
   396     fix x y assume "x:B" "y:B" "?g x = ?g y"
   397     from s `x:B` obtain a1 where a1: "?P x a1" unfolding image_def by blast
   398     from s `y:B` obtain a2 where a2: "?P y a2" unfolding image_def by blast
   399     from g[OF a1] a1 g[OF a2] a2 `?g x = ?g y` show "x=y" by simp
   400   qed
   401   moreover have "?g ` B = A"
   402   proof(auto simp:image_def)
   403     fix b assume "b:B"
   404     with s obtain a where P: "?P b a" unfolding image_def by blast
   405     thus "?g b \<in> A" using g[OF P] by auto
   406   next
   407     fix a assume "a:A"
   408     then obtain b where P: "?P b a" using s unfolding image_def by blast
   409     then have "b:B" using s unfolding image_def by blast
   410     with g[OF P] show "\<exists>b\<in>B. a = ?g b" by blast
   411   qed
   412   ultimately show ?thesis by(auto simp:bij_betw_def)
   413 qed
   414 
   415 lemma bij_betw_cong:
   416   "(\<And> a. a \<in> A \<Longrightarrow> f a = g a) \<Longrightarrow> bij_betw f A A' = bij_betw g A A'"
   417 unfolding bij_betw_def inj_on_def by force
   418 
   419 lemma bij_betw_id[intro, simp]:
   420   "bij_betw id A A"
   421 unfolding bij_betw_def id_def by auto
   422 
   423 lemma bij_betw_id_iff:
   424   "bij_betw id A B \<longleftrightarrow> A = B"
   425 by(auto simp add: bij_betw_def)
   426 
   427 lemma bij_betw_combine:
   428   assumes "bij_betw f A B" "bij_betw f C D" "B \<inter> D = {}"
   429   shows "bij_betw f (A \<union> C) (B \<union> D)"
   430   using assms unfolding bij_betw_def inj_on_Un image_Un by auto
   431 
   432 lemma bij_betw_UNION_chain:
   433   assumes CH: "\<And> i j. \<lbrakk>i \<in> I; j \<in> I\<rbrakk> \<Longrightarrow> A i \<le> A j \<or> A j \<le> A i" and
   434          BIJ: "\<And> i. i \<in> I \<Longrightarrow> bij_betw f (A i) (A' i)"
   435   shows "bij_betw f (\<Union> i \<in> I. A i) (\<Union> i \<in> I. A' i)"
   436 proof(unfold bij_betw_def, auto simp add: image_def)
   437   have "\<And> i. i \<in> I \<Longrightarrow> inj_on f (A i)"
   438   using BIJ bij_betw_def[of f] by auto
   439   thus "inj_on f (\<Union> i \<in> I. A i)"
   440   using CH inj_on_UNION_chain[of I A f] by auto
   441 next
   442   fix i x
   443   assume *: "i \<in> I" "x \<in> A i"
   444   hence "f x \<in> A' i" using BIJ bij_betw_def[of f] by auto
   445   thus "\<exists>j \<in> I. f x \<in> A' j" using * by blast
   446 next
   447   fix i x'
   448   assume *: "i \<in> I" "x' \<in> A' i"
   449   hence "\<exists>x \<in> A i. x' = f x" using BIJ bij_betw_def[of f] by blast
   450   thus "\<exists>j \<in> I. \<exists>x \<in> A j. x' = f x"
   451   using * by blast
   452 qed
   453 
   454 lemma bij_betw_Disj_Un:
   455   assumes DISJ: "A \<inter> B = {}" and DISJ': "A' \<inter> B' = {}" and
   456           B1: "bij_betw f A A'" and B2: "bij_betw f B B'"
   457   shows "bij_betw f (A \<union> B) (A' \<union> B')"
   458 proof-
   459   have 1: "inj_on f A \<and> inj_on f B"
   460   using B1 B2 by (auto simp add: bij_betw_def)
   461   have 2: "f`A = A' \<and> f`B = B'"
   462   using B1 B2 by (auto simp add: bij_betw_def)
   463   hence "f`(A - B) \<inter> f`(B - A) = {}"
   464   using DISJ DISJ' by blast
   465   hence "inj_on f (A \<union> B)"
   466   using 1 by (auto simp add: inj_on_Un)
   467   (*  *)
   468   moreover
   469   have "f`(A \<union> B) = A' \<union> B'"
   470   using 2 by auto
   471   ultimately show ?thesis
   472   unfolding bij_betw_def by auto
   473 qed
   474 
   475 lemma bij_betw_subset:
   476   assumes BIJ: "bij_betw f A A'" and
   477           SUB: "B \<le> A" and IM: "f ` B = B'"
   478   shows "bij_betw f B B'"
   479 using assms
   480 by(unfold bij_betw_def inj_on_def, auto simp add: inj_on_def)
   481 
   482 lemma surj_image_vimage_eq: "surj f ==> f ` (f -` A) = A"
   483 by simp
   484 
   485 lemma surj_vimage_empty:
   486   assumes "surj f" shows "f -` A = {} \<longleftrightarrow> A = {}"
   487   using surj_image_vimage_eq[OF `surj f`, of A]
   488   by (intro iffI) fastsimp+
   489 
   490 lemma inj_vimage_image_eq: "inj f ==> f -` (f ` A) = A"
   491 by (simp add: inj_on_def, blast)
   492 
   493 lemma vimage_subsetD: "surj f ==> f -` B <= A ==> B <= f ` A"
   494 by (blast intro: sym)
   495 
   496 lemma vimage_subsetI: "inj f ==> B <= f ` A ==> f -` B <= A"
   497 by (unfold inj_on_def, blast)
   498 
   499 lemma vimage_subset_eq: "bij f ==> (f -` B <= A) = (B <= f ` A)"
   500 apply (unfold bij_def)
   501 apply (blast del: subsetI intro: vimage_subsetI vimage_subsetD)
   502 done
   503 
   504 lemma inj_on_Un_image_eq_iff: "inj_on f (A \<union> B) \<Longrightarrow> f ` A = f ` B \<longleftrightarrow> A = B"
   505 by(blast dest: inj_onD)
   506 
   507 lemma inj_on_image_Int:
   508    "[| inj_on f C;  A<=C;  B<=C |] ==> f`(A Int B) = f`A Int f`B"
   509 apply (simp add: inj_on_def, blast)
   510 done
   511 
   512 lemma inj_on_image_set_diff:
   513    "[| inj_on f C;  A<=C;  B<=C |] ==> f`(A-B) = f`A - f`B"
   514 apply (simp add: inj_on_def, blast)
   515 done
   516 
   517 lemma image_Int: "inj f ==> f`(A Int B) = f`A Int f`B"
   518 by (simp add: inj_on_def, blast)
   519 
   520 lemma image_set_diff: "inj f ==> f`(A-B) = f`A - f`B"
   521 by (simp add: inj_on_def, blast)
   522 
   523 lemma inj_image_mem_iff: "inj f ==> (f a : f`A) = (a : A)"
   524 by (blast dest: injD)
   525 
   526 lemma inj_image_subset_iff: "inj f ==> (f`A <= f`B) = (A<=B)"
   527 by (simp add: inj_on_def, blast)
   528 
   529 lemma inj_image_eq_iff: "inj f ==> (f`A = f`B) = (A = B)"
   530 by (blast dest: injD)
   531 
   532 (*injectivity's required.  Left-to-right inclusion holds even if A is empty*)
   533 lemma image_INT:
   534    "[| inj_on f C;  ALL x:A. B x <= C;  j:A |]
   535     ==> f ` (INTER A B) = (INT x:A. f ` B x)"
   536 apply (simp add: inj_on_def, blast)
   537 done
   538 
   539 (*Compare with image_INT: no use of inj_on, and if f is surjective then
   540   it doesn't matter whether A is empty*)
   541 lemma bij_image_INT: "bij f ==> f ` (INTER A B) = (INT x:A. f ` B x)"
   542 apply (simp add: bij_def)
   543 apply (simp add: inj_on_def surj_def, blast)
   544 done
   545 
   546 lemma surj_Compl_image_subset: "surj f ==> -(f`A) <= f`(-A)"
   547 by auto
   548 
   549 lemma inj_image_Compl_subset: "inj f ==> f`(-A) <= -(f`A)"
   550 by (auto simp add: inj_on_def)
   551 
   552 lemma bij_image_Compl_eq: "bij f ==> f`(-A) = -(f`A)"
   553 apply (simp add: bij_def)
   554 apply (rule equalityI)
   555 apply (simp_all (no_asm_simp) add: inj_image_Compl_subset surj_Compl_image_subset)
   556 done
   557 
   558 lemma inj_vimage_singleton: "inj f \<Longrightarrow> f -` {a} \<subseteq> {THE x. f x = a}"
   559   -- {* The inverse image of a singleton under an injective function
   560          is included in a singleton. *}
   561   apply (auto simp add: inj_on_def)
   562   apply (blast intro: the_equality [symmetric])
   563   done
   564 
   565 lemma (in ordered_ab_group_add) inj_uminus[simp, intro]: "inj_on uminus A"
   566   by (auto intro!: inj_onI)
   567 
   568 lemma (in linorder) strict_mono_imp_inj_on: "strict_mono f \<Longrightarrow> inj_on f A"
   569   by (auto intro!: inj_onI dest: strict_mono_eq)
   570 
   571 
   572 subsection{*Function Updating*}
   573 
   574 definition
   575   fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)" where
   576   "fun_upd f a b == % x. if x=a then b else f x"
   577 
   578 nonterminal updbinds and updbind
   579 
   580 syntax
   581   "_updbind" :: "['a, 'a] => updbind"             ("(2_ :=/ _)")
   582   ""         :: "updbind => updbinds"             ("_")
   583   "_updbinds":: "[updbind, updbinds] => updbinds" ("_,/ _")
   584   "_Update"  :: "['a, updbinds] => 'a"            ("_/'((_)')" [1000, 0] 900)
   585 
   586 translations
   587   "_Update f (_updbinds b bs)" == "_Update (_Update f b) bs"
   588   "f(x:=y)" == "CONST fun_upd f x y"
   589 
   590 (* Hint: to define the sum of two functions (or maps), use sum_case.
   591          A nice infix syntax could be defined (in Datatype.thy or below) by
   592 notation
   593   sum_case  (infixr "'(+')"80)
   594 *)
   595 
   596 lemma fun_upd_idem_iff: "(f(x:=y) = f) = (f x = y)"
   597 apply (simp add: fun_upd_def, safe)
   598 apply (erule subst)
   599 apply (rule_tac [2] ext, auto)
   600 done
   601 
   602 (* f x = y ==> f(x:=y) = f *)
   603 lemmas fun_upd_idem = fun_upd_idem_iff [THEN iffD2, standard]
   604 
   605 (* f(x := f x) = f *)
   606 lemmas fun_upd_triv = refl [THEN fun_upd_idem]
   607 declare fun_upd_triv [iff]
   608 
   609 lemma fun_upd_apply [simp]: "(f(x:=y))z = (if z=x then y else f z)"
   610 by (simp add: fun_upd_def)
   611 
   612 (* fun_upd_apply supersedes these two,   but they are useful
   613    if fun_upd_apply is intentionally removed from the simpset *)
   614 lemma fun_upd_same: "(f(x:=y)) x = y"
   615 by simp
   616 
   617 lemma fun_upd_other: "z~=x ==> (f(x:=y)) z = f z"
   618 by simp
   619 
   620 lemma fun_upd_upd [simp]: "f(x:=y,x:=z) = f(x:=z)"
   621 by (simp add: fun_eq_iff)
   622 
   623 lemma fun_upd_twist: "a ~= c ==> (m(a:=b))(c:=d) = (m(c:=d))(a:=b)"
   624 by (rule ext, auto)
   625 
   626 lemma inj_on_fun_updI: "\<lbrakk> inj_on f A; y \<notin> f`A \<rbrakk> \<Longrightarrow> inj_on (f(x:=y)) A"
   627 by (fastsimp simp:inj_on_def image_def)
   628 
   629 lemma fun_upd_image:
   630      "f(x:=y) ` A = (if x \<in> A then insert y (f ` (A-{x})) else f ` A)"
   631 by auto
   632 
   633 lemma fun_upd_comp: "f \<circ> (g(x := y)) = (f \<circ> g)(x := f y)"
   634 by (auto intro: ext)
   635 
   636 
   637 subsection {* @{text override_on} *}
   638 
   639 definition
   640   override_on :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'b"
   641 where
   642   "override_on f g A = (\<lambda>a. if a \<in> A then g a else f a)"
   643 
   644 lemma override_on_emptyset[simp]: "override_on f g {} = f"
   645 by(simp add:override_on_def)
   646 
   647 lemma override_on_apply_notin[simp]: "a ~: A ==> (override_on f g A) a = f a"
   648 by(simp add:override_on_def)
   649 
   650 lemma override_on_apply_in[simp]: "a : A ==> (override_on f g A) a = g a"
   651 by(simp add:override_on_def)
   652 
   653 
   654 subsection {* @{text swap} *}
   655 
   656 definition
   657   swap :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)"
   658 where
   659   "swap a b f = f (a := f b, b:= f a)"
   660 
   661 lemma swap_self [simp]: "swap a a f = f"
   662 by (simp add: swap_def)
   663 
   664 lemma swap_commute: "swap a b f = swap b a f"
   665 by (rule ext, simp add: fun_upd_def swap_def)
   666 
   667 lemma swap_nilpotent [simp]: "swap a b (swap a b f) = f"
   668 by (rule ext, simp add: fun_upd_def swap_def)
   669 
   670 lemma swap_triple:
   671   assumes "a \<noteq> c" and "b \<noteq> c"
   672   shows "swap a b (swap b c (swap a b f)) = swap a c f"
   673   using assms by (simp add: fun_eq_iff swap_def)
   674 
   675 lemma comp_swap: "f \<circ> swap a b g = swap a b (f \<circ> g)"
   676 by (rule ext, simp add: fun_upd_def swap_def)
   677 
   678 lemma swap_image_eq [simp]:
   679   assumes "a \<in> A" "b \<in> A" shows "swap a b f ` A = f ` A"
   680 proof -
   681   have subset: "\<And>f. swap a b f ` A \<subseteq> f ` A"
   682     using assms by (auto simp: image_iff swap_def)
   683   then have "swap a b (swap a b f) ` A \<subseteq> (swap a b f) ` A" .
   684   with subset[of f] show ?thesis by auto
   685 qed
   686 
   687 lemma inj_on_imp_inj_on_swap:
   688   "\<lbrakk>inj_on f A; a \<in> A; b \<in> A\<rbrakk> \<Longrightarrow> inj_on (swap a b f) A"
   689   by (simp add: inj_on_def swap_def, blast)
   690 
   691 lemma inj_on_swap_iff [simp]:
   692   assumes A: "a \<in> A" "b \<in> A" shows "inj_on (swap a b f) A \<longleftrightarrow> inj_on f A"
   693 proof
   694   assume "inj_on (swap a b f) A"
   695   with A have "inj_on (swap a b (swap a b f)) A"
   696     by (iprover intro: inj_on_imp_inj_on_swap)
   697   thus "inj_on f A" by simp
   698 next
   699   assume "inj_on f A"
   700   with A show "inj_on (swap a b f) A" by (iprover intro: inj_on_imp_inj_on_swap)
   701 qed
   702 
   703 lemma surj_imp_surj_swap: "surj f \<Longrightarrow> surj (swap a b f)"
   704   by simp
   705 
   706 lemma surj_swap_iff [simp]: "surj (swap a b f) \<longleftrightarrow> surj f"
   707   by simp
   708 
   709 lemma bij_betw_swap_iff [simp]:
   710   "\<lbrakk> x \<in> A; y \<in> A \<rbrakk> \<Longrightarrow> bij_betw (swap x y f) A B \<longleftrightarrow> bij_betw f A B"
   711   by (auto simp: bij_betw_def)
   712 
   713 lemma bij_swap_iff [simp]: "bij (swap a b f) \<longleftrightarrow> bij f"
   714   by simp
   715 
   716 hide_const (open) swap
   717 
   718 subsection {* Inversion of injective functions *}
   719 
   720 definition the_inv_into :: "'a set => ('a => 'b) => ('b => 'a)" where
   721 "the_inv_into A f == %x. THE y. y : A & f y = x"
   722 
   723 lemma the_inv_into_f_f:
   724   "[| inj_on f A;  x : A |] ==> the_inv_into A f (f x) = x"
   725 apply (simp add: the_inv_into_def inj_on_def)
   726 apply blast
   727 done
   728 
   729 lemma f_the_inv_into_f:
   730   "inj_on f A ==> y : f`A  ==> f (the_inv_into A f y) = y"
   731 apply (simp add: the_inv_into_def)
   732 apply (rule the1I2)
   733  apply(blast dest: inj_onD)
   734 apply blast
   735 done
   736 
   737 lemma the_inv_into_into:
   738   "[| inj_on f A; x : f ` A; A <= B |] ==> the_inv_into A f x : B"
   739 apply (simp add: the_inv_into_def)
   740 apply (rule the1I2)
   741  apply(blast dest: inj_onD)
   742 apply blast
   743 done
   744 
   745 lemma the_inv_into_onto[simp]:
   746   "inj_on f A ==> the_inv_into A f ` (f ` A) = A"
   747 by (fast intro:the_inv_into_into the_inv_into_f_f[symmetric])
   748 
   749 lemma the_inv_into_f_eq:
   750   "[| inj_on f A; f x = y; x : A |] ==> the_inv_into A f y = x"
   751   apply (erule subst)
   752   apply (erule the_inv_into_f_f, assumption)
   753   done
   754 
   755 lemma the_inv_into_comp:
   756   "[| inj_on f (g ` A); inj_on g A; x : f ` g ` A |] ==>
   757   the_inv_into A (f o g) x = (the_inv_into A g o the_inv_into (g ` A) f) x"
   758 apply (rule the_inv_into_f_eq)
   759   apply (fast intro: comp_inj_on)
   760  apply (simp add: f_the_inv_into_f the_inv_into_into)
   761 apply (simp add: the_inv_into_into)
   762 done
   763 
   764 lemma inj_on_the_inv_into:
   765   "inj_on f A \<Longrightarrow> inj_on (the_inv_into A f) (f ` A)"
   766 by (auto intro: inj_onI simp: image_def the_inv_into_f_f)
   767 
   768 lemma bij_betw_the_inv_into:
   769   "bij_betw f A B \<Longrightarrow> bij_betw (the_inv_into A f) B A"
   770 by (auto simp add: bij_betw_def inj_on_the_inv_into the_inv_into_into)
   771 
   772 abbreviation the_inv :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a)" where
   773   "the_inv f \<equiv> the_inv_into UNIV f"
   774 
   775 lemma the_inv_f_f:
   776   assumes "inj f"
   777   shows "the_inv f (f x) = x" using assms UNIV_I
   778   by (rule the_inv_into_f_f)
   779 
   780 subsection {* Cantor's Paradox *}
   781 
   782 lemma Cantors_paradox [no_atp]:
   783   "\<not>(\<exists>f. f ` A = Pow A)"
   784 proof clarify
   785   fix f assume "f ` A = Pow A" hence *: "Pow A \<le> f ` A" by blast
   786   let ?X = "{a \<in> A. a \<notin> f a}"
   787   have "?X \<in> Pow A" unfolding Pow_def by auto
   788   with * obtain x where "x \<in> A \<and> f x = ?X" by blast
   789   thus False by best
   790 qed
   791 
   792 subsection {* Setup *} 
   793 
   794 subsubsection {* Proof tools *}
   795 
   796 text {* simplifies terms of the form
   797   f(...,x:=y,...,x:=z,...) to f(...,x:=z,...) *}
   798 
   799 simproc_setup fun_upd2 ("f(v := w, x := y)") = {* fn _ =>
   800 let
   801   fun gen_fun_upd NONE T _ _ = NONE
   802     | gen_fun_upd (SOME f) T x y = SOME (Const (@{const_name fun_upd}, T) $ f $ x $ y)
   803   fun dest_fun_T1 (Type (_, T :: Ts)) = T
   804   fun find_double (t as Const (@{const_name fun_upd},T) $ f $ x $ y) =
   805     let
   806       fun find (Const (@{const_name fun_upd},T) $ g $ v $ w) =
   807             if v aconv x then SOME g else gen_fun_upd (find g) T v w
   808         | find t = NONE
   809     in (dest_fun_T1 T, gen_fun_upd (find f) T x y) end
   810 
   811   fun proc ss ct =
   812     let
   813       val ctxt = Simplifier.the_context ss
   814       val t = Thm.term_of ct
   815     in
   816       case find_double t of
   817         (T, NONE) => NONE
   818       | (T, SOME rhs) =>
   819           SOME (Goal.prove ctxt [] [] (Logic.mk_equals (t, rhs))
   820             (fn _ =>
   821               rtac eq_reflection 1 THEN
   822               rtac ext 1 THEN
   823               simp_tac (Simplifier.inherit_context ss @{simpset}) 1))
   824     end
   825 in proc end
   826 *}
   827 
   828 
   829 subsubsection {* Code generator *}
   830 
   831 types_code
   832   "fun"  ("(_ ->/ _)")
   833 attach (term_of) {*
   834 fun term_of_fun_type _ aT _ bT _ = Free ("<function>", aT --> bT);
   835 *}
   836 attach (test) {*
   837 fun gen_fun_type aF aT bG bT i =
   838   let
   839     val tab = Unsynchronized.ref [];
   840     fun mk_upd (x, (_, y)) t = Const ("Fun.fun_upd",
   841       (aT --> bT) --> aT --> bT --> aT --> bT) $ t $ aF x $ y ()
   842   in
   843     (fn x =>
   844        case AList.lookup op = (!tab) x of
   845          NONE =>
   846            let val p as (y, _) = bG i
   847            in (tab := (x, p) :: !tab; y) end
   848        | SOME (y, _) => y,
   849      fn () => Basics.fold mk_upd (!tab) (Const ("HOL.undefined", aT --> bT)))
   850   end;
   851 *}
   852 
   853 code_const "op \<circ>"
   854   (SML infixl 5 "o")
   855   (Haskell infixr 9 ".")
   856 
   857 code_const "id"
   858   (Haskell "id")
   859 
   860 
   861 subsubsection {* Functorial structure of types *}
   862 
   863 use "Tools/enriched_type.ML"
   864 
   865 end