Theory func1

(* 
    This file is a part of IsarMathLib - 
    a library of formalized mathematics written for Isabelle/Isar.

    Copyright (C) 2005 - 2024  Slawomir Kolodynski

    This program is free software Redistribution and use in source and binary forms, 
    with or without modification, are permitted provided that the following conditions are met:

   1. Redistributions of source code must retain the above copyright notice, 
   this list of conditions and the following disclaimer.
   2. Redistributions in binary form must reproduce the above copyright notice, 
   this list of conditions and the following disclaimer in the documentation and/or 
   other materials provided with the distribution.
   3. The name of the author may not be used to endorse or promote products 
   derived from this software without specific prior written permission.

THIS SOFTWARE IS PROVIDED BY THE AUTHOR ``AS IS'' AND ANY EXPRESS OR IMPLIED 
WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE IMPLIED WARRANTIES OF 
MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE DISCLAIMED. 
IN NO EVENT SHALL THE AUTHOR BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, 
SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO, 
PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES LOSS OF USE, DATA, OR PROFITS 
OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF LIABILITY, 
WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING NEGLIGENCE OR 
OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS SOFTWARE, 
EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. *)

section Functions - introduction

theory func1 imports ZF.func Fol1 ZF1

begin

textThis theory covers basic properties of function spaces.
  A set of functions with domain $X$ and values in the set $Y$
  is denoted in Isabelle as $X\rightarrow Y$. It just happens
  that the colon ":" is a synonym of the set membership symbol
  $\in$ in Isabelle/ZF so we can write $f:X\rightarrow Y$ instead of 
  $f \in X\rightarrow Y$. This is the only case that we use the colon 
  instead of the regular set membership symbol.

subsectionProperties of functions, function spaces and (inverse) images.

textFunctions in ZF are sets of pairs. This means
  that if $f: X\rightarrow Y $ then $f\subseteq X\times Y$.
  This section is mostly about consequences of this understanding 
  of the notion of function.


textWe define the notion of function that preserves a collection here.
  Given two collection of sets a function preserves the collections if 
  the inverse image of sets in one collection belongs to the second one.
  This notion does not have a name in romantic math. It is used to define 
  continuous functions in Topology_ZF_2› theory. 
  We define it here so that we can 
  use it for other purposes, like defining measurable functions.
  Recall that f-``(A)› means the inverse image of the set $A$.

definition
  "PresColl(f,S,T)   AT. f-``(A)S"

textA definition that allows to get the first factor of the
  domain of a binary function $f: X\times Y \rightarrow Z$.

definition
  "fstdom(f)  domain(domain(f))"

textIf a function maps $A$ into another set, then $A$ is the 
  domain of the function.

lemma func1_1_L1: assumes "f:AC" shows "domain(f) = A"
  using assms domain_of_fun by simp

textStandard Isabelle defines a function(f)› predicate.
  The next lemma shows that our functions satisfy that predicate. 
  It is a special version of Isabelle's fun_is_function›.

lemma fun_is_fun: assumes "f:XY" shows "function(f)"
  using assms fun_is_function by simp

textA lemma explains what fstdom› is for.

lemma fstdomdef: assumes A1: "f: X×Y  Z" and A2: "Y" 
  shows "fstdom(f) = X"
proof -
  from A1 have "domain(f) = X×Y" using func1_1_L1
    by simp
  with A2 show "fstdom(f) = X" unfolding fstdom_def by auto
qed

textA version of the Pi_type› lemma from the standard Isabelle/ZF library.

lemma func1_1_L1A: assumes A1: "f:XY" and A2: "xX. f`(x)  Z"
  shows "f:XZ"
proof -
  { fix x assume "xX" 
    with A2 have "f`(x)  Z" by simp }
  with A1 show "f:XZ" by (rule Pi_type)
qed

textA variant of func1_1_L1A›.

lemma func1_1_L1B: assumes A1: "f:XY" and A2: "YZ"
  shows "f:XZ"
proof -
  from A1 A2 have "xX. f`(x)  Z"
    using apply_funtype by auto
  with A1 show  "f:XZ" using func1_1_L1A by blast
qed

textThere is a value for each argument.

lemma func1_1_L2: assumes A1: "f:XY"  "xX" 
  shows "yY. x,y  f"  
proof -
  from A1 have "f`(x)  Y" using apply_type by simp
  moreover from A1 have " x,f`(x) f" using apply_Pair by simp
  ultimately show ?thesis by auto
qed

textThe inverse image is the image of converse. True for relations as well.

lemma vimage_converse: shows "r-``(A) = converse(r)``(A)"
  using vimage_iff image_iff converse_iff by auto

textThe image is the inverse image of converse.

lemma image_converse: shows "converse(r)-``(A) = r``(A)"
  using vimage_iff image_iff converse_iff by auto

textThe inverse image by a composition is the composition of inverse images.

lemma vimage_comp: shows "(r O s)-``(A) = s-``(r-``(A))"
  using vimage_converse converse_comp image_comp image_converse by simp 

textA version of vimage_comp› for three functions.

lemma vimage_comp3: shows "(r O s O t)-``(A) = t-``(s-``(r-``(A)))"
  using vimage_comp by simp

textInverse image of any set is contained in the domain.

lemma func1_1_L3: assumes A1: "f:XY" shows "f-``(D)  X"
proof -
   have "x. xf-``(D)  x  domain(f)"
      using  vimage_iff domain_iff by auto
    with A1 have "x. (x  f-``(D))  (xX)" using func1_1_L1 by simp
    then show ?thesis by auto
qed

textThe inverse image of the range is the domain.

lemma func1_1_L4: assumes "f:XY" shows "f-``(Y) = X"
  using assms func1_1_L3 func1_1_L2 vimage_iff by blast

textThe arguments belongs to the domain and values to the range.

lemma func1_1_L5: 
  assumes A1: " x,y  f" and A2: "f:XY"  
  shows "xX  yY" 
proof
  from A1 A2 show "xX" using apply_iff by simp
  with A2 have "f`(x) Y" using apply_type by simp
  with A1 A2 show "yY" using apply_iff by simp
qed

textFunction is a subset of cartesian product.

lemma fun_subset_prod: assumes A1: "f:XY" shows "f  X×Y"
proof
  fix p assume "p  f"
  with A1 have "xX. p = x, f`(x)"
    using Pi_memberD by simp
  then obtain x where I: "p = x, f`(x)"
    by auto
  with A1 p  f have "xX  f`(x)  Y"
    using func1_1_L5 by blast
  with I show "p  X×Y" by auto
qed
  
textThe (argument, value) pair belongs to the graph of the function.

lemma func1_1_L5A: 
  assumes A1: "f:XY"  "xX"  "y = f`(x)"
  shows "x,y  f"  "y  range(f)" 
proof -
  from A1 show "x,y  f" using apply_Pair by simp
  then show "y  range(f)" using rangeI by simp
qed

textThe next theorem illustrates the meaning of the concept of 
  function in ZF.

theorem fun_is_set_of_pairs: assumes A1: "f:XY"
  shows "f = {x, f`(x). x  X}"
proof
  from A1 show "{x, f`(x). x  X}  f" using func1_1_L5A
    by auto
next
  { fix p assume "p  f"
    with A1 have "p  X×Y" using fun_subset_prod
      by auto
    with A1 p  f have "p  {x, f`(x). x  X}" 
      using apply_equality by auto
  } thus "f  {x, f`(x). x  X}" by auto
qed

textThe range of function that maps $X$ into $Y$ is contained in $Y$.

lemma func1_1_L5B: 
  assumes  A1: "f:XY" shows "range(f)  Y"
proof
  fix y assume "y  range(f)"
  then obtain x where " x,y  f"
    using range_def converse_def domain_def by auto
  with A1 show "yY" using func1_1_L5 by blast
qed

textThe image of any set is contained in the range.

lemma func1_1_L6: assumes A1: "f:XY" 
  shows "f``(B)  range(f)" and "f``(B)  Y"
proof -
  show "f``(B)  range(f)" using image_iff rangeI by auto
  with A1 show "f``(B)  Y" using func1_1_L5B by blast
qed  

textThe inverse image of any set is contained in the domain.

lemma func1_1_L6A: assumes A1: "f:XY" shows "f-``(A)X"
proof
  fix x
  assume A2: "xf-``(A)" then obtain y where " x,y  f" 
    using vimage_iff by auto
  with A1 show  "xX" using func1_1_L5 by fast
qed

textImage of a greater set is greater.

lemma func1_1_L8: assumes A1: "AB"  shows "f``(A) f``(B)"
  using assms image_Un by auto

textA set is contained in the the inverse image of its image.
  There is similar theorem in equalities.thy›
  (function_image_vimage›)
  which shows that the image of inverse image of a set 
  is contained in the set.

lemma func1_1_L9: assumes A1: "f:XY" and A2: "AX"
  shows "A  f-``(f``(A))"
proof -
  from A1 A2 have "xA.  x,f`(x)  f"  using apply_Pair by auto
  then show ?thesis using image_iff by auto
qed

textThe inverse image of the image of the domain is the domain.

