src/HOL/Fun.thy
author nipkow
Wed, 08 Sep 2010 10:45:55 +0200
changeset 39443 297cd703f1f0
parent 39428 f967a16dfcdd
child 39535 d7728f65b353
permissions -rw-r--r--
put expand_(fun/set)_eq back in as synonyms, for compatibility
     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 begin
    11 
    12 text{*As a simplification rule, it replaces all function equalities by
    13   first-order equalities.*}
    14 lemma ext_iff: "f = g \<longleftrightarrow> (\<forall>x. f x = g x)"
    15 apply (rule iffI)
    16 apply (simp (no_asm_simp))
    17 apply (rule ext)
    18 apply (simp (no_asm_simp))
    19 done
    20 
    21 lemmas expand_fun_eq = ext_iff
    22 
    23 lemma apply_inverse:
    24   "f x = u \<Longrightarrow> (\<And>x. P x \<Longrightarrow> g (f x) = x) \<Longrightarrow> P x \<Longrightarrow> x = g u"
    25   by auto
    26 
    27 
    28 subsection {* The Identity Function @{text id} *}
    29 
    30 definition
    31   id :: "'a \<Rightarrow> 'a"
    32 where
    33   "id = (\<lambda>x. x)"
    34 
    35 lemma id_apply [simp]: "id x = x"
    36   by (simp add: id_def)
    37 
    38 lemma image_ident [simp]: "(%x. x) ` Y = Y"
    39 by blast
    40 
    41 lemma image_id [simp]: "id ` Y = Y"
    42 by (simp add: id_def)
    43 
    44 lemma vimage_ident [simp]: "(%x. x) -` Y = Y"
    45 by blast
    46 
    47 lemma vimage_id [simp]: "id -` A = A"
    48 by (simp add: id_def)
    49 
    50 
    51 subsection {* The Composition Operator @{text "f \<circ> g"} *}
    52 
    53 definition
    54   comp :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "o" 55)
    55 where
    56   "f o g = (\<lambda>x. f (g x))"
    57 
    58 notation (xsymbols)
    59   comp  (infixl "\<circ>" 55)
    60 
    61 notation (HTML output)
    62   comp  (infixl "\<circ>" 55)
    63 
    64 text{*compatibility*}
    65 lemmas o_def = comp_def
    66 
    67 lemma o_apply [simp]: "(f o g) x = f (g x)"
    68 by (simp add: comp_def)
    69 
    70 lemma o_assoc: "f o (g o h) = f o g o h"
    71 by (simp add: comp_def)
    72 
    73 lemma id_o [simp]: "id o g = g"
    74 by (simp add: comp_def)
    75 
    76 lemma o_id [simp]: "f o id = f"
    77 by (simp add: comp_def)
    78 
    79 lemma o_eq_dest:
    80   "a o b = c o d \<Longrightarrow> a (b v) = c (d v)"
    81   by (simp only: o_def) (fact fun_cong)
    82 
    83 lemma o_eq_elim:
    84   "a o b = c o d \<Longrightarrow> ((\<And>v. a (b v) = c (d v)) \<Longrightarrow> R) \<Longrightarrow> R"
    85   by (erule meta_mp) (fact o_eq_dest) 
    86 
    87 lemma image_compose: "(f o g) ` r = f`(g`r)"
    88 by (simp add: comp_def, blast)
    89 
    90 lemma vimage_compose: "(g \<circ> f) -` x = f -` (g -` x)"
    91   by auto
    92 
    93 lemma UN_o: "UNION A (g o f) = UNION (f`A) g"
    94 by (unfold comp_def, blast)
    95 
    96 
    97 subsection {* The Forward Composition Operator @{text fcomp} *}
    98 
    99 definition
   100   fcomp :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "\<circ>>" 60)
   101 where
   102   "f \<circ>> g = (\<lambda>x. g (f x))"
   103 
   104 lemma fcomp_apply [simp]:  "(f \<circ>> g) x = g (f x)"
   105   by (simp add: fcomp_def)
   106 
   107 lemma fcomp_assoc: "(f \<circ>> g) \<circ>> h = f \<circ>> (g \<circ>> h)"
   108   by (simp add: fcomp_def)
   109 
   110 lemma id_fcomp [simp]: "id \<circ>> g = g"
   111   by (simp add: fcomp_def)
   112 
   113 lemma fcomp_id [simp]: "f \<circ>> id = f"
   114   by (simp add: fcomp_def)
   115 
   116 code_const fcomp
   117   (Eval infixl 1 "#>")
   118 
   119 no_notation fcomp (infixl "\<circ>>" 60)
   120 
   121 
   122 subsection {* Injectivity, Surjectivity and Bijectivity *}
   123 
   124 definition inj_on :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> bool" where -- "injective"
   125   "inj_on f A \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>A. f x = f y \<longrightarrow> x = y)"
   126 
   127 definition surj_on :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b set \<Rightarrow> bool" where -- "surjective"
   128   "surj_on f B \<longleftrightarrow> B \<subseteq> range f"
   129 
   130 definition bij_betw :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> bool" where -- "bijective"
   131   "bij_betw f A B \<longleftrightarrow> inj_on f A \<and> f ` A = B"
   132 
   133 text{*A common special case: functions injective over the entire domain type.*}
   134 
   135 abbreviation
   136   "inj f \<equiv> inj_on f UNIV"
   137 
   138 abbreviation
   139   "surj f \<equiv> surj_on f UNIV"
   140 
   141 abbreviation
   142   "bij f \<equiv> bij_betw f UNIV UNIV"
   143 
   144 lemma injI:
   145   assumes "\<And>x y. f x = f y \<Longrightarrow> x = y"
   146   shows "inj f"
   147   using assms unfolding inj_on_def by auto
   148 
   149 text{*For Proofs in @{text "Tools/Datatype/datatype_rep_proofs"}*}
   150 lemma datatype_injI:
   151     "(!! x. ALL y. f(x) = f(y) --> x=y) ==> inj(f)"
   152 by (simp add: inj_on_def)
   153 
   154 theorem range_ex1_eq: "inj f \<Longrightarrow> b : range f = (EX! x. b = f x)"
   155   by (unfold inj_on_def, blast)
   156 
   157 lemma injD: "[| inj(f); f(x) = f(y) |] ==> x=y"
   158 by (simp add: inj_on_def)
   159 
   160 lemma inj_on_eq_iff: "inj_on f A ==> x:A ==> y:A ==> (f(x) = f(y)) = (x=y)"
   161 by (force simp add: inj_on_def)
   162 
   163 lemma inj_comp:
   164   "inj f \<Longrightarrow> inj g \<Longrightarrow> inj (f \<circ> g)"
   165   by (simp add: inj_on_def)
   166 
   167 lemma inj_fun: "inj f \<Longrightarrow> inj (\<lambda>x y. f x)"
   168   by (simp add: inj_on_def ext_iff)
   169 
   170 lemma inj_eq: "inj f ==> (f(x) = f(y)) = (x=y)"
   171 by (simp add: inj_on_eq_iff)
   172 
   173 lemma inj_on_id[simp]: "inj_on id A"
   174   by (simp add: inj_on_def)
   175 
   176 lemma inj_on_id2[simp]: "inj_on (%x. x) A"
   177 by (simp add: inj_on_def)
   178 
   179 lemma surj_id[simp]: "surj_on id A"
   180 by (simp add: surj_on_def)
   181 
   182 lemma bij_id[simp]: "bij id"
   183 by (simp add: bij_betw_def)
   184 
   185 lemma inj_onI:
   186     "(!! x y. [|  x:A;  y:A;  f(x) = f(y) |] ==> x=y) ==> inj_on f A"
   187 by (simp add: inj_on_def)
   188 
   189 lemma inj_on_inverseI: "(!!x. x:A ==> g(f(x)) = x) ==> inj_on f A"
   190 by (auto dest:  arg_cong [of concl: g] simp add: inj_on_def)
   191 
   192 lemma inj_onD: "[| inj_on f A;  f(x)=f(y);  x:A;  y:A |] ==> x=y"
   193 by (unfold inj_on_def, blast)
   194 
   195 lemma inj_on_iff: "[| inj_on f A;  x:A;  y:A |] ==> (f(x)=f(y)) = (x=y)"
   196 by (blast dest!: inj_onD)
   197 
   198 lemma comp_inj_on:
   199      "[| inj_on f A;  inj_on g (f`A) |] ==> inj_on (g o f) A"
   200 by (simp add: comp_def inj_on_def)
   201 
   202 lemma inj_on_imageI: "inj_on (g o f) A \<Longrightarrow> inj_on g (f ` A)"
   203 apply(simp add:inj_on_def image_def)
   204 apply blast
   205 done
   206 
   207 lemma inj_on_image_iff: "\<lbrakk> ALL x:A. ALL y:A. (g(f x) = g(f y)) = (g x = g y);
   208   inj_on f A \<rbrakk> \<Longrightarrow> inj_on g (f ` A) = inj_on g A"
   209 apply(unfold inj_on_def)
   210 apply blast
   211 done
   212 
   213 lemma inj_on_contraD: "[| inj_on f A;  ~x=y;  x:A;  y:A |] ==> ~ f(x)=f(y)"
   214 by (unfold inj_on_def, blast)
   215 
   216 lemma inj_singleton: "inj (%s. {s})"
   217 by (simp add: inj_on_def)
   218 
   219 lemma inj_on_empty[iff]: "inj_on f {}"
   220 by(simp add: inj_on_def)
   221 
   222 lemma subset_inj_on: "[| inj_on f B; A <= B |] ==> inj_on f A"
   223 by (unfold inj_on_def, blast)
   224 
   225 lemma inj_on_Un:
   226  "inj_on f (A Un B) =
   227   (inj_on f A & inj_on f B & f`(A-B) Int f`(B-A) = {})"
   228 apply(unfold inj_on_def)
   229 apply (blast intro:sym)
   230 done
   231 
   232 lemma inj_on_insert[iff]:
   233   "inj_on f (insert a A) = (inj_on f A & f a ~: f`(A-{a}))"
   234 apply(unfold inj_on_def)
   235 apply (blast intro:sym)
   236 done
   237 
   238 lemma inj_on_diff: "inj_on f A ==> inj_on f (A-B)"
   239 apply(unfold inj_on_def)
   240 apply (blast)
   241 done
   242 
   243 lemma surj_onI: "(\<And>x. x \<in> B \<Longrightarrow> g (f x) = x) \<Longrightarrow> surj_on g B"
   244   by (simp add: surj_on_def) (blast intro: sym)
   245 
   246 lemma surj_onD: "surj_on f B \<Longrightarrow> y \<in> B \<Longrightarrow> \<exists>x. y = f x"
   247   by (auto simp: surj_on_def)
   248 
   249 lemma surj_on_range_iff: "surj_on f B \<longleftrightarrow> (\<exists>A. f ` A = B)"
   250   unfolding surj_on_def by (auto intro!: exI[of _ "f -` B"])
   251 
   252 lemma surj_def: "surj f \<longleftrightarrow> (\<forall>y. \<exists>x. y = f x)"
   253   by (simp add: surj_on_def subset_eq image_iff)
   254 
   255 lemma surjI: "(\<And> x. g (f x) = x) \<Longrightarrow> surj g"
   256   by (blast intro: surj_onI)
   257 
   258 lemma surjD: "surj f \<Longrightarrow> \<exists>x. y = f x"
   259   by (simp add: surj_def)
   260 
   261 lemma surjE: "surj f \<Longrightarrow> (\<And>x. y = f x \<Longrightarrow> C) \<Longrightarrow> C"
   262   by (simp add: surj_def, blast)
   263 
   264 lemma comp_surj: "[| surj f;  surj g |] ==> surj (g o f)"
   265 apply (simp add: comp_def surj_def, clarify)
   266 apply (drule_tac x = y in spec, clarify)
   267 apply (drule_tac x = x in spec, blast)
   268 done
   269 
   270 lemma surj_range: "surj f \<Longrightarrow> range f = UNIV"
   271   by (auto simp add: surj_on_def)
   272 
   273 lemma surj_range_iff: "surj f \<longleftrightarrow> range f = UNIV"
   274   unfolding surj_on_def by auto
   275 
   276 lemma bij_betw_imp_surj: "bij_betw f A UNIV \<Longrightarrow> surj f"
   277   unfolding bij_betw_def surj_range_iff by auto
   278 
   279 lemma bij_def: "bij f \<longleftrightarrow> inj f \<and> surj f"
   280   unfolding surj_range_iff bij_betw_def ..
   281 
   282 lemma bijI: "[| inj f; surj f |] ==> bij f"
   283 by (simp add: bij_def)
   284 
   285 lemma bij_is_inj: "bij f ==> inj f"
   286 by (simp add: bij_def)
   287 
   288 lemma bij_is_surj: "bij f ==> surj f"
   289 by (simp add: bij_def)
   290 
   291 lemma bij_betw_imp_inj_on: "bij_betw f A B \<Longrightarrow> inj_on f A"
   292 by (simp add: bij_betw_def)
   293 
   294 lemma bij_betw_imp_surj_on: "bij_betw f A B \<Longrightarrow> surj_on f B"
   295 by (auto simp: bij_betw_def surj_on_range_iff)
   296 
   297 lemma bij_comp: "bij f \<Longrightarrow> bij g \<Longrightarrow> bij (g o f)"
   298 by(fastsimp intro: comp_inj_on comp_surj simp: bij_def surj_range)
   299 
   300 lemma bij_betw_trans:
   301   "bij_betw f A B \<Longrightarrow> bij_betw g B C \<Longrightarrow> bij_betw (g o f) A C"
   302 by(auto simp add:bij_betw_def comp_inj_on)
   303 
   304 lemma bij_betw_inv: assumes "bij_betw f A B" shows "EX g. bij_betw g B A"
   305 proof -
   306   have i: "inj_on f A" and s: "f ` A = B"
   307     using assms by(auto simp:bij_betw_def)
   308   let ?P = "%b a. a:A \<and> f a = b" let ?g = "%b. The (?P b)"
   309   { fix a b assume P: "?P b a"
   310     hence ex1: "\<exists>a. ?P b a" using s unfolding image_def by blast
   311     hence uex1: "\<exists>!a. ?P b a" by(blast dest:inj_onD[OF i])
   312     hence " ?g b = a" using the1_equality[OF uex1, OF P] P by simp
   313   } note g = this
   314   have "inj_on ?g B"
   315   proof(rule inj_onI)
   316     fix x y assume "x:B" "y:B" "?g x = ?g y"
   317     from s `x:B` obtain a1 where a1: "?P x a1" unfolding image_def by blast
   318     from s `y:B` obtain a2 where a2: "?P y a2" unfolding image_def by blast
   319     from g[OF a1] a1 g[OF a2] a2 `?g x = ?g y` show "x=y" by simp
   320   qed
   321   moreover have "?g ` B = A"
   322   proof(auto simp:image_def)
   323     fix b assume "b:B"
   324     with s obtain a where P: "?P b a" unfolding image_def by blast
   325     thus "?g b \<in> A" using g[OF P] by auto
   326   next
   327     fix a assume "a:A"
   328     then obtain b where P: "?P b a" using s unfolding image_def by blast
   329     then have "b:B" using s unfolding image_def by blast
   330     with g[OF P] show "\<exists>b\<in>B. a = ?g b" by blast
   331   qed
   332   ultimately show ?thesis by(auto simp:bij_betw_def)
   333 qed
   334 
   335 lemma bij_betw_combine:
   336   assumes "bij_betw f A B" "bij_betw f C D" "B \<inter> D = {}"
   337   shows "bij_betw f (A \<union> C) (B \<union> D)"
   338   using assms unfolding bij_betw_def inj_on_Un image_Un by auto
   339 
   340 lemma surj_image_vimage_eq: "surj f ==> f ` (f -` A) = A"
   341 by (simp add: surj_range)
   342 
   343 lemma inj_vimage_image_eq: "inj f ==> f -` (f ` A) = A"
   344 by (simp add: inj_on_def, blast)
   345 
   346 lemma vimage_subsetD: "surj f ==> f -` B <= A ==> B <= f ` A"
   347 apply (unfold surj_def)
   348 apply (blast intro: sym)
   349 done
   350 
   351 lemma vimage_subsetI: "inj f ==> B <= f ` A ==> f -` B <= A"
   352 by (unfold inj_on_def, blast)
   353 
   354 lemma vimage_subset_eq: "bij f ==> (f -` B <= A) = (B <= f ` A)"
   355 apply (unfold bij_def)
   356 apply (blast del: subsetI intro: vimage_subsetI vimage_subsetD)
   357 done
   358 
   359 lemma inj_on_Un_image_eq_iff: "inj_on f (A \<union> B) \<Longrightarrow> f ` A = f ` B \<longleftrightarrow> A = B"
   360 by(blast dest: inj_onD)
   361 
   362 lemma inj_on_image_Int:
   363    "[| inj_on f C;  A<=C;  B<=C |] ==> f`(A Int B) = f`A Int f`B"
   364 apply (simp add: inj_on_def, blast)
   365 done
   366 
   367 lemma inj_on_image_set_diff:
   368    "[| inj_on f C;  A<=C;  B<=C |] ==> f`(A-B) = f`A - f`B"
   369 apply (simp add: inj_on_def, blast)
   370 done
   371 
   372 lemma image_Int: "inj f ==> f`(A Int B) = f`A Int f`B"
   373 by (simp add: inj_on_def, blast)
   374 
   375 lemma image_set_diff: "inj f ==> f`(A-B) = f`A - f`B"
   376 by (simp add: inj_on_def, blast)
   377 
   378 lemma inj_image_mem_iff: "inj f ==> (f a : f`A) = (a : A)"
   379 by (blast dest: injD)
   380 
   381 lemma inj_image_subset_iff: "inj f ==> (f`A <= f`B) = (A<=B)"
   382 by (simp add: inj_on_def, blast)
   383 
   384 lemma inj_image_eq_iff: "inj f ==> (f`A = f`B) = (A = B)"
   385 by (blast dest: injD)
   386 
   387 (*injectivity's required.  Left-to-right inclusion holds even if A is empty*)
   388 lemma image_INT:
   389    "[| inj_on f C;  ALL x:A. B x <= C;  j:A |]
   390     ==> f ` (INTER A B) = (INT x:A. f ` B x)"
   391 apply (simp add: inj_on_def, blast)
   392 done
   393 
   394 (*Compare with image_INT: no use of inj_on, and if f is surjective then
   395   it doesn't matter whether A is empty*)
   396 lemma bij_image_INT: "bij f ==> f ` (INTER A B) = (INT x:A. f ` B x)"
   397 apply (simp add: bij_def)
   398 apply (simp add: inj_on_def surj_def, blast)
   399 done
   400 
   401 lemma surj_Compl_image_subset: "surj f ==> -(f`A) <= f`(-A)"
   402 by (auto simp add: surj_def)
   403 
   404 lemma inj_image_Compl_subset: "inj f ==> f`(-A) <= -(f`A)"
   405 by (auto simp add: inj_on_def)
   406 
   407 lemma bij_image_Compl_eq: "bij f ==> f`(-A) = -(f`A)"
   408 apply (simp add: bij_def)
   409 apply (rule equalityI)
   410 apply (simp_all (no_asm_simp) add: inj_image_Compl_subset surj_Compl_image_subset)
   411 done
   412 
   413 lemma (in ordered_ab_group_add) inj_uminus[simp, intro]: "inj_on uminus A"
   414   by (auto intro!: inj_onI)
   415 
   416 lemma (in linorder) strict_mono_imp_inj_on: "strict_mono f \<Longrightarrow> inj_on f A"
   417   by (auto intro!: inj_onI dest: strict_mono_eq)
   418 
   419 subsection{*Function Updating*}
   420 
   421 definition
   422   fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)" where
   423   "fun_upd f a b == % x. if x=a then b else f x"
   424 
   425 nonterminals
   426   updbinds updbind
   427 syntax
   428   "_updbind" :: "['a, 'a] => updbind"             ("(2_ :=/ _)")
   429   ""         :: "updbind => updbinds"             ("_")
   430   "_updbinds":: "[updbind, updbinds] => updbinds" ("_,/ _")
   431   "_Update"  :: "['a, updbinds] => 'a"            ("_/'((_)')" [1000, 0] 900)
   432 
   433 translations
   434   "_Update f (_updbinds b bs)" == "_Update (_Update f b) bs"
   435   "f(x:=y)" == "CONST fun_upd f x y"
   436 
   437 (* Hint: to define the sum of two functions (or maps), use sum_case.
   438          A nice infix syntax could be defined (in Datatype.thy or below) by
   439 notation
   440   sum_case  (infixr "'(+')"80)
   441 *)
   442 
   443 lemma fun_upd_idem_iff: "(f(x:=y) = f) = (f x = y)"
   444 apply (simp add: fun_upd_def, safe)
   445 apply (erule subst)
   446 apply (rule_tac [2] ext, auto)
   447 done
   448 
   449 (* f x = y ==> f(x:=y) = f *)
   450 lemmas fun_upd_idem = fun_upd_idem_iff [THEN iffD2, standard]
   451 
   452 (* f(x := f x) = f *)
   453 lemmas fun_upd_triv = refl [THEN fun_upd_idem]
   454 declare fun_upd_triv [iff]
   455 
   456 lemma fun_upd_apply [simp]: "(f(x:=y))z = (if z=x then y else f z)"
   457 by (simp add: fun_upd_def)
   458 
   459 (* fun_upd_apply supersedes these two,   but they are useful
   460    if fun_upd_apply is intentionally removed from the simpset *)
   461 lemma fun_upd_same: "(f(x:=y)) x = y"
   462 by simp
   463 
   464 lemma fun_upd_other: "z~=x ==> (f(x:=y)) z = f z"
   465 by simp
   466 
   467 lemma fun_upd_upd [simp]: "f(x:=y,x:=z) = f(x:=z)"
   468 by (simp add: ext_iff)
   469 
   470 lemma fun_upd_twist: "a ~= c ==> (m(a:=b))(c:=d) = (m(c:=d))(a:=b)"
   471 by (rule ext, auto)
   472 
   473 lemma inj_on_fun_updI: "\<lbrakk> inj_on f A; y \<notin> f`A \<rbrakk> \<Longrightarrow> inj_on (f(x:=y)) A"
   474 by (fastsimp simp:inj_on_def image_def)
   475 
   476 lemma fun_upd_image:
   477      "f(x:=y) ` A = (if x \<in> A then insert y (f ` (A-{x})) else f ` A)"
   478 by auto
   479 
   480 lemma fun_upd_comp: "f \<circ> (g(x := y)) = (f \<circ> g)(x := f y)"
   481 by (auto intro: ext)
   482 
   483 
   484 subsection {* @{text override_on} *}
   485 
   486 definition
   487   override_on :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'b"
   488 where
   489   "override_on f g A = (\<lambda>a. if a \<in> A then g a else f a)"
   490 
   491 lemma override_on_emptyset[simp]: "override_on f g {} = f"
   492 by(simp add:override_on_def)
   493 
   494 lemma override_on_apply_notin[simp]: "a ~: A ==> (override_on f g A) a = f a"
   495 by(simp add:override_on_def)
   496 
   497 lemma override_on_apply_in[simp]: "a : A ==> (override_on f g A) a = g a"
   498 by(simp add:override_on_def)
   499 
   500 
   501 subsection {* @{text swap} *}
   502 
   503 definition
   504   swap :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)"
   505 where
   506   "swap a b f = f (a := f b, b:= f a)"
   507 
   508 lemma swap_self [simp]: "swap a a f = f"
   509 by (simp add: swap_def)
   510 
   511 lemma swap_commute: "swap a b f = swap b a f"
   512 by (rule ext, simp add: fun_upd_def swap_def)
   513 
   514 lemma swap_nilpotent [simp]: "swap a b (swap a b f) = f"
   515 by (rule ext, simp add: fun_upd_def swap_def)
   516 
   517 lemma swap_triple:
   518   assumes "a \<noteq> c" and "b \<noteq> c"
   519   shows "swap a b (swap b c (swap a b f)) = swap a c f"
   520   using assms by (simp add: ext_iff swap_def)
   521 
   522 lemma comp_swap: "f \<circ> swap a b g = swap a b (f \<circ> g)"
   523 by (rule ext, simp add: fun_upd_def swap_def)
   524 
   525 lemma swap_image_eq [simp]:
   526   assumes "a \<in> A" "b \<in> A" shows "swap a b f ` A = f ` A"
   527 proof -
   528   have subset: "\<And>f. swap a b f ` A \<subseteq> f ` A"
   529     using assms by (auto simp: image_iff swap_def)
   530   then have "swap a b (swap a b f) ` A \<subseteq> (swap a b f) ` A" .
   531   with subset[of f] show ?thesis by auto
   532 qed
   533 
   534 lemma inj_on_imp_inj_on_swap:
   535   "\<lbrakk>inj_on f A; a \<in> A; b \<in> A\<rbrakk> \<Longrightarrow> inj_on (swap a b f) A"
   536   by (simp add: inj_on_def swap_def, blast)
   537 
   538 lemma inj_on_swap_iff [simp]:
   539   assumes A: "a \<in> A" "b \<in> A" shows "inj_on (swap a b f) A \<longleftrightarrow> inj_on f A"
   540 proof
   541   assume "inj_on (swap a b f) A"
   542   with A have "inj_on (swap a b (swap a b f)) A"
   543     by (iprover intro: inj_on_imp_inj_on_swap)
   544   thus "inj_on f A" by simp
   545 next
   546   assume "inj_on f A"
   547   with A show "inj_on (swap a b f) A" by (iprover intro: inj_on_imp_inj_on_swap)
   548 qed
   549 
   550 lemma surj_imp_surj_swap: "surj f \<Longrightarrow> surj (swap a b f)"
   551   unfolding surj_range_iff by simp
   552 
   553 lemma surj_swap_iff [simp]: "surj (swap a b f) \<longleftrightarrow> surj f"
   554   unfolding surj_range_iff by simp
   555 
   556 lemma bij_betw_swap_iff [simp]:
   557   "\<lbrakk> x \<in> A; y \<in> A \<rbrakk> \<Longrightarrow> bij_betw (swap x y f) A B \<longleftrightarrow> bij_betw f A B"
   558   by (auto simp: bij_betw_def)
   559 
   560 lemma bij_swap_iff [simp]: "bij (swap a b f) \<longleftrightarrow> bij f"
   561   by simp
   562 
   563 hide_const (open) swap
   564 
   565 subsection {* Inversion of injective functions *}
   566 
   567 definition the_inv_into :: "'a set => ('a => 'b) => ('b => 'a)" where
   568 "the_inv_into A f == %x. THE y. y : A & f y = x"
   569 
   570 lemma the_inv_into_f_f:
   571   "[| inj_on f A;  x : A |] ==> the_inv_into A f (f x) = x"
   572 apply (simp add: the_inv_into_def inj_on_def)
   573 apply blast
   574 done
   575 
   576 lemma f_the_inv_into_f:
   577   "inj_on f A ==> y : f`A  ==> f (the_inv_into A f y) = y"
   578 apply (simp add: the_inv_into_def)
   579 apply (rule the1I2)
   580  apply(blast dest: inj_onD)
   581 apply blast
   582 done
   583 
   584 lemma the_inv_into_into:
   585   "[| inj_on f A; x : f ` A; A <= B |] ==> the_inv_into A f x : B"
   586 apply (simp add: the_inv_into_def)
   587 apply (rule the1I2)
   588  apply(blast dest: inj_onD)
   589 apply blast
   590 done
   591 
   592 lemma the_inv_into_onto[simp]:
   593   "inj_on f A ==> the_inv_into A f ` (f ` A) = A"
   594 by (fast intro:the_inv_into_into the_inv_into_f_f[symmetric])
   595 
   596 lemma the_inv_into_f_eq:
   597   "[| inj_on f A; f x = y; x : A |] ==> the_inv_into A f y = x"
   598   apply (erule subst)
   599   apply (erule the_inv_into_f_f, assumption)
   600   done
   601 
   602 lemma the_inv_into_comp:
   603   "[| inj_on f (g ` A); inj_on g A; x : f ` g ` A |] ==>
   604   the_inv_into A (f o g) x = (the_inv_into A g o the_inv_into (g ` A) f) x"
   605 apply (rule the_inv_into_f_eq)
   606   apply (fast intro: comp_inj_on)
   607  apply (simp add: f_the_inv_into_f the_inv_into_into)
   608 apply (simp add: the_inv_into_into)
   609 done
   610 
   611 lemma inj_on_the_inv_into:
   612   "inj_on f A \<Longrightarrow> inj_on (the_inv_into A f) (f ` A)"
   613 by (auto intro: inj_onI simp: image_def the_inv_into_f_f)
   614 
   615 lemma bij_betw_the_inv_into:
   616   "bij_betw f A B \<Longrightarrow> bij_betw (the_inv_into A f) B A"
   617 by (auto simp add: bij_betw_def inj_on_the_inv_into the_inv_into_into)
   618 
   619 abbreviation the_inv :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a)" where
   620   "the_inv f \<equiv> the_inv_into UNIV f"
   621 
   622 lemma the_inv_f_f:
   623   assumes "inj f"
   624   shows "the_inv f (f x) = x" using assms UNIV_I
   625   by (rule the_inv_into_f_f)
   626 
   627 
   628 subsection {* Proof tool setup *} 
   629 
   630 text {* simplifies terms of the form
   631   f(...,x:=y,...,x:=z,...) to f(...,x:=z,...) *}
   632 
   633 simproc_setup fun_upd2 ("f(v := w, x := y)") = {* fn _ =>
   634 let
   635   fun gen_fun_upd NONE T _ _ = NONE
   636     | gen_fun_upd (SOME f) T x y = SOME (Const (@{const_name fun_upd}, T) $ f $ x $ y)
   637   fun dest_fun_T1 (Type (_, T :: Ts)) = T
   638   fun find_double (t as Const (@{const_name fun_upd},T) $ f $ x $ y) =
   639     let
   640       fun find (Const (@{const_name fun_upd},T) $ g $ v $ w) =
   641             if v aconv x then SOME g else gen_fun_upd (find g) T v w
   642         | find t = NONE
   643     in (dest_fun_T1 T, gen_fun_upd (find f) T x y) end
   644 
   645   fun proc ss ct =
   646     let
   647       val ctxt = Simplifier.the_context ss
   648       val t = Thm.term_of ct
   649     in
   650       case find_double t of
   651         (T, NONE) => NONE
   652       | (T, SOME rhs) =>
   653           SOME (Goal.prove ctxt [] [] (Logic.mk_equals (t, rhs))
   654             (fn _ =>
   655               rtac eq_reflection 1 THEN
   656               rtac ext 1 THEN
   657               simp_tac (Simplifier.inherit_context ss @{simpset}) 1))
   658     end
   659 in proc end
   660 *}
   661 
   662 
   663 subsection {* Code generator setup *}
   664 
   665 types_code
   666   "fun"  ("(_ ->/ _)")
   667 attach (term_of) {*
   668 fun term_of_fun_type _ aT _ bT _ = Free ("<function>", aT --> bT);
   669 *}
   670 attach (test) {*
   671 fun gen_fun_type aF aT bG bT i =
   672   let
   673     val tab = Unsynchronized.ref [];
   674     fun mk_upd (x, (_, y)) t = Const ("Fun.fun_upd",
   675       (aT --> bT) --> aT --> bT --> aT --> bT) $ t $ aF x $ y ()
   676   in
   677     (fn x =>
   678        case AList.lookup op = (!tab) x of
   679          NONE =>
   680            let val p as (y, _) = bG i
   681            in (tab := (x, p) :: !tab; y) end
   682        | SOME (y, _) => y,
   683      fn () => Basics.fold mk_upd (!tab) (Const ("HOL.undefined", aT --> bT)))
   684   end;
   685 *}
   686 
   687 code_const "op \<circ>"
   688   (SML infixl 5 "o")
   689   (Haskell infixr 9 ".")
   690 
   691 code_const "id"
   692   (Haskell "id")
   693 
   694 end