2015-02-01 16:39:47 +00:00
|
|
|
/-
|
|
|
|
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
|
|
|
Module: data.sigma
|
|
|
|
Author: Leonardo de Moura, Jeremy Avigad, Floris van Doorn
|
|
|
|
|
|
|
|
Sigma types, aka dependent sum.
|
|
|
|
-/
|
2014-12-01 04:34:12 +00:00
|
|
|
import logic.cast
|
2014-11-21 04:22:19 +00:00
|
|
|
open inhabited eq.ops sigma.ops
|
2014-08-20 02:32:44 +00:00
|
|
|
|
2014-08-15 03:12:54 +00:00
|
|
|
namespace sigma
|
2014-11-04 00:22:30 +00:00
|
|
|
universe variables u v
|
|
|
|
variables {A A' : Type.{u}} {B : A → Type.{v}} {B' : A' → Type.{v}}
|
2014-08-15 03:12:54 +00:00
|
|
|
|
2014-11-21 04:22:19 +00:00
|
|
|
definition unpack {C : (Σa, B a) → Type} {u : Σa, B a} (H : C ⟨u.1 , u.2⟩) : C u :=
|
2014-11-09 22:22:42 +00:00
|
|
|
destruct u (λx y H, H) H
|
|
|
|
|
2014-11-04 00:22:30 +00:00
|
|
|
theorem dpair_heq {a : A} {a' : A'} {b : B a} {b' : B' a'}
|
2014-12-11 23:48:09 +00:00
|
|
|
(HB : B == B') (Ha : a == a') (Hb : b == b') : ⟨a, b⟩ == ⟨a', b'⟩ :=
|
2014-12-20 02:23:08 +00:00
|
|
|
hcongr_arg4 @mk (heq.type_eq Ha) HB Ha Hb
|
2014-11-04 00:22:30 +00:00
|
|
|
|
|
|
|
protected theorem hequal {p : Σa : A, B a} {p' : Σa' : A', B' a'} (HB : B == B') :
|
2014-12-11 23:48:09 +00:00
|
|
|
∀(H₁ : p.1 == p'.1) (H₂ : p.2 == p'.2), p == p' :=
|
2014-11-04 00:22:30 +00:00
|
|
|
destruct p (take a₁ b₁, destruct p' (take a₂ b₂ H₁ H₂, dpair_heq HB H₁ H₂))
|
|
|
|
|
2015-02-24 23:25:02 +00:00
|
|
|
protected definition is_inhabited [instance] [H₁ : inhabited A] [H₂ : inhabited (B (default A))] :
|
2014-08-15 03:12:54 +00:00
|
|
|
inhabited (sigma B) :=
|
2014-12-11 23:48:09 +00:00
|
|
|
inhabited.destruct H₁ (λa, inhabited.destruct H₂ (λb, inhabited.mk ⟨default A, b⟩))
|
feat(library): add constructions of categories, some changes in eq, sigma and path
in eq, add theorem for proof irrelevance and congruence for binary functions
in sigma, add some support for triplets
in path, comment out two unneccesary definitions
in category, add Cat, slice, coslice, product and arrow categories, also add fully bundled approach
2014-09-22 18:48:09 +00:00
|
|
|
|
2014-11-04 00:22:30 +00:00
|
|
|
theorem eq_rec_dpair_commute {C : Πa, B a → Type} {a a' : A} (H : a = a') (b : B a) (c : C a b) :
|
2014-12-11 23:48:09 +00:00
|
|
|
eq.rec_on H ⟨b, c⟩ = ⟨eq.rec_on H b, eq.rec_on (dcongr_arg2 C H rfl) c⟩ :=
|
2014-11-04 00:22:30 +00:00
|
|
|
eq.drec_on H (dpair_eq rfl (!eq.rec_on_id⁻¹))
|
|
|
|
|
2014-10-10 23:33:58 +00:00
|
|
|
variables {C : Πa, B a → Type} {D : Πa b, C a b → Type}
|
feat(library): add constructions of categories, some changes in eq, sigma and path
in eq, add theorem for proof irrelevance and congruence for binary functions
in sigma, add some support for triplets
in path, comment out two unneccesary definitions
in category, add Cat, slice, coslice, product and arrow categories, also add fully bundled approach
2014-09-22 18:48:09 +00:00
|
|
|
|
2014-11-21 04:22:19 +00:00
|
|
|
definition dtrip (a : A) (b : B a) (c : C a b) := ⟨a, b, c⟩
|
|
|
|
definition dquad (a : A) (b : B a) (c : C a b) (d : D a b c) := ⟨a, b, c, d⟩
|
feat(library): add constructions of categories, some changes in eq, sigma and path
in eq, add theorem for proof irrelevance and congruence for binary functions
in sigma, add some support for triplets
in path, comment out two unneccesary definitions
in category, add Cat, slice, coslice, product and arrow categories, also add fully bundled approach
2014-09-22 18:48:09 +00:00
|
|
|
|
2015-01-09 02:47:44 +00:00
|
|
|
definition pr1' [reducible] (x : Σ a, B a) := x.1
|
|
|
|
definition pr2' [reducible] (x : Σ a b, C a b) := x.2.1
|
|
|
|
definition pr3 [reducible] (x : Σ a b, C a b) := x.2.2
|
|
|
|
definition pr3' [reducible] (x : Σ a b c, D a b c) := x.2.2.1
|
|
|
|
definition pr4 [reducible] (x : Σ a b c, D a b c) := x.2.2.2
|
feat(library): add constructions of categories, some changes in eq, sigma and path
in eq, add theorem for proof irrelevance and congruence for binary functions
in sigma, add some support for triplets
in path, comment out two unneccesary definitions
in category, add Cat, slice, coslice, product and arrow categories, also add fully bundled approach
2014-09-22 18:48:09 +00:00
|
|
|
|
2014-09-26 23:36:04 +00:00
|
|
|
theorem dtrip_eq {a₁ a₂ : A} {b₁ : B a₁} {b₂ : B a₂} {c₁ : C a₁ b₁} {c₂ : C a₂ b₂}
|
2014-11-04 00:22:30 +00:00
|
|
|
(H₁ : a₁ = a₂) (H₂ : eq.rec_on H₁ b₁ = b₂) (H₃ : cast (dcongr_arg2 C H₁ H₂) c₁ = c₂) :
|
2014-11-21 04:22:19 +00:00
|
|
|
⟨a₁, b₁, c₁⟩ = ⟨a₂, b₂, c₂⟩ :=
|
2014-11-04 00:22:30 +00:00
|
|
|
dcongr_arg3 dtrip H₁ H₂ H₃
|
feat(library): add constructions of categories, some changes in eq, sigma and path
in eq, add theorem for proof irrelevance and congruence for binary functions
in sigma, add some support for triplets
in path, comment out two unneccesary definitions
in category, add Cat, slice, coslice, product and arrow categories, also add fully bundled approach
2014-09-22 18:48:09 +00:00
|
|
|
|
2014-11-04 00:22:30 +00:00
|
|
|
theorem ndtrip_eq {A B : Type} {C : A → B → Type} {a₁ a₂ : A} {b₁ b₂ : B}
|
|
|
|
{c₁ : C a₁ b₁} {c₂ : C a₂ b₂} (H₁ : a₁ = a₂) (H₂ : b₁ = b₂)
|
2014-11-21 04:22:19 +00:00
|
|
|
(H₃ : cast (congr_arg2 C H₁ H₂) c₁ = c₂) : ⟨a₁, b₁, c₁⟩ = ⟨a₂, b₂, c₂⟩ :=
|
2014-12-12 21:20:27 +00:00
|
|
|
hdcongr_arg3 dtrip H₁ (heq.of_eq H₂) H₃
|
2014-09-26 23:36:04 +00:00
|
|
|
|
2014-11-04 00:22:30 +00:00
|
|
|
theorem ndtrip_equal {A B : Type} {C : A → B → Type} {p₁ p₂ : Σa b, C a b} :
|
2014-12-20 02:23:08 +00:00
|
|
|
∀(H₁ : pr1 p₁ = pr1 p₂) (H₂ : pr2' p₁ = pr2' p₂)
|
|
|
|
(H₃ : eq.rec_on (congr_arg2 C H₁ H₂) (pr3 p₁) = pr3 p₂), p₁ = p₂ :=
|
feat(library): add constructions of categories, some changes in eq, sigma and path
in eq, add theorem for proof irrelevance and congruence for binary functions
in sigma, add some support for triplets
in path, comment out two unneccesary definitions
in category, add Cat, slice, coslice, product and arrow categories, also add fully bundled approach
2014-09-22 18:48:09 +00:00
|
|
|
destruct p₁ (take a₁ q₁, destruct q₁ (take b₁ c₁, destruct p₂ (take a₂ q₂, destruct q₂
|
2014-11-04 00:22:30 +00:00
|
|
|
(take b₂ c₂ H₁ H₂ H₃, ndtrip_eq H₁ H₂ H₃))))
|
feat(library): add constructions of categories, some changes in eq, sigma and path
in eq, add theorem for proof irrelevance and congruence for binary functions
in sigma, add some support for triplets
in path, comment out two unneccesary definitions
in category, add Cat, slice, coslice, product and arrow categories, also add fully bundled approach
2014-09-22 18:48:09 +00:00
|
|
|
|
2014-08-15 03:12:54 +00:00
|
|
|
end sigma
|