lemma inv_im_dom: assumes A1: "f:XY" shows "f-``(f``(X)) = X"
proof
  from A1 show "f-``(f``(X))  X" using func1_1_L3 by simp
  from A1 show "X  f-``(f``(X))" using func1_1_L9 by simp
qed

textA technical lemma needed to make the func1_1_L11› 
  proof more clear.

lemma func1_1_L10: 
  assumes A1: "f  X×Y" and A2: "∃!y. (yY  x,y  f)"
  shows "∃!y. x,y  f"
proof
  from A2 show "y. x, y  f" by auto
  fix y n assume "x,y  f" and "x,n  f"
  with A1 A2 show "y=n" by auto
qed

textIf $f\subseteq X\times Y$ and for every $x\in X$ there is exactly 
one $y\in Y$ such that $(x,y)\in f$ then $f$ maps $X$ to $Y$.

lemma func1_1_L11: 
  assumes "f  X×Y" and "xX. ∃!y. yY  x,y  f"
  shows "f: XY" using assms func1_1_L10 Pi_iff_old by simp

textA set defined by a lambda-type expression is a fuction. There is a 
  similar lemma in func.thy, but I had problems with lambda expressions syntax
  so I could not apply it. This lemma is a workaround for this. Besides, lambda
  expressions are not readable.


lemma func1_1_L11A: assumes A1: "xX. b(x)  Y"
  shows "{x,y  X×Y. b(x) = y} : XY"
proof -
  let ?f = "{ x,y  X×Y. b(x) = y}"
  have "?f  X×Y" by auto
  moreover have "xX. ∃!y. yY   x,y  ?f"
  proof
    fix x assume A2: "xX"
    show "∃!y. yY  x, y  {x,y  X×Y . b(x) = y}"
    proof
      from A2 A1 show 
        "y. yY  x, y  {x,y  X×Y . b(x) = y}"
	by simp
    next
      fix y y1
      assume "yY  x, y  {x,y  X×Y . b(x) = y}"
	and "y1Y  x, y1  {x,y  X×Y . b(x) = y}"
      then show "y = y1" by simp
    qed
  qed
  ultimately show "{ x,y  X×Y. b(x) = y} : XY" 
    using func1_1_L11 by simp
qed

textThe next lemma will replace func1_1_L11A› one day.

lemma ZF_fun_from_total: assumes A1: "xX. b(x)  Y"
  shows "{x,b(x). xX} : XY"
proof -
  let ?f = "{x,b(x). xX}"
  { fix x assume A2: "xX"
    have "∃!y. yY  x, y  ?f"
    proof
	from A1 A2 show "y. yY  x, y  ?f"
	by simp
    next fix y y1 assume "yY  x, y  ?f"
	and "y1Y  x, y1  ?f"
      then show "y = y1" by simp
    qed
  } then have "xX. ∃!y. yY   x,y  ?f"
    by simp
  moreover from A1 have "?f  X×Y" by auto
  ultimately show ?thesis using func1_1_L11
    by simp
qed
 
textThe value of a function defined by a meta-function is this 
  meta-function (deprecated, use ZF_fun_from_tot_val(1)› instead).

lemma func1_1_L11B: 
  assumes A1: "f:XY"   "xX"
  and A2: "f = {x,y  X×Y. b(x) = y}"
  shows "f`(x) = b(x)"
proof -
  from A1 have " x,f`(x)  f" using apply_iff by simp
  with A2 show ?thesis by simp
qed

textThe next lemma will replace func1_1_L11B› one day.

lemma ZF_fun_from_tot_val: 
  assumes "f:XY"   "xX"
  and "f = {x,b(x). xX}"
  shows "f`(x) = b(x)" and "b(x)Y"
proof -
  from assms(1,2) have "x,f`(x)  f" using apply_iff by simp
  with assms(3) show "f`(x) = b(x)" by simp
  from assms(1,2) have "f`(x)Y" by (rule apply_funtype)
  with f`(x) = b(x) show "b(x)Y" by simp 
qed

textIdentical meaning as ZF_fun_from_tot_val›, but
  phrased a bit differently.

lemma ZF_fun_from_tot_val0: 
  assumes "f:XY" and "f = {x,b(x). xX}"
  shows "xX. f`(x) = b(x)"
  using assms ZF_fun_from_tot_val by simp

textAnother way of expressing that lambda expression is a function.

lemma lam_is_fun_range: assumes "f={x,g(x). xX}"
  shows "f:Xrange(f)"
proof -
  have "xX. g(x)  range({x,g(x). xX})" unfolding range_def 
    by auto
  then have "{x,g(x). xX} : Xrange({x,g(x). xX})" by (rule ZF_fun_from_total)
  with assms show ?thesis by auto
qed

textYet another way of expressing value of a function.

lemma ZF_fun_from_tot_val1:
  assumes "xX" shows "{x,b(x). xX}`(x)=b(x)"
proof -
  let ?f = "{x,b(x). xX}"
  have "?f:Xrange(?f)" using lam_is_fun_range by simp
  with assms show ?thesis using ZF_fun_from_tot_val0 by simp
qed

textAn hypotheses-free form of ZF_fun_from_tot_val1›: the value of a function
  $X\ni x \mapsto p(x)$ is $p(x)$ for all $x\in X$. 

lemma ZF_fun_from_tot_val2: shows "xX. {x,b(x). xX}`(x) = b(x)"
  using ZF_fun_from_tot_val1 by simp

textThe range of a function defined by set comprehension is the set of its values.

lemma range_fun: shows "range({x,b(x). xX}) = {b(x). xX}" 
  by blast

textIn Isabelle/ZF and Metamath if $x$ is not in the domain of a function $f$
  then $f(x)$ is the empty set. This allows us to conclude that if $y\in f(x)$, then
  $x$ must be en element of the domain of $f$. 

lemma arg_in_domain: assumes "f:XY" "yf`(x)" shows "xX"
proof -
  { assume "xX"
    with assms have False using func1_1_L1 apply_0 by simp
  } thus ?thesis by auto
qed

textWe can extend a function by specifying its values on a set
  disjoint with the domain.

lemma func1_1_L11C: assumes A1: "f:XY" and A2: "xA. b(x)B"
  and A3: "XA = " and Dg: "g = f  {x,b(x). xA}"
  shows 
  "g : XA  YB"
  "xX. g`(x) = f`(x)"
  "xA. g`(x) = b(x)"
proof -
  let ?h = "{x,b(x). xA}"
  from A1 A2 A3 have 
    I: "f:XY"  "?h : AB"  "XA = "
    using ZF_fun_from_total by auto
  then have "f?h : XA  YB"
    by (rule fun_disjoint_Un)
  with Dg show "g : XA  YB" by simp
  { fix x assume A4: "xA"
    with A1 A3 have "(f?h)`(x) = ?h`(x)"
      using func1_1_L1 fun_disjoint_apply2
      by blast
    moreover from I A4 have "?h`(x) = b(x)"
      using ZF_fun_from_tot_val by simp
    ultimately have "(f?h)`(x) = b(x)"
      by simp
  } with Dg show "xA. g`(x) = b(x)" by simp
  { fix x assume A5: "xX"
    with A3 I have "x  domain(?h)"
      using func1_1_L1 by auto
    then have "(f?h)`(x) = f`(x)"
      using fun_disjoint_apply1 by simp
  } with Dg show "xX. g`(x) = f`(x)" by simp
qed

textWe can extend a function by specifying its value at a point that
  does not belong to the domain.

lemma func1_1_L11D: assumes A1: "f:XY" and A2: "aX"
  and Dg: "g = f  {a,b}"
  shows 
  "g : X{a}  Y{b}"
  "xX. g`(x) = f`(x)"
  "g`(a) = b"
proof -
  let ?h = "{a,b}"
  from A1 A2 Dg have I:
    "f:XY"  "x{a}. b{b}"  "X{a} = "  "g = f  {x,b. x{a}}"
    by auto
  then show "g : X{a}  Y{b}"
    by (rule func1_1_L11C)
  from I show "xX. g`(x) = f`(x)"
    by (rule func1_1_L11C)
  from I have "x{a}. g`(x) = b"
    by (rule func1_1_L11C)
  then show "g`(a) = b" by auto
qed

textA technical lemma about extending a function both by defining
  on a set disjoint with the domain and on a point that does not belong
  to any of those sets.

lemma func1_1_L11E:
  assumes A1: "f:XY" and 
  A2: "xA. b(x)B" and 
  A3: "XA = " and A4: "a  XA"
  and Dg: "g = f  {x,b(x). xA}  {a,c}"
  shows
  "g : XA{a}  YB{c}"
  "xX. g`(x) = f`(x)"
  "xA. g`(x) = b(x)"
  "g`(a) = c"
proof -
  let ?h = "f  {x,b(x). xA}"
  from assms show "g : XA{a}  YB{c}"
    using func1_1_L11C func1_1_L11D by simp
  from A1 A2 A3 have I:
    "f:XY"  "xA. b(x)B"  "XA = "  "?h = f  {x,b(x). xA}"
    by auto
  from assms have 
    II: "?h : XA  YB"  "a XA"  "g = ?h  {a,c}"
    using func1_1_L11C by auto
  then have III: "xXA. g`(x) = ?h`(x)" by (rule func1_1_L11D)
  moreover from I have  "xX. ?h`(x) = f`(x)"
    by (rule func1_1_L11C)
  ultimately show "xX. g`(x) = f`(x)" by simp
  from I have "xA. ?h`(x) = b(x)" by (rule func1_1_L11C)
  with III show "xA. g`(x) = b(x)" by simp
  from II show "g`(a) = c" by (rule func1_1_L11D)
