Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 3fcba7d

Browse files
jcommelindigama0
authored andcommitted
feat(logic/basic): add class 'unique' for uniquely inhabited types (#605)
1 parent 41b3fdd commit 3fcba7d

File tree

2 files changed

+80
-1
lines changed

2 files changed

+80
-1
lines changed

src/data/equiv/basic.lean

Lines changed: 25 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -8,7 +8,7 @@ We say two types are equivalent if they are isomorphic.
88
99
Two equivalent types have the same cardinality.
1010
-/
11-
import logic.function data.set.basic data.bool
11+
import logic.function logic.unique data.set.basic data.bool
1212

1313
open function
1414

@@ -471,6 +471,15 @@ def decidable_eq_of_equiv [decidable_eq β] (e : α ≃ β) : decidable_eq α
471471
def inhabited_of_equiv [inhabited β] (e : α ≃ β) : inhabited α :=
472472
⟨e.symm (default _)⟩
473473

474+
def unique_of_equiv [unique β] (e : α ≃ β) : unique α :=
475+
unique.of_surjective e.symm.bijective.2
476+
477+
def unique_congr (e : α ≃ β) : unique α ≃ unique β :=
478+
{ to_fun := λ h, by resetI; exact e.symm.unique_of_equiv,
479+
inv_fun := λ h, by resetI; exact e.unique_of_equiv,
480+
left_inv := λ _, subsingleton.elim _ _,
481+
right_inv := λ _, subsingleton.elim _ _ }
482+
474483
section
475484
open subtype
476485

@@ -706,3 +715,18 @@ instance {α} [subsingleton α] : subsingleton (plift α) := equiv.plift.subsing
706715

707716
instance {α} [decidable_eq α] : decidable_eq (ulift α) := equiv.ulift.decidable_eq
708717
instance {α} [decidable_eq α] : decidable_eq (plift α) := equiv.plift.decidable_eq
718+
719+
def unique_unique_equiv : unique (unique α) ≃ unique α :=
720+
{ to_fun := λ h, h.default,
721+
inv_fun := λ h, { default := h, uniq := λ _, subsingleton.elim _ _ },
722+
left_inv := λ _, subsingleton.elim _ _,
723+
right_inv := λ _, subsingleton.elim _ _ }
724+
725+
def equiv_of_unique_of_unique [unique α] [unique β] : α ≃ β :=
726+
{ to_fun := λ _, default β,
727+
inv_fun := λ _, default α,
728+
left_inv := λ _, subsingleton.elim _ _,
729+
right_inv := λ _, subsingleton.elim _ _ }
730+
731+
def equiv_punit_of_unique [unique α] : α ≃ punit.{v} :=
732+
equiv_of_unique_of_unique

src/logic/unique.lean

Lines changed: 55 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,55 @@
1+
/-
2+
Copyright (c) 2019 Johan Commelin. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Johan Commelin
5+
-/
6+
7+
import logic.basic
8+
9+
universes u v w
10+
11+
variables {α : Sort u} {β : Sort v} {γ : Sort w}
12+
13+
structure unique (α : Sort u) extends inhabited α :=
14+
(uniq : ∀ a:α, a = default)
15+
16+
attribute [class] unique
17+
18+
instance punit.unique : unique punit.{u} :=
19+
{ default := punit.star,
20+
uniq := λ x, punit_eq x _ }
21+
22+
namespace unique
23+
open function
24+
25+
section
26+
27+
variables [unique α]
28+
29+
instance : inhabited α := to_inhabited ‹unique α›
30+
31+
lemma eq_default (a : α) : a = default α := uniq _ a
32+
33+
lemma default_eq (a : α) : default α = a := (uniq _ a).symm
34+
35+
instance : subsingleton α := ⟨λ a b, by rw [eq_default a, eq_default b]⟩
36+
37+
end
38+
39+
protected lemma subsingleton_unique' : ∀ (h₁ h₂ : unique α), h₁ = h₂
40+
| ⟨⟨x⟩, h⟩ ⟨⟨y⟩, _⟩ := by congr; rw [h x, h y]
41+
42+
instance subsingleton_unique : subsingleton (unique α) :=
43+
⟨unique.subsingleton_unique'⟩
44+
45+
def of_surjective {f : α → β} (hf : surjective f) [unique α] : unique β :=
46+
{ default := f (default _),
47+
uniq := λ b,
48+
begin
49+
cases hf b with a ha,
50+
subst ha,
51+
exact congr_arg f (eq_default a)
52+
end }
53+
54+
end unique
55+

0 commit comments

Comments
 (0)