2014-11-28 12:06:46 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
2015-06-04 08:51:34 +00:00
|
|
|
|
Author: Leonardo de Moura, Jeremy Avigad, Haitao Zhang
|
2014-11-28 12:06:46 +00:00
|
|
|
|
|
|
|
|
|
General operations on functions.
|
|
|
|
|
-/
|
2015-07-06 14:29:56 +00:00
|
|
|
|
prelude
|
|
|
|
|
import init.prod init.funext init.logic
|
2015-05-23 06:16:36 +00:00
|
|
|
|
|
2014-07-24 23:29:39 +00:00
|
|
|
|
namespace function
|
|
|
|
|
|
2014-10-10 23:33:58 +00:00
|
|
|
|
variables {A : Type} {B : Type} {C : Type} {D : Type} {E : Type}
|
2014-07-24 23:29:39 +00:00
|
|
|
|
|
2015-07-07 23:37:06 +00:00
|
|
|
|
definition compose [reducible] [unfold-full] (f : B → C) (g : A → B) : A → C :=
|
2014-10-10 23:33:58 +00:00
|
|
|
|
λx, f (g x)
|
2014-07-24 23:29:39 +00:00
|
|
|
|
|
2015-07-07 23:37:06 +00:00
|
|
|
|
definition compose_right [reducible] [unfold-full] (f : B → B → B) (g : A → B) : B → A → B :=
|
2015-04-09 17:54:28 +00:00
|
|
|
|
λ b a, f b (g a)
|
|
|
|
|
|
2015-07-07 23:37:06 +00:00
|
|
|
|
definition compose_left [reducible] [unfold-full] (f : B → B → B) (g : A → B) : A → B → B :=
|
2015-04-09 17:54:28 +00:00
|
|
|
|
λ a b, f (g a) b
|
|
|
|
|
|
2015-07-07 23:37:06 +00:00
|
|
|
|
definition id [reducible] [unfold-full] (a : A) : A :=
|
2014-10-10 23:33:58 +00:00
|
|
|
|
a
|
2014-07-24 23:29:39 +00:00
|
|
|
|
|
2015-07-07 23:37:06 +00:00
|
|
|
|
definition on_fun [reducible] [unfold-full] (f : B → B → C) (g : A → B) : A → A → C :=
|
2014-10-10 23:33:58 +00:00
|
|
|
|
λx y, f (g x) (g y)
|
2014-07-24 23:29:39 +00:00
|
|
|
|
|
2015-07-07 23:37:06 +00:00
|
|
|
|
definition combine [reducible] [unfold-full] (f : A → B → C) (op : C → D → E) (g : A → B → D)
|
2015-05-07 05:23:39 +00:00
|
|
|
|
: A → B → E :=
|
2014-10-10 23:33:58 +00:00
|
|
|
|
λx y, op (f x y) (g x y)
|
2014-07-24 23:29:39 +00:00
|
|
|
|
|
2015-07-07 23:37:06 +00:00
|
|
|
|
definition const [reducible] [unfold-full] (B : Type) (a : A) : B → A :=
|
2014-07-29 02:58:57 +00:00
|
|
|
|
λx, a
|
2014-07-28 04:01:59 +00:00
|
|
|
|
|
2015-07-07 23:37:06 +00:00
|
|
|
|
definition dcompose [reducible] [unfold-full] {B : A → Type} {C : Π {x : A}, B x → Type}
|
2014-11-28 12:06:46 +00:00
|
|
|
|
(f : Π {x : A} (y : B x), C y) (g : Πx, B x) : Πx, C (g x) :=
|
2014-07-29 02:58:57 +00:00
|
|
|
|
λx, f (g x)
|
2014-07-24 23:29:39 +00:00
|
|
|
|
|
2015-07-07 23:37:06 +00:00
|
|
|
|
definition swap [reducible] [unfold-full] {C : A → B → Type} (f : Πx y, C x y) : Πy x, C x y :=
|
2014-07-29 02:58:57 +00:00
|
|
|
|
λy x, f x y
|
2014-07-24 23:29:39 +00:00
|
|
|
|
|
2015-02-17 02:52:41 +00:00
|
|
|
|
definition app [reducible] {B : A → Type} (f : Πx, B x) (x : A) : B x :=
|
2014-07-29 02:58:57 +00:00
|
|
|
|
f x
|
2014-07-24 23:29:39 +00:00
|
|
|
|
|
2015-07-07 23:37:06 +00:00
|
|
|
|
definition curry [reducible] [unfold-full] : (A × B → C) → A → B → C :=
|
2015-04-11 23:45:07 +00:00
|
|
|
|
λ f a b, f (a, b)
|
|
|
|
|
|
2015-07-07 23:37:06 +00:00
|
|
|
|
definition uncurry [reducible] [unfold 5] : (A → B → C) → (A × B → C) :=
|
2015-04-11 23:45:07 +00:00
|
|
|
|
λ f p, match p with (a, b) := f a b end
|
|
|
|
|
|
|
|
|
|
theorem curry_uncurry (f : A → B → C) : curry (uncurry f) = f :=
|
|
|
|
|
rfl
|
|
|
|
|
|
|
|
|
|
theorem uncurry_curry (f : A × B → C) : uncurry (curry f) = f :=
|
|
|
|
|
funext (λ p, match p with (a, b) := rfl end)
|
|
|
|
|
|
2014-07-24 23:29:39 +00:00
|
|
|
|
precedence `∘'`:60
|
|
|
|
|
precedence `on`:1
|
|
|
|
|
precedence `$`:1
|
|
|
|
|
|
|
|
|
|
infixr ∘ := compose
|
|
|
|
|
infixr ∘' := dcompose
|
|
|
|
|
infixl on := on_fun
|
|
|
|
|
infixr $ := app
|
|
|
|
|
notation f `-[` op `]-` g := combine f op g
|
|
|
|
|
|
2015-06-04 08:51:34 +00:00
|
|
|
|
lemma left_id (f : A → B) : id ∘ f = f := rfl
|
2015-04-03 22:43:44 +00:00
|
|
|
|
|
2015-06-04 08:51:34 +00:00
|
|
|
|
lemma right_id (f : A → B) : f ∘ id = f := rfl
|
2015-04-03 22:43:44 +00:00
|
|
|
|
|
2015-06-04 08:51:34 +00:00
|
|
|
|
theorem compose.assoc (f : C → D) (g : B → C) (h : A → B) : (f ∘ g) ∘ h = f ∘ (g ∘ h) := rfl
|
|
|
|
|
|
|
|
|
|
theorem compose.left_id (f : A → B) : id ∘ f = f := rfl
|
|
|
|
|
|
|
|
|
|
theorem compose.right_id (f : A → B) : f ∘ id = f := rfl
|
|
|
|
|
|
|
|
|
|
theorem compose_const_right (f : B → C) (b : B) : f ∘ (const A b) = const A (f b) := rfl
|
|
|
|
|
|
2015-06-09 22:30:58 +00:00
|
|
|
|
definition injective [reducible] (f : A → B) : Prop := ∀ ⦃a₁ a₂⦄, f a₁ = f a₂ → a₁ = a₂
|
2015-06-04 08:51:34 +00:00
|
|
|
|
|
|
|
|
|
theorem injective_compose {g : B → C} {f : A → B} (Hg : injective g) (Hf : injective f) :
|
|
|
|
|
injective (g ∘ f) :=
|
|
|
|
|
take a₁ a₂, assume Heq, Hf (Hg Heq)
|
2015-04-17 03:52:18 +00:00
|
|
|
|
|
2015-06-09 22:30:58 +00:00
|
|
|
|
definition surjective [reducible] (f : A → B) : Prop := ∀ b, ∃ a, f a = b
|
2015-04-17 03:52:18 +00:00
|
|
|
|
|
2015-06-04 08:51:34 +00:00
|
|
|
|
theorem surjective_compose {g : B → C} {f : A → B} (Hg : surjective g) (Hf : surjective f) :
|
|
|
|
|
surjective (g ∘ f) :=
|
|
|
|
|
take c,
|
|
|
|
|
obtain b (Hb : g b = c), from Hg c,
|
|
|
|
|
obtain a (Ha : f a = b), from Hf b,
|
|
|
|
|
exists.intro a (eq.trans (congr_arg g Ha) Hb)
|
|
|
|
|
|
|
|
|
|
definition bijective (f : A → B) := injective f ∧ surjective f
|
2015-04-17 03:52:18 +00:00
|
|
|
|
|
2015-06-04 08:51:34 +00:00
|
|
|
|
theorem bijective_compose {g : B → C} {f : A → B} (Hg : bijective g) (Hf : bijective f) :
|
|
|
|
|
bijective (g ∘ f) :=
|
|
|
|
|
obtain Hginj Hgsurj, from Hg,
|
|
|
|
|
obtain Hfinj Hfsurj, from Hf,
|
|
|
|
|
and.intro (injective_compose Hginj Hfinj) (surjective_compose Hgsurj Hfsurj)
|
2015-04-17 03:52:18 +00:00
|
|
|
|
|
2015-06-04 08:51:34 +00:00
|
|
|
|
-- g is a left inverse to f
|
|
|
|
|
definition left_inverse (g : B → A) (f : A → B) : Prop := ∀x, g (f x) = x
|
|
|
|
|
|
2015-07-06 19:17:57 +00:00
|
|
|
|
definition id_of_left_inverse {g : B → A} {f : A → B} : left_inverse g f → g ∘ f = id :=
|
|
|
|
|
assume h, funext h
|
|
|
|
|
|
2015-06-04 08:51:34 +00:00
|
|
|
|
definition has_left_inverse (f : A → B) : Prop := ∃ finv : B → A, left_inverse finv f
|
|
|
|
|
|
|
|
|
|
-- g is a right inverse to f
|
|
|
|
|
definition right_inverse (g : B → A) (f : A → B) : Prop := left_inverse f g
|
|
|
|
|
|
2015-07-06 19:17:57 +00:00
|
|
|
|
definition id_of_righ_inverse {g : B → A} {f : A → B} : right_inverse g f → f ∘ g = id :=
|
|
|
|
|
assume h, funext h
|
|
|
|
|
|
2015-06-04 08:51:34 +00:00
|
|
|
|
definition has_right_inverse (f : A → B) : Prop := ∃ finv : B → A, right_inverse finv f
|
|
|
|
|
|
2015-07-06 19:17:57 +00:00
|
|
|
|
theorem injective_of_left_inverse {g : B → A} {f : A → B} : left_inverse g f → injective f :=
|
2015-04-17 03:52:18 +00:00
|
|
|
|
assume h, take a b, assume faeqfb,
|
2015-07-06 19:17:57 +00:00
|
|
|
|
calc a = g (f a) : by rewrite h
|
|
|
|
|
... = g (f b) : faeqfb
|
|
|
|
|
... = b : by rewrite h
|
|
|
|
|
|
|
|
|
|
theorem injective_of_has_left_inverse {f : A → B} : has_left_inverse f → injective f :=
|
|
|
|
|
assume h, obtain (finv : B → A) (inv : left_inverse finv f), from h,
|
|
|
|
|
injective_of_left_inverse inv
|
2015-06-04 08:51:34 +00:00
|
|
|
|
|
|
|
|
|
theorem right_inverse_of_injective_of_left_inverse {f : A → B} {g : B → A}
|
|
|
|
|
(injf : injective f) (lfg : left_inverse f g) :
|
|
|
|
|
right_inverse f g :=
|
|
|
|
|
take x,
|
|
|
|
|
have H : f (g (f x)) = f x, from lfg (f x),
|
|
|
|
|
injf H
|
2015-04-03 22:43:44 +00:00
|
|
|
|
|
2015-06-04 08:51:34 +00:00
|
|
|
|
theorem surjective_of_has_right_inverse {f : A → B} : has_right_inverse f → surjective f :=
|
2015-04-17 03:52:18 +00:00
|
|
|
|
assume h, take b,
|
2015-06-04 08:51:34 +00:00
|
|
|
|
obtain (finv : B → A) (inv : right_inverse finv f), from h,
|
2015-04-17 03:52:18 +00:00
|
|
|
|
let a : A := finv b in
|
|
|
|
|
have h : f a = b, from calc
|
|
|
|
|
f a = (f ∘ finv) b : rfl
|
2015-06-04 08:51:34 +00:00
|
|
|
|
... = id b : by rewrite inv
|
2015-04-17 03:52:18 +00:00
|
|
|
|
... = b : rfl,
|
|
|
|
|
exists.intro a h
|
2015-05-23 06:16:36 +00:00
|
|
|
|
|
2015-06-04 08:51:34 +00:00
|
|
|
|
theorem left_inverse_of_surjective_of_right_inverse {f : A → B} {g : B → A}
|
|
|
|
|
(surjf : surjective f) (rfg : right_inverse f g) :
|
|
|
|
|
left_inverse f g :=
|
|
|
|
|
take y,
|
|
|
|
|
obtain x (Hx : f x = y), from surjf y,
|
|
|
|
|
calc
|
|
|
|
|
f (g y) = f (g (f x)) : Hx
|
|
|
|
|
... = f x : rfg
|
|
|
|
|
... = y : Hx
|
2015-05-23 06:16:36 +00:00
|
|
|
|
|
2015-06-04 08:51:34 +00:00
|
|
|
|
theorem injective_id : injective (@id A) := take a₁ a₂ H, H
|
2015-05-23 06:16:36 +00:00
|
|
|
|
|
2015-06-04 08:51:34 +00:00
|
|
|
|
theorem surjective_id : surjective (@id A) := take a, exists.intro a rfl
|
2015-05-23 06:16:36 +00:00
|
|
|
|
|
2015-06-04 08:51:34 +00:00
|
|
|
|
theorem bijective_id : bijective (@id A) := and.intro injective_id surjective_id
|
2015-05-23 06:16:36 +00:00
|
|
|
|
|
2014-08-07 23:59:08 +00:00
|
|
|
|
end function
|
2015-02-17 02:52:41 +00:00
|
|
|
|
|
|
|
|
|
-- copy reducible annotations to top-level
|
2015-05-07 05:23:39 +00:00
|
|
|
|
export [reduce-hints] [unfold-hints] function
|