qed

textA way of defining a function on a union of two possibly overlapping sets. We decompose the 
union into two differences and the intersection and define a function separately on each part.

lemma fun_union_overlap: assumes "xAB. h(x)  Y"  "xAB. f(x)  Y"  "xBA. g(x)  Y"
  shows "{x,if xAB then f(x) else if xBA then g(x) else h(x). x  AB}: AB  Y"
proof -
  let ?F = "{x,if xAB then f(x) else if xBA then g(x) else h(x). x  AB}"
  from assms have "xAB. (if xAB then f(x) else if xBA then g(x) else h(x))  Y"
    by auto
  then show ?thesis by (rule ZF_fun_from_total)
qed

textInverse image of intersection is the intersection of inverse images.

lemma invim_inter_inter_invim: assumes "f:XY"
  shows "f-``(AB) = f-``(A)  f-``(B)"
  using assms fun_is_fun function_vimage_Int by simp

textThe inverse image of an intersection of a nonempty collection of sets 
  is the intersection of the 
  inverse images. This generalizes invim_inter_inter_invim› 
  which is proven for the case of two sets.

lemma func1_1_L12:
  assumes A1: "B  Pow(Y)" and A2: "B" and A3: "f:XY"
  shows "f-``(B) = (UB. f-``(U))"
proof
  from A2 show  "f-``(B)  (UB. f-``(U))" by blast
  show "(UB. f-``(U))  f-``(B)"
  proof
    fix x assume A4: "x  (UB. f-``(U))"
    from A3 have "UB. f-``(U)  X" using func1_1_L6A by simp
    with A4 have "UB. xX" by auto
    with A2 have "xX" by auto
    with A3 have "∃!y.  x,y  f" using Pi_iff_old by simp
    with A2 A4 show "x  f-``(B)" using vimage_iff by blast
  qed
qed

textThe inverse image of a set does not change when we intersect
  the set with the image of the domain.

lemma inv_im_inter_im: assumes "f:XY" 
  shows "f-``(A  f``(X)) = f-``(A)"
  using assms invim_inter_inter_invim inv_im_dom func1_1_L6A
  by blast

textIf the inverse image of a set is not empty, then the set is not empty.
  Proof by contradiction.

lemma func1_1_L13: assumes A1:"f-``(A)  " shows "A"
  using assms by auto

textIf the image of a set is not empty, then the set is not empty.
  Proof by contradiction.

lemma func1_1_L13A: assumes A1: "f``(A)" shows "A"
  using assms by auto

textWhat is the inverse image of a singleton?

lemma func1_1_L14: assumes "f:XY" 
  shows "f-``({y}) = {xX. f`(x) = y}" 
  using assms func1_1_L6A vimage_singleton_iff apply_iff by auto

textA lemma that can be used instead fun_extension_iff›
  to show that two functions are equal

lemma func_eq: 
  assumes "f: XY"  "g: XZ"and  "xX. f`(x) = g`(x)"
  shows "f = g" using assms fun_extension_iff by simp

textAn alternative syntax for defining a function: instead of writing 
  $\{\langle x,p(x)\rangle. x\in X\}$ we can write $\lambda x\in X. p(x)$. 

lemma lambda_fun_alt: shows "{x,p(x). xX} = (λxX. p(x))"
proof -
  let ?L = "{x,p(x). xX}" 
  let ?R = "λxX. p(x)"
  have "?L:Xrange(?L)" and "?R:Xrange(?L)" 
    using lam_is_fun_range range_fun lam_funtype by simp_all
  moreover have "xX. ?L`(x) = ?R`(x)" using ZF_fun_from_tot_val1 beta by simp
  ultimately show "?L = ?R" using func_eq by blast
qed

textIf a function is equal to an expression $b(x)$ on $X$, then it has to be
  of the form $\{ \langle x, b(x)\rangle | x\in X\}$. 

lemma func_eq_set_of_pairs: assumes "f:XY" "xX. f`(x) = b(x)"
  shows "f = {x, b(x). x  X}" 
proof -
  from assms(1) have "f = {x, f`(x). x  X}" using fun_is_set_of_pairs
    by simp
  with assms(2) show ?thesis by simp
qed

textFunction defined on a singleton is a single pair.

lemma func_singleton_pair: assumes A1: "f : {a}X"
  shows "f = {a, f`(a)}"
proof -
  let ?g = "{a, f`(a)}"
  note A1
  moreover have "?g : {a}  {f`(a)}" using singleton_fun by simp
  moreover have "x  {a}. f`(x) = ?g`(x)" using singleton_apply
    by simp
  ultimately show "f = ?g" by (rule func_eq)
qed

textA single pair is a function on a singleton. This is
  similar to singleton_fun› from standard Isabelle/ZF.

lemma pair_func_singleton: assumes A1: "y  Y"
  shows "{x,y} : {x}  Y"
proof -
  have "{x,y} : {x}  {y}" using singleton_fun by simp
  moreover from A1 have "{y}  Y" by simp
  ultimately show "{x,y} : {x}  Y"
    by (rule func1_1_L1B)
qed

textThe value of a pair on the first element is the second one.

lemma pair_val: shows "{x,y}`(x) = y"
  using singleton_fun apply_equality by simp
  
textA more familiar definition of inverse image.

lemma func1_1_L15: assumes A1: "f:XY"
  shows "f-``(A) = {xX. f`(x)  A}"
proof -
  have "f-``(A) = (yA . f-``{y})" 
    by (rule vimage_eq_UN)
  with A1 show ?thesis using func1_1_L14 by auto
qed

textA more familiar definition of image.

lemma func_imagedef: assumes A1: "f:XY" and A2: "AX"
  shows "f``(A) = {f`(x). x  A}"
proof
  from A1 show "f``(A)  {f`(x). x  A}"
    using image_iff apply_iff by auto
  show "{f`(x). x  A}  f``(A)"
  proof
    fix y assume "y  {f`(x). x  A}"
    then obtain x where "xA" and  "y = f`(x)"
      by auto
    with A1 A2 have "x,y  f" using apply_iff by force  
    with A1 A2 xA show "y  f``(A)" using image_iff by auto
  qed
qed  

textIf all elements of a nonempty set map to the same element of the codomain,
  then the image of this set is a singleton.

lemma image_constant_singleton: 
  assumes "f:XY" "AX" "A" "xA. f`(x) = c"
  shows "f``(A) = {c}"
  using assms func_imagedef by auto

textA technical lemma about graphs of functions: if we have two disjoint sets $A$ and $B$
  then the cartesian product of the inverse image of $A$ and $B$ is disjoint
  with (the graph of) $f$.

lemma vimage_prod_dis_graph: assumes "f:XY" "AB = "
  shows "f-``(A)×B  f = "
proof -
  { assume "f-``(A)×B  f  "
    then obtain p where "p  f-``(A)×B" and "pf" by blast
    from assms(1) pf have "p  {x, f`(x). x  X}" 
      using fun_is_set_of_pairs by simp
    then obtain x where "p = x, f`(x)" by blast
    with assms p  f-``(A)×B have False using func1_1_L15 by auto
  } thus ?thesis by auto
qed

text For two functions with the same domain $X$ and the codomain $Y,Z$ resp., we can define
  a third one that maps $X$ to the cartesian product of $Y$ and $Z$. 

lemma prod_fun_val: 
  assumes "{x,p(x). xX}: XY" "{x,q(x). xX}: XZ"
  defines "h  {x,p(x),q(x). xX}"
  shows "h:XY×Z" and "xX. h`(x) = p(x),q(x)"
proof -
  from assms(1,2) have "xX. p(x),q(x)  Y×Z"
    using ZF_fun_from_tot_val(2) by auto
  with assms(3) show "h:XY×Z" using ZF_fun_from_total by simp
  with assms(3) show "xX. h`(x) = p(x),q(x)" using ZF_fun_from_tot_val0
    by simp
qed
  
textSuppose we have two functions $f:X\rightarrow Y$ and $g:X\rightarrow Z$ and 
  the third one is defined as $h:X\rightarrow Y\times Z$, $x\mapsto \langle f(x),g(x)\rangle$.
  Given two sets $U$, $V$ we have $h^{-1}(U\times V) = (f^{-1}(U)) \cap (g^{-1}(V))$. 
  We also show that the set where the function $f,g$ are equal is the same as 
  $h^{-1}(\{ \langle y,y\rangle : y\in X\}$. 
  It is a bit surprising that we get the last identity without the assumption that $Y=Z$. 

lemma vimage_prod: 
  assumes "f:XY" "g:XZ" 
  defines "h  {x,f`(x),g`(x). xX}"
  shows 
    "h:XY×Z" 
    "xX. h`(x) = f`(x),g`(x)" 
    "h-``(U×V) = f-``(U)  g-``(V)"
    "{xX. f`(x) = g`(x)} = h-``({y,y. yY})"
proof - 
  from assms show "h:XY×Z" using apply_funtype ZF_fun_from_total
    by simp
  with assms(3) show I: "xX. h`(x) = f`(x),g`(x)" 
    using ZF_fun_from_tot_val by simp
  with assms(1,2) h:XY×Z show "h-``(U×V) = f-``(U)  g-``(V)"
    using func1_1_L15 by auto
  from assms(1) I h:XY×Z show "{xX. f`(x) = g`(x)} = h-``({y,y. yY})"
    using apply_funtype func1_1_L15 by auto
qed

textThe image of a set contained in domain under identity is the same set.

lemma image_id_same: assumes "AX" shows "id(X)``(A) = A"
  using assms id_type id_conv by auto

textThe inverse image of a set contained in domain under identity is the same set.

lemma vimage_id_same: assumes "AX" shows "id(X)-``(A) = A"
  using assms id_type id_conv by auto

textWhat is the image of a singleton?

lemma singleton_image: 
  assumes "fXY" and "xX"
  shows "f``{x} = {f`(x)}"
  using assms func_imagedef by auto

textIf an element of the domain of a function belongs to a set, 
  then its value belongs to the image of that set.

lemma func1_1_L15D: assumes "f:XY"  "xA"  "AX"
  shows "f`(x)  f``(A)"
  using assms func_imagedef by auto

textRange is the image of the domain. Isabelle/ZF defines
  range(f)› as domain(converse(f))›,
  and that's why we have something to prove here.

lemma range_image_domain: 
  assumes A1: "f:XY" shows "f``(X) = range(f)"
proof
  show "f``(X)  range(f)" using image_def by auto
  { fix y assume "y  range(f)"
    then obtain x where "y,x  converse(f)" by auto
    with A1 have "xX" using func1_1_L5 by blast
    with A1 have "f`(x)  f``(X)" using func_imagedef
      by auto
    with A1  y,x  converse(f) have "y  f``(X)"
      using apply_equality by auto
  } then show "range(f)  f``(X)" by auto
qed
    
textThe difference of images is contained in the image of difference.

lemma diff_image_diff: assumes A1: "f: XY" and A2: "AX"
  shows "f``(X)f``(A)  f``(XA)"
proof
  fix y assume "y  f``(X)  f``(A)"
  hence "y  f``(X)" and I: "y  f``(A)" by auto
  with A1 obtain x where "xX" and II: "y = f`(x)"
    using func_imagedef by auto
  with A1 A2 I have "xA"
    using func1_1_L15D by auto
  with xX have "x  XA" "XA  X" by auto
  with A1 II show "y  f``(XA)"
    using func1_1_L15D by simp
qed

textThe image of an intersection is contained in the 
  intersection of the images.

lemma image_of_Inter: assumes  A1: "f:XY" and
  A2: "I" and A3: "iI. P(i)  X"
  shows "f``(iI. P(i))  ( iI. f``(P(i)) )"
proof
  fix y assume A4: "y  f``(iI. P(i))"
  from A1 A2 A3 have "f``(iI. P(i)) = {f`(x). x  ( iI. P(i) )}"
    using ZF1_1_L7 func_imagedef by simp
  with A4 obtain x where "x  ( iI. P(i) )" and "y = f`(x)"
    by auto
  with A1 A2 A3 show "y  ( iI. f``(P(i)) )" using func_imagedef
    by auto
qed

textThe image of union is the union of images.

lemma image_of_Union: assumes A1: "f:XY" and A2: "AM. AX"
  shows "f``(M) = {f``(A). AM}"
proof
  from A2 have "M  X" by auto
  { fix y assume "y  f``(M)"
    with A1 M  X obtain x where "xM" and I: "y = f`(x)" 
      using func_imagedef by auto
    then obtain A where "AM" and "xA" by auto
    with assms I have "y  {f``(A). AM}" using func_imagedef by auto
  } thus "f``(M)  {f``(A). AM}" by auto
  { fix y assume "y  {f``(A). AM}"
    then obtain A where "AM" and "y  f``(A)" by auto
    with assms M  X have "y  f``(M)" using func_imagedef by auto
  } thus "{f``(A). AM}  f``(M)" by auto
qed

textIf the domain of a function is nonempty, then the codomain is as well.

lemma codomain_nonempty: assumes "f:XY" "X" shows "Y"
  using assms apply_funtype by blast

textThe image of a nonempty subset of domain is nonempty.

lemma func1_1_L15A: 
  assumes A1: "f: XY" and A2: "AX" and A3: "A"
  shows "f``(A)  "
proof -
  from A3 obtain x where "xA" by auto
  with A1 A2 have "f`(x)  f``(A)"
    using func_imagedef by auto
  then show "f``(A)  " by auto
qed

textThe next lemma allows to prove statements about the values in the
  domain of a function given a statement about values in the range.

lemma func1_1_L15B: 
  assumes "f:XY" and "AX" and "yf``(A). P(y)"
  shows "xA. P(f`(x))"
  using assms func_imagedef by simp
  
textAn image of an image is the image of a composition.

lemma func1_1_L15C: assumes  A1: "f:XY" and A2: "g:YZ"
  and A3: "AX"
  shows 
  "g``(f``(A)) =  {g`(f`(x)). xA}"
  "g``(f``(A)) = (g O f)``(A)"
proof -
  from A1 A3 have "{f`(x). xA}  Y"
    using apply_funtype by auto
  with A2 have "g``{f`(x). xA} = {g`(f`(x)). xA}"
    using func_imagedef by auto
  with A1 A3 show I: "g``(f``(A)) =  {g`(f`(x)). xA}" 
    using func_imagedef by simp
  from A1 A3 have "xA. (g O f)`(x) = g`(f`(x))"
    using comp_fun_apply by auto
  with I have "g``(f``(A)) = {(g O f)`(x). xA}"
    by simp
  moreover from A1 A2 A3 have "(g O f)``(A) = {(g O f)`(x). xA}"
    using comp_fun func_imagedef by blast
  ultimately show "g``(f``(A)) = (g O f)``(A)"
    by simp
qed
 
textWhat is the image of a set defined by a meta-fuction?

lemma func1_1_L17: 
  assumes A1: "f  XY" and A2: "xA. b(x)  X"
  shows "f``({b(x). xA}) = {f`(b(x)). xA}"
proof -
  from A2 have "{b(x). xA}  X" by auto
  with A1 show ?thesis using func_imagedef by auto
qed

textWhat are the values of composition of three functions?

lemma func1_1_L18: assumes A1: "f:AB"  "g:BC"  "h:CD"
  and A2: "xA"
  shows
  "(h O g O f)`(x)  D"
  "(h O g O f)`(x) = h`(g`(f`(x)))"  
proof -
  from A1 have "(h O g O f) : AD"
    using comp_fun by blast
  with A2 show "(h O g O f)`(x)  D" using apply_funtype
    by simp
  from A1 A2 have "(h O g O f)`(x) = h`( (g O f)`(x))"
    using comp_fun comp_fun_apply by blast
  with A1 A2 show "(h O g O f)`(x) = h`(g`(f`(x)))"
    using comp_fun_apply by simp
qed

textA composition of functions is a function. This is a slight
  generalization of standard Isabelle's comp_fun›. 

lemma comp_fun_subset: 
  assumes A1: "g:AB"  and A2: "f:CD" and A3: "B  C"
  shows "f O g : A  D"
proof -
  from A1 A3 have "g:AC" by (rule func1_1_L1B) 
  with A2 show "f O g : A  D" using comp_fun by simp
qed

textThis lemma supersedes the lemma comp_eq_id_iff› 
  in Isabelle/ZF. Contributed by Victor Porton.

lemma comp_eq_id_iff1: assumes A1: "g: BA" and A2: "f: AC"
  shows "(yB. f`(g`(y)) = y)  f O g = id(B)"
proof -
  from assms have "f O g: BC" and "id(B): BB"
    using comp_fun id_type by auto
  then have "(yB. (f O g)`y = id(B)`(y))  f O g = id(B)" 
    by (rule fun_extension_iff)
  moreover from A1 have 
    "yB. (f O g)`y = f`(g`y)" and "yB. id(B)`(y) = y"
    by auto
  ultimately show "(yB. f`(g`y) = y)  f O g = id(B)" by simp
qed
  
textA lemma about a value of a function that is a union of 
  some collection of functions.

lemma fun_Union_apply: assumes A1: "F : XY" and 
  A2: "fF" and A3: "f:AB" and A4: "xA"
  shows "(F)`(x) = f`(x)"
proof -
  from A3 A4 have "x, f`(x)  f" using apply_Pair
    by simp
  with A2 have "x, f`(x)  F" by auto
  with A1 show "(F)`(x) = f`(x)" using apply_equality
    by simp
qed

subsectionFunctions restricted to a set

textStandard Isabelle/ZF defines the notion restrict(f,A)› 
  of to mean a function (or relation) $f$ restricted to a set.
  This means that if $f$ is a function defined on $X$ and $A$
  is a subset of $X$ then restrict(f,A)› is a function 
  whith the same values as $f$, but whose domain is $A$.
 
textWhat is the inverse image of a set under a restricted fuction?

lemma func1_2_L1: assumes A1: "f:XY" and A2: "BX"
  shows "restrict(f,B)-``(A) = f-``(A)  B"
proof -
  let ?g = "restrict(f,B)"
  from A1 A2 have "?g:BY" 
    using restrict_type2 by simp
  with A2 A1 show "?g-``(A) = f-``(A)  B"
    using func1_1_L15 restrict_if by auto
qed
   
textA criterion for when one function is a restriction of another.
  The lemma below provides a result useful in the actual proof of the 
  criterion and applications.

lemma func1_2_L2: 
  assumes A1: "f:XY" and A2: "g  AZ" 
  and A3: "AX" and A4: "f  A×Z = g"
  shows "xA. g`(x) = f`(x)"
proof
  fix x assume "xA"
  with A2 have "x,g`(x)  g" using apply_Pair by simp
  with A4 A1 show "g`(x) = f`(x)"  using apply_iff by auto 
qed

textHere is the actual criterion.

lemma func1_2_L3: 
  assumes A1: "f:XY" and A2: "g:AZ" 
  and A3: "AX" and A4: "f  A×Z = g"
  shows "g = restrict(f,A)"
proof
  from A4 show "g  restrict(f, A)" using restrict_iff by auto
  show "restrict(f, A)  g"
  proof
    fix z assume A5:"z  restrict(f,A)"
    then obtain x y where D1:"zf  xA   z = x,y"
      using restrict_iff by auto
    with A1 have "y = f`(x)" using apply_iff by auto
    with A1 A2 A3 A4 D1 have "y = g`(x)" using func1_2_L2 by simp
    with A2 D1 show "zg" using apply_Pair by simp
  qed
qed

textWhich function space a restricted function belongs to?

lemma func1_2_L4: 
  assumes A1: "f:XY" and A2: "AX" and A3: "xA. f`(x)  Z"
  shows "restrict(f,A) : AZ"
proof -
  let ?g = "restrict(f,A)"
  from A1 A2 have "?g : AY" 
    using restrict_type2 by simp
  moreover { 
    fix x assume "xA" 
    with A1 A3 have "?g`(x)  Z" using restrict by simp}
  ultimately show ?thesis by (rule Pi_type)
qed

textA simpler case of func1_2_L4›, where
  the range of the original and restricted function are the same.

corollary restrict_fun: assumes A1: "f:XY" and A2: "AX"
  shows "restrict(f,A) : A  Y"
proof -
  from assms have "xA. f`(x)  Y" using apply_funtype
    by auto
  with assms show ?thesis using func1_2_L4 by simp
qed

textA function restricted to its domain is itself.

lemma restrict_domain: assumes "f:XY"
  shows "restrict(f,X) = f"
proof - 
  have "xX. restrict(f,X)`(x) = f`(x)" using restrict by simp
  with assms show ?thesis using func_eq restrict_fun by blast
qed

textSuppose a function $f:X\rightarrow Y$ is defined by an expression $q$, i.e.
  $f = \{\langle x,y\rangle : x\in X\}$. Then a function that is defined by the same expression,
   but on a smaller set is the same as the restriction of $f$ to that smaller set.

lemma restrict_def_alt: assumes "AX"
  shows "restrict({x,q(x). xX},A) = {x,q(x). xA}"
proof -
  let ?Y = "{q(x). xX}"
  let ?f = "{x,q(x). xX}"
  have "xX. q(x)?Y" by blast
  with assms have "?f:X?Y" using ZF_fun_from_total by simp
  with assms have "restrict(?f,A):A?Y" using restrict_fun by simp
  moreover 
  from assms have "xA. q(x)?Y" by blast
  then have "{x,q(x). xA}:A?Y" using ZF_fun_from_total by simp
  moreover from assms have 
    "xA. restrict(?f,A)`(x) = {x,q(x). xA}`(x)"  
    using restrict ZF_fun_from_tot_val1 by auto
  ultimately show ?thesis by (rule func_eq)
qed

textA composition of two functions is the same as 
  composition with a restriction.

lemma comp_restrict: 
  assumes A1: "f : AB" and A2: "g : X  C" and A3: "BX"
  shows "g O f = restrict(g,B) O f"
proof -
  from assms have "g O f : A  C" using comp_fun_subset
    by simp
  moreover from assms have "restrict(g,B) O f : A  C"
    using restrict_fun comp_fun by simp
  moreover from A1 have 
    "xA. (g O f)`(x) = (restrict(g,B) O f)`(x)"
    using comp_fun_apply apply_funtype restrict
    by simp
  ultimately show "g O f = restrict(g,B) O f"
    by (rule func_eq)
qed

textA way to look at restriction. Contributed by Victor Porton.

lemma right_comp_id_any: shows "r O id(C) = restrict(r,C)"
  unfolding restrict_def by auto
 
subsectionConstant functions

textConstant functions are trivial, but still we need to 
  prove some properties to shorten proofs.

textWe define constant($=c$) functions on a set $X$ 
  in a natural way as ConstantFunction$(X,c)$.

definition
  "ConstantFunction(X,c)  X×{c}"

textConstant function is a function (i.e. belongs to a function space).

lemma func1_3_L1: 
  assumes A1: "cY" shows "ConstantFunction(X,c) : XY"
proof -
   from A1 have "X×{c} = { x,y  X×Y. c = y}" 
     by auto
   with A1 show ?thesis using func1_1_L11A ConstantFunction_def
     by simp
qed

textConstant function is equal to the constant on its domain.

lemma func1_3_L2: assumes A1: "xX"
  shows "ConstantFunction(X,c)`(x) = c"
proof -
  have "ConstantFunction(X,c)  X{c}"
    using func1_3_L1 by simp
  moreover from A1 have "x,c  ConstantFunction(X,c)"
    using ConstantFunction_def by simp
  ultimately show ?thesis using apply_iff by simp
qed

textAnother way of looking at the constant function - it's a set of pairs
  $\langle x,c\rangle$ as $x$ ranges over $X$. 

lemma const_fun_def_alt: shows "ConstantFunction(X,c) = {x,c. xX}"
  unfolding ConstantFunction_def by auto

textIf $c\in A$ then the inverse image of $A$ by the constant function $x\mapsto c$ 
  is the whole domain. 

lemma const_vimage_domain: assumes "cA" 
  shows "ConstantFunction(X,c)-``(A) = X"
proof -
  let ?C = "ConstantFunction(X,c)"
  have "?C-``(A) = {xX. ?C`(x)  A}" using func1_3_L1 func1_1_L15 
    by blast
  with assms show ?thesis using func1_3_L2 by simp
qed

textIf $c$ is not an element of $A$  then the inverse image of $A$ by the constant 
  function $x\mapsto c$ is empty. 

lemma const_vimage_empty: assumes "cA"
  shows "ConstantFunction(X,c)-``(A) = "
proof -
  let ?C = "ConstantFunction(X,c)"
  have "?C-``(A) = {xX. ?C`(x)  A}" using func1_3_L1 func1_1_L15 
    by blast
  with assms show ?thesis using func1_3_L2 by simp
qed

subsectionInjections, surjections, bijections etc.

textIn this section we prove the properties of the spaces
  of injections, surjections and bijections that we can't find in the
  standard Isabelle's Perm.thy›.

textFor injections the image a difference of two sets is
  the difference of images

lemma inj_image_dif: 
  assumes A1: "f  inj(A,B)" and A2: "C  A"
  shows "f``(AC) = f``(A)f``(C)"
proof
  show "f``(AC)  f``(A)f``(C)"
  proof
    fix y assume A3: "y  f``(AC)"
    from A1 have "f:AB" using inj_def by simp
    moreover have "AC  A" by auto
    ultimately have "f``(AC) = {f`(x). x  AC}"
      using func_imagedef by simp
    with A3 obtain x where I: "f`(x) = y" and "x  AC" 
      by auto
    hence "xA" by auto
    with f:AB I have "y  f``(A)"
      using func_imagedef by auto
    moreover have "y   f``(C)"
    proof -
      { assume "y  f``(C)"
	with A2 f:AB obtain x0 
	  where II: "f`(x0) = y" and "x0  C"
	  using func_imagedef by auto
	with A1 A2 I xA have
	  "f  inj(A,B)" "f`(x) = f`(x0)"  "xA" "x0  A"
	  by auto
	then have "x = x0" by (rule inj_apply_equality)
	with x  AC x0  C have False by simp
      } thus ?thesis by auto
    qed
    ultimately show "y  f``(A)f``(C)" by simp
  qed
  from A1 A2 show "f``(A)f``(C)  f``(AC)"
    using inj_def diff_image_diff by auto
qed

textFor injections the image of intersection is the intersection of images.

lemma inj_image_inter: assumes A1: "f  inj(X,Y)" and A2: "AX" "BX"
  shows "f``(AB) = f``(A)  f``(B)"
proof
  show "f``(AB)  f``(A)  f``(B)" using image_Int_subset by simp
  { from A1 have "f:XY" using inj_def by simp 
    fix y assume "y  f``(A)  f``(B)"
    then have "y  f``(A)" and  "y  f``(B)" by auto
    with A2 f:XY obtain xA xB where 
    "xA  A" "xB  B" and I: "y = f`(xA)"  "y = f`(xB)"
      using func_imagedef by auto
    with A2 have "xA  X" "xB  X" and " f`(xA) =  f`(xB)" by auto 
    with A1 have "xA = xB" using inj_def by auto
    with xA  A xB  B have "f`(xA)  {f`(x). x  AB}" by auto
    moreover from A2 f:XY have "f``(AB) = {f`(x). x  AB}"
      using func_imagedef by blast
    ultimately have "f`(xA)  f``(AB)" by simp 
    with I have "y  f``(AB)" by simp 
  } thus "f``(A)  f``(B)  f``(A  B)" by auto
qed

textFor surjection from $A$ to $B$ the image of 
  the domain is $B$.

lemma surj_range_image_domain: assumes A1: "f  surj(A,B)"
  shows "f``(A) = B"
proof -
  from A1 have "f``(A) = range(f)" 
    using surj_def range_image_domain by auto
  with A1 show "f``(A) = B"  using surj_range
    by simp
qed

textSurjections are functions that map the domain onto the codomain.

lemma surj_def_alt: shows "surj(X,Y) = {fXY. f``(X) = Y}"
proof
  show "surj(X,Y)  {fXY. f``(X) = Y}"
    using surj_range_image_domain unfolding surj_def by auto
  show "{fXY. f``(X) = Y}  surj(X,Y)"
    using range_image_domain fun_is_surj by auto
qed

textBijections are functions that preserve complements.

lemma bij_def_alt: 
  shows "bij(X,Y) = {fXY. APow(X). f``(XA) = Yf``(A)}"
proof
  let ?R = "{fXY. APow(X). f``(XA) = Yf``(A)}"
  show "bij(X,Y)  ?R"
    using inj_image_dif surj_range_image_domain surj_is_fun
    unfolding bij_def by auto
  { fix f assume "f?R"
    hence "f:XY" and I: "APow(X). f``(XA) = Yf``(A)"
      by auto
    { fix x1 x2 assume "x1X" "x2X" "f`(x1) = f`(x2)"
      with f:XY have 
        "f``{x1} = {f`(x1)}" "f``{x2} = {f`(x2)}" "f``{x1} = f``{x2}"
        using singleton_image by simp_all
      { assume "x1x2"
        from f:XY have "f``(X{x1}) = {f`(t). tX{x1}}"
          using func_imagedef by blast
        with I x2X x1X x1x2 f``{x1} = f``{x2} 
          have "f`(x2)  Yf``{x2}" by auto
        with f``{x2} = {f`(x2)} have False by auto
      } hence "x1=x2" by auto
    } with f:XY have "finj(X,Y)" unfolding inj_def 
      by auto
    moreover
    from I have "f``(X) = Yf``()" by blast
    with f:XY have "fsurj(X,Y)" using surj_def_alt by simp
    ultimately have "f  bij(X,Y)" unfolding bij_def by simp
  } thus "?R  bij(X,Y)" by auto
qed

textFor injections the inverse image of an image is the same set.

lemma inj_vimage_image: assumes "f  inj(X,Y)" and "AX"
  shows "f-``(f``(A)) = A"
proof -
  have "f-``(f``(A)) = (converse(f) O f)``(A)" 
    using vimage_converse image_comp by simp
  with assms show ?thesis using left_comp_inverse image_id_same
    by simp
qed

textFor surjections the image of an inverse image is the same set.

lemma surj_image_vimage: assumes A1: "f  surj(X,Y)" and A2: "AY"
  shows "f``(f-``(A)) = A"
proof -
  have "f``(f-``(A)) = (f O converse(f))``(A)"
    using vimage_converse image_comp by simp
  with assms show ?thesis using right_comp_inverse image_id_same
    by simp
qed

textA lemma about how a surjection maps collections of subsets in domain and range.

lemma surj_subsets: assumes A1: "f  surj(X,Y)" and A2: "B  Pow(Y)"
  shows "{ f``(U). U  {f-``(V). VB} } = B"
proof
  { fix W assume "W  { f``(U). U  {f-``(V). VB} }"
    then obtain U where I: "U  {f-``(V). VB}" and II: "W = f``(U)" by auto
    then obtain V where "VB" and "U = f-``(V)" by auto
    with II have "W = f``(f-``(V))" by simp
    moreover from assms VB have "f  surj(X,Y)" and "VY" by auto 
    ultimately have "W=V" using surj_image_vimage by simp
    with VB have "W  B" by simp 
  } thus "{ f``(U). U  {f-``(V). VB} }  B" by auto
  { fix W assume "WB"
    let ?U = "f-``(W)"
    from WB have "?U  {f-``(V). VB}" by auto
    moreover from A1 A2 WB have "W = f``(?U)" using surj_image_vimage by auto  
    ultimately have "W  { f``(U). U  {f-``(V). VB} }" by auto 
  } thus "B  { f``(U). U  {f-``(V). VB} }" by auto
qed

textRestriction of an bijection to a set without a point
  is a a bijection.

lemma bij_restrict_rem: 
  assumes A1: "f  bij(A,B)" and A2: "aA"
  shows "restrict(f, A{a})  bij(A{a}, B{f`(a)})"
proof -
  let ?C = "A{a}"
  from A1 have "f  inj(A,B)"  "?C  A"
    using bij_def by auto
  then have "restrict(f,?C)  bij(?C, f``(?C))"
    using restrict_bij by simp
  moreover have "f``(?C) =  B{f`(a)}"
  proof -
    from A2 f  inj(A,B) have "f``(?C) = f``(A)f``{a}"
      using inj_image_dif by simp
    moreover from A1 have "f``(A) = B" 
      using bij_def surj_range_image_domain by auto
    moreover from A1 A2 have "f``{a} = {f`(a)}"
      using bij_is_fun singleton_image by blast
    ultimately show "f``(?C) =  B{f`(a)}" by simp
  qed
  ultimately show ?thesis by simp
qed

textThe domain of a bijection between $X$ and $Y$ is $X$.

lemma domain_of_bij: 
  assumes A1: "f  bij(X,Y)" shows "domain(f) = X"
proof -
  from A1 have "f:XY" using bij_is_fun by simp
  then show "domain(f) = X" using func1_1_L1 by simp
qed

textThe value of the inverse of an injection on a point of the image 
  of a set belongs to that set.

lemma inj_inv_back_in_set: 
  assumes A1: "f  inj(A,B)" and A2: "CA" and A3: "y  f``(C)"
  shows 
  "converse(f)`(y)  C"
  "f`(converse(f)`(y)) = y"
proof -
  from A1 have I: "f:AB" using inj_is_fun by simp
  with A2 A3 obtain x where II: "xC"   "y = f`(x)"
    using func_imagedef by auto
  with A1 A2 show "converse(f)`(y)  C" using left_inverse
    by auto
  from A1 A2 I II show "f`(converse(f)`(y)) = y"
    using func1_1_L5A right_inverse by auto
qed

textFor a bijection between $Y$ and $X$ and a set $A\subseteq X$ 
  an element $y\in Y$ is in the image $f(A)$ if and only if $f^{-1}(y)$ 
  is an element of $A$. Note this is false with the weakened assumption that 
  $f$ is an injection, for example consider $f:\{ 0,1\}\rightarrow \mathbb{N}, f(n)= n+1$
  and $y=3$. Then $f^{-1}:\{1, 2\}\rightarrow \{ 0,1\}$ and (since $3$ is not in the domain 
  of the inverse function) $f^{-1}(3) = \emptyset = 0 \in \{ 0,1\}$, but $3$ is not in the 
  image $f(\{ 0,1\})$.  

lemma bij_val_image_vimage: assumes "f  bij(X,Y)" "AX" "yY"
  shows "yf``(A)  converse(f)`(y)  A"
proof 
  assume "yf``(A)"
  with assms(1,2) show "converse(f)`(y)  A" 
    unfolding bij_def using inj_inv_back_in_set by blast
next
  assume "converse(f)`(y)  A"
  with assms(1,3) have "y{f`(x). xA}" using right_inverse_bij 
    by force
  with assms(1,2) show "yf``(A)" using bij_is_fun func_imagedef 
    by force
qed

textFor injections if a value at a point 
  belongs to the image of a set, then the point
  belongs to the set.

lemma inj_point_of_image: 
  assumes A1: "f  inj(A,B)" and A2: "CA" and
  A3: "xA" and A4: "f`(x)  f``(C)"
  shows "x  C"
proof -
  from A1 A2 A4 have "converse(f)`(f`(x))  C"
    using inj_inv_back_in_set by simp
  moreover from A1 A3 have "converse(f)`(f`(x)) = x"
    using left_inverse_eq by simp
  ultimately show "x  C" by simp
qed

textFor injections the image of intersection is 
  the intersection of images.

lemma inj_image_of_Inter: assumes A1: "f  inj(A,B)" and
  A2: "I" and A3: "iI. P(i)  A"
  shows "f``(iI. P(i)) = ( iI. f``(P(i)) )"
proof
  from A1 A2 A3 show "f``(iI. P(i))  ( iI. f``(P(i)) )"
    using inj_is_fun image_of_Inter by auto
  from A1 A2 A3 have "f:AB"  and "( iI. P(i) )  A"
    using inj_is_fun ZF1_1_L7 by auto
  then have I: "f``(iI. P(i)) = { f`(x). x  ( iI. P(i) ) }"
    using func_imagedef by simp
  { fix y assume A4: "y  ( iI. f``(P(i)) )"
    let ?x = "converse(f)`(y)"
    from A2 obtain i0 where "i0  I" by auto
    with A1 A4 have II: "y  range(f)" using inj_is_fun func1_1_L6
      by auto
    with A1 have III: "f`(?x) = y" using right_inverse by simp
    from A1 II have IV: "?x  A" using inj_converse_fun apply_funtype 
      by blast
    { fix i assume "iI"
      with A3 A4 III have "P(i)  A" and "f`(?x)   f``(P(i))" 
	by auto
      with A1 IV have "?x  P(i)" using inj_point_of_image
	by blast
    } then have "iI. ?x  P(i)" by simp
    with A2 I have "f`(?x)  f``( iI. P(i) )"
      by auto
    with III have "y   f``( iI. P(i) )" by simp
  } then show "( iI. f``(P(i)) )   f``( iI. P(i) )"
    by auto
qed

textAn injection is injective onto its range. Suggested by Victor Porton.

lemma inj_inj_range: assumes "f  inj(A,B)"
  shows "f  inj(A,range(f))"
  using assms inj_def range_of_fun by auto


textAn injection is a bijection on its range. Suggested by Victor Porton.

lemma inj_bij_range: assumes "f  inj(A,B)" 
  shows "f  bij(A,range(f))"
proof -
  from assms have "f  surj(A,range(f))" using inj_def fun_is_surj
    by auto
  with assms show ?thesis using inj_inj_range bij_def by simp
qed
  
textA lemma about extending a surjection by one point.

lemma surj_extend_point: 
  assumes A1: "f  surj(X,Y)" and A2: "aX" and
  A3: "g = f  {a,b}"
  shows "g  surj(X{a},Y{b})"
proof -
  from A1 A2 A3 have "g : X{a}  Y{b}"
    using surj_def func1_1_L11D by simp
  moreover have "y  Y{b}. x  X{a}. y = g`(x)"
  proof
    fix y assume "y   Y  {b}"
    then have "y  Y  y = b" by auto
    moreover
    { assume "y  Y"
      with A1 obtain x where "xX" and "y = f`(x)"
	using surj_def by auto
      with A1 A2 A3 have "x   X{a}" and "y = g`(x)"
	using surj_def func1_1_L11D by auto
      then have "x  X{a}. y = g`(x)" by auto }
    moreover
    { assume "y = b"
      with A1 A2 A3 have "y = g`(a)"
	using surj_def func1_1_L11D by auto
      then have "x  X{a}. y = g`(x)" by auto }
    ultimately show "x  X{a}. y = g`(x)"
      by auto
  qed
  ultimately show "g  surj(X{a},Y{b})"
    using surj_def by auto
qed

textA lemma about extending an injection by one point. 
  Essentially the same as standard Isabelle's inj_extend›.


lemma inj_extend_point: assumes "f  inj(X,Y)" "aX" "bY"
  shows "(f  {a,b})  inj(X{a},Y{b})"
proof -
  from assms have "cons(a,b,f)  inj(cons(a, X), cons(b, Y))"
    using assms inj_extend by simp
  moreover have "cons(a,b,f) = f  {a,b}" and
    "cons(a, X) = X{a}" and "cons(b, Y) = Y{b}"
    by auto
  ultimately show ?thesis by simp
qed

textA lemma about extending a bijection by one point.

lemma bij_extend_point: assumes "f  bij(X,Y)" "aX" "bY"
  shows "(f  {a,b})  bij(X{a},Y{b})"
  using assms surj_extend_point inj_extend_point bij_def
  by simp

textA quite general form of the $a^{-1}b = 1$ 
  implies $a=b$ law.

lemma comp_inv_id_eq: 
  assumes A1: "converse(b) O a = id(A)" and
  A2: "a  A×B" "b  surj(A,B)"
  shows "a = b"
proof -
  from A1 have "(b O converse(b)) O a = b O id(A)"
    using comp_assoc by simp
  with A2 have "id(B) O a = b O id(A)" 
    using right_comp_inverse by simp
  moreover
  from A2 have "a  A×B" and "b  A×B"
    using surj_def fun_subset_prod
    by auto
  then have "id(B) O a = a" and "b O id(A) = b"
    using left_comp_id right_comp_id by auto
  ultimately show "a = b" by simp
qed

textA special case of comp_inv_id_eq› - 
  the $a^{-1}b = 1$ implies $a=b$ law for bijections.

lemma comp_inv_id_eq_bij: 
  assumes A1: "a  bij(A,B)" "b  bij(A,B)" and
  A2: "converse(b) O a = id(A)"
  shows "a = b"
proof -
  from A1 have  "a  A×B" and "b  surj(A,B)"
    using bij_def surj_def fun_subset_prod
    by auto
  with A2 show "a = b" by (rule comp_inv_id_eq)
qed

textConverse of a converse of a bijection is the same bijection. 
This is a special case of converse_converse› from standard Isabelle's 
equalities› theory where it is proved for relations.

lemma bij_converse_converse: assumes "a  bij(A,B)" 
  shows "converse(converse(a)) = a"
proof -
  from assms have "a  A×B" using bij_def surj_def fun_subset_prod by simp
  then show ?thesis using converse_converse by simp
qed

textIf a composition of bijections is identity, then one is the inverse
  of the other.

lemma comp_id_conv: assumes A1: "a  bij(A,B)" "b  bij(B,A)" and
  A2: "b O a = id(A)"
  shows "a = converse(b)" and "b = converse(a)"
proof -
  from A1 have "a  bij(A,B)" and "converse(b)  bij(A,B)" using bij_converse_bij 
    by auto
  moreover from assms have "converse(converse(b)) O a = id(A)" 
    using bij_converse_converse by simp
  ultimately show "a = converse(b)" by (rule comp_inv_id_eq_bij)
  with assms show "b = converse(a)" using bij_converse_converse by simp
qed

textA version of comp_id_conv› with weaker assumptions.

lemma comp_conv_id: assumes A1: "a  bij(A,B)" and A2: "b:BA" and
  A3: "xA. b`(a`(x)) = x"
  shows "b  bij(B,A)" and  "a = converse(b)" and "b = converse(a)"
proof -
  have "b  surj(B,A)"
  proof -
    have "xA. yB. b`(y) = x"
    proof -
      { fix x assume "xA"
        let ?y = "a`(x)"
        from A1 A3 xA have "?yB" and "b`(?y) = x" 
          using bij_def inj_def apply_funtype by auto
        hence "yB. b`(y) = x" by auto
      } thus ?thesis by simp 
    qed
    with A2 show "b  surj(B,A)" using surj_def by simp
  qed
  moreover have "b  inj(B,A)"
  proof -
    have "wB.yB. b`(w) = b`(y)  w=y"
    proof -
      { fix w y assume "wB"  "yB" and I: "b`(w) = b`(y)"
        from A1 have "a  surj(A,B)" unfolding bij_def by simp
        with wB obtain xw where "xw  A" and II: "a`(xw) = w"
          using surj_def by auto
        with I have "b`(a`(xw)) = b`(y)" by simp 
        moreover from a  surj(A,B) yB obtain xy where 
          "xy  A" and III: "a`(xy) = y"
          using surj_def by auto
        moreover from A3 xw  A  xy  A have "b`(a`(xw)) = xw" and  "b`(a`(xy)) = xy"
          by auto
        ultimately have "xw = xy" by simp
        with II III have "w=y" by simp 
      } thus ?thesis by auto  
    qed
    with A2 show "b  inj(B,A)" using inj_def by auto
  qed
  ultimately show "b  bij(B,A)" using bij_def by simp
  from assms have "b O a = id(A)" using bij_def inj_def comp_eq_id_iff1 by auto
  with A1 b  bij(B,A) show "a = converse(b)" and "b = converse(a)"
    using comp_id_conv by auto
qed  
 
textFor a surjection the union if images of singletons
  is the whole range.

lemma surj_singleton_image: assumes A1: "f  surj(X,Y)"
  shows "(xX. {f`(x)}) = Y"
proof
  from A1 show "(xX. {f`(x)})  Y"
    using surj_def apply_funtype by auto
next 
  { fix y assume "y  Y"
    with A1 have "y  (xX. {f`(x)})"
      using surj_def by auto
  } then show  "Y  (xX. {f`(x)})" by auto
qed

subsectionFunctions of two variables

textIn this section we consider functions whose domain is a cartesian product
  of two sets. Such functions are called functions of two variables (although really 
  in ZF all functions admit only one argument). 
  For every function of two variables we can define families of 
  functions of one variable by fixing the other variable. This section 
  establishes basic definitions and results for this concept.


textWe can create functions of two variables by combining functions of one variable.

lemma cart_prod_fun: assumes "f1:X1Y1"  "f2:X2Y2" and
  "g = {p,f1`(fst(p)),f2`(snd(p)). p  X1×X2}"
  shows "g: X1×X2  Y1×Y2" using assms apply_funtype  ZF_fun_from_total by simp

textA reformulation of cart_prod_fun› above in a sligtly different notation.

lemma prod_fun:
  assumes "f:X1X2"  "g:X3X4"
  shows "{x,y,f`x,g`y. x,yX1×X3}:X1×X3X2×X4" 
proof -
  have "{x,y,f`x,g`y. x,yX1×X3} = {p,f`(fst(p)),g`(snd(p)). p  X1×X3}"
    by auto
  with assms show ?thesis using cart_prod_fun by simp 
qed

textProduct of two surjections is a surjection.

theorem prod_functions_surj:
  assumes "fsurj(A,B)" "gsurj(C,D)"
  shows "{a1,a2,f`a1,g`a2.a1,a2A×C}  surj(A×C,B×D)"
proof -
  let ?h = "{x, y, f`(x), g`(y) . x,y  A × C}"
  from assms have fun: "f:AB""g:CD" unfolding surj_def by auto
  then have pfun: "?h : A × C  B × D" using prod_fun by auto
  {
    fix b assume "bB×D"
    then obtain b1 b2 where "b=b1,b2" "b1B" "b2D" by auto
    with assms obtain a1 a2 where "f`(a1)=b1" "g`(a2)=b2" "a1A" "a2C" 
      unfolding surj_def by blast
    hence "a1,a2,b1,b2  ?h" by auto
    with pfun have "?h`a1,a2=b1,b2" using apply_equality by auto
    with b=b1,b2 a1A a2C have "aA×C. ?h`(a)=b" 
      by auto
  } hence "bB×D. aA×C. ?h`(a) = b" by auto
  with pfun show ?thesis unfolding surj_def by auto
qed


textFor a function of two variables created from functions of one variable as in 
  cart_prod_fun› above, the inverse image of a cartesian product of sets is the 
  cartesian product of inverse images.

lemma cart_prod_fun_vimage: assumes "f1:X1Y1"  "f2:X2Y2" and
  "g = {p,f1`(fst(p)),f2`(snd(p)). p  X1×X2}"
  shows "g-``(A1×A2) = f1-``(A1) × f2-``(A2)"
proof -
  from assms have "g: X1×X2  Y1×Y2" using cart_prod_fun 
    by simp
  then have "g-``(A1×A2) = {p  X1×X2. g`(p)  A1×A2}" using func1_1_L15 
    by simp
  with assms g: X1×X2  Y1×Y2 show "g-``(A1×A2) = f1-``(A1) × f2-``(A2)" 
    using ZF_fun_from_tot_val func1_1_L15 by auto
qed
  
textFor a function of two variables defined on $X\times Y$, if we fix an 
  $x\in X$ we obtain a function on $Y$.
  Note that if domain(f)› is $X\times Y$, range(domain(f))› 
  extracts $Y$ from $X\times Y$.

definition
  "Fix1stVar(f,x)  {y,f`x,y. y  range(domain(f))}"
  
textFor every $y\in Y$ we can fix the second variable in a binary function
  $f: X\times Y \rightarrow Z$ to get a function on $X$.

definition
  "Fix2ndVar(f,y)  {x,f`x,y. x  domain(domain(f))}"

textWe defined Fix1stVar› and Fix2ndVar› so that
  the domain of the function is not listed in the arguments, but is recovered 
  from the function. The next lemma is a technical fact that makes it easier
  to use this definition.

lemma fix_var_fun_domain: assumes A1: "f : X×Y  Z"
  shows
  "xX  Fix1stVar(f,x) = {y,f`x,y. y  Y}"
  "yY  Fix2ndVar(f,y) = {x,f`x,y. x  X}"
proof -
  from A1 have I: "domain(f) = X×Y" using func1_1_L1 by simp
  { assume "xX"
    with I have "range(domain(f)) = Y" by auto
    then have "Fix1stVar(f,x) = {y,f`x,y. y  Y}"
      using Fix1stVar_def by simp
  } then show "xX  Fix1stVar(f,x) = {y,f`x,y. y  Y}"
    by simp
  { assume "yY"
    with I have "domain(domain(f)) = X" by auto
    then have "Fix2ndVar(f,y) = {x,f`x,y. x  X}"
      using Fix2ndVar_def by simp
  } then show "yY  Fix2ndVar(f,y) = {x,f`x,y. x  X}"
    by simp
qed
    
textIf we fix the first variable, we get a function of the second variable.

lemma fix_1st_var_fun: assumes A1: "f : X×Y  Z" and A2: "xX"
  shows "Fix1stVar(f,x) : Y  Z"
proof -
  from A1 A2 have "yY. f`x,y  Z"
    using apply_funtype by simp
  then have "{y,f`x,y. y  Y} :  Y  Z" using ZF_fun_from_total by simp
  with A1 A2 show "Fix1stVar(f,x) : Y  Z" using fix_var_fun_domain by simp
qed

textIf we fix the second variable, we get a function of the first
  variable.

lemma fix_2nd_var_fun: assumes A1: "f : X×Y  Z" and A2: "yY"
  shows "Fix2ndVar(f,y) : X  Z"
proof -
  from A1 A2 have "xX. f`x,y  Z"
    using apply_funtype by simp
  then have "{x,f`x,y. x  X} :  X  Z"
    using ZF_fun_from_total by simp
  with A1 A2 show "Fix2ndVar(f,y) : X  Z"
    using fix_var_fun_domain by simp 
qed

textWhat is the value of Fix1stVar(f,x)› at $y\in Y$
  and the value of Fix2ndVar(f,y)› at $x\in X$"?

lemma fix_var_val: 
  assumes A1: "f : X×Y  Z" and A2: "xX"  "yY"
  shows 
  "Fix1stVar(f,x)`(y) = f`x,y"
  "Fix2ndVar(f,y)`(x) = f`x,y"
proof -
  let ?f1 = "{y,f`x,y. y  Y}"
  let ?f2 = "{x,f`x,y. x  X}"
  from A1 A2 have I:
    "Fix1stVar(f,x) = ?f1"
    "Fix2ndVar(f,y) = ?f2"
    using fix_var_fun_domain by auto
  moreover from A1 A2 have
    "Fix1stVar(f,x) : Y  Z"
    "Fix2ndVar(f,y) : X  Z"
    using fix_1st_var_fun fix_2nd_var_fun by auto
  ultimately have "?f1 : Y  Z" and  "?f2 : X  Z"
    by auto
  with A2 have "?f1`(y) = f`x,y" and "?f2`(x) = f`x,y"
    using ZF_fun_from_tot_val by auto
  with I show
    "Fix1stVar(f,x)`(y) = f`x,y"
    "Fix2ndVar(f,y)`(x) = f`x,y"
    by auto
qed

textFixing the second variable commutes with restrictig the domain.

lemma fix_2nd_var_restr_comm: 
  assumes A1: "f : X×Y  Z" and A2: "yY" and A3: "X1  X"
  shows "Fix2ndVar(restrict(f,X1×Y),y) = restrict(Fix2ndVar(f,y),X1)"
proof -
  let ?g = "Fix2ndVar(restrict(f,X1×Y),y)"
  let ?h = "restrict(Fix2ndVar(f,y),X1)"
  from A3 have I: "X1×Y  X×Y" by auto
  with A1 have II: "restrict(f,X1×Y) : X1×Y  Z"
    using restrict_type2 by simp
  with A2 have "?g : X1  Z"
    using fix_2nd_var_fun by simp
  moreover
  from A1 A2 have III: "Fix2ndVar(f,y) : X  Z"
    using fix_2nd_var_fun by simp
  with A3 have "?h : X1  Z"
    using restrict_type2 by simp
  moreover
  { fix z assume A4: "z  X1"
    with A2 I II have "?g`(z) = f`z,y"
      using restrict fix_var_val by simp
    also from A1 A2 A3 A4 have "f`z,y = ?h`(z)"
      using restrict fix_var_val by auto
    finally have "?g`(z) = ?h`(z)" by simp
  } then have "z  X1. ?g`(z) = ?h`(z)" by simp
  ultimately show "?g = ?h" by (rule func_eq)
qed

textThe next lemma expresses the inverse image of a set by function with fixed 
first variable in terms of the original function.

lemma fix_1st_var_vimage:
  assumes A1: "f : X×Y  Z" and A2: "xX" 
  shows "Fix1stVar(f,x)-``(A) = {yY. x,y  f-``(A)}"
proof -
  from assms have "Fix1stVar(f,x)-``(A) = {yY. Fix1stVar(f,x)`(y)  A}"
    using fix_1st_var_fun func1_1_L15 by blast
  with assms show ?thesis using fix_var_val func1_1_L15 by auto
qed

textThe next lemma expresses the inverse image of a set by function with fixed 
second variable in terms of the original function.

lemma fix_2nd_var_vimage:
  assumes A1: "f : X×Y  Z" and A2: "yY" 
  shows "Fix2ndVar(f,y)-``(A) = {xX. x,y  f-``(A)}"
proof -
  from assms have I: "Fix2ndVar(f,y)-``(A) = {xX. Fix2ndVar(f,y)`(x)  A}"
    using fix_2nd_var_fun func1_1_L15 by blast
  with assms show ?thesis using fix_var_val func1_1_L15 by auto
qed

end