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

Commit f8d8171

Browse files
committed
refactor(logic/relator): turn *_unique and *_total into defs, not classes (#9135)
We had (almost) no instances for these classes and (almost) no lemmas taking these assumptions as TC arguments.
1 parent f1bf7b8 commit f8d8171

File tree

8 files changed

+41
-55
lines changed

8 files changed

+41
-55
lines changed

src/computability/turing_machine.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -605,9 +605,9 @@ theorem reaches₁_eq {σ} {f : σ → option σ} {a b c}
605605
trans_gen.head'_iff.trans (trans_gen.head'_iff.trans $ by rw h).symm
606606

607607
theorem reaches_total {σ} {f : σ → option σ}
608-
{a b c} : reaches f a breaches f a c
608+
{a b c} (hab : reaches f a b) (hac : reaches f a c) :
609609
reaches f b c ∨ reaches f c b :=
610-
refl_trans_gen.total_of_right_unique λ _ _ _, option.mem_unique
610+
refl_trans_gen.total_of_right_unique (λ _ _ _, option.mem_unique) hab hac
611611

612612
theorem reaches₁_fwd {σ} {f : σ → option σ}
613613
{a b c} (h₁ : reaches₁ f a c) (h₂ : b ∈ f a) : reaches f b c :=

src/data/list/forall2.lean

Lines changed: 9 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -93,24 +93,24 @@ lemma forall₂_and_left {r : α → β → Prop} {p : α → Prop} :
9393
| _ (b :: u) := by simp only [map, forall₂_cons_right_iff, forall₂_map_right_iff]
9494

9595
lemma left_unique_forall₂' (hr : left_unique r) :
96-
∀ {a b c}, forall₂ r a b → forall₂ r c b → a = c
96+
∀ {a b c}, forall₂ r a c → forall₂ r b c → a = b
9797
| a₀ nil a₁ forall₂.nil forall₂.nil := rfl
9898
| (a₀ :: l₀) (b :: l) (a₁ :: l₁) (forall₂.cons ha₀ h₀) (forall₂.cons ha₁ h₁) :=
99-
hr.unique ha₀ ha₁ ▸ left_unique_forall₂' h₀ h₁ ▸ rfl
99+
hr ha₀ ha₁ ▸ left_unique_forall₂' h₀ h₁ ▸ rfl
100100

101-
lemma left_unique_forall (hr : left_unique r) : left_unique (forall₂ r) :=
102-
@left_unique_forall₂' _ _ _ hr
101+
lemma _root_.relator.left_unique.forall (hr : left_unique r) : left_unique (forall₂ r) :=
102+
@left_unique_forall₂' _ _ _ hr
103103

104104
lemma right_unique_forall₂' (hr : right_unique r) : ∀ {a b c}, forall₂ r a b → forall₂ r a c → b = c
105105
| nil a₀ a₁ forall₂.nil forall₂.nil := rfl
106106
| (b :: l) (a₀ :: l₀) (a₁ :: l₁) (forall₂.cons ha₀ h₀) (forall₂.cons ha₁ h₁) :=
107-
hr.unique ha₀ ha₁ ▸ right_unique_forall₂' h₀ h₁ ▸ rfl
107+
hr ha₀ ha₁ ▸ right_unique_forall₂' h₀ h₁ ▸ rfl
108108

109-
lemma right_unique_forall (hr : right_unique r) : right_unique (forall₂ r) :=
110-
@right_unique_forall₂' _ _ _ hr
109+
lemma _root_.relator.right_unique.forall (hr : right_unique r) : right_unique (forall₂ r) :=
110+
@right_unique_forall₂' _ _ _ hr
111111

112-
lemma bi_unique_forall (hr : bi_unique r) : bi_unique (forall₂ r) :=
113-
@@bi_unique.mk _ (left_unique_forall₂ hr.1) (right_unique_forall₂ hr.2)
112+
lemma _root_.relator.bi_unique.forall (hr : bi_unique r) : bi_unique (forall₂ r) :=
113+
hr.left.forall₂, hr.right.forall₂⟩
114114

115115
theorem forall₂_length_eq {R : α → β → Prop} :
116116
∀ {l₁ l₂}, forall₂ R l₁ l₂ → length l₁ = length l₂

src/data/list/perm.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -296,7 +296,7 @@ this ▸ hbd
296296
lemma rel_perm (hr : bi_unique r) : (forall₂ r ⇒ forall₂ r ⇒ (↔)) perm perm :=
297297
assume a b hab c d hcd, iff.intro
298298
(rel_perm_imp hr.2 hab hcd)
299-
(rel_perm_imp (left_unique_flip hr.1) hab.flip hcd.flip)
299+
(rel_perm_imp hr.left.flip hab.flip hcd.flip)
300300

301301
end rel
302302

src/data/option/basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -71,7 +71,7 @@ theorem mem_unique {o : option α} {a b : α} (ha : a ∈ o) (hb : b ∈ o) : a
7171
option.some.inj $ ha.symm.trans hb
7272

7373
theorem mem.left_unique : relator.left_unique ((∈) : α → option α → Prop) :=
74-
λ a o b, mem_unique
74+
λ a o b, mem_unique
7575

7676
theorem some_injective (α : Type*) : function.injective (@some α) :=
7777
λ _ _, some_inj.mp

src/data/part.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -69,7 +69,7 @@ theorem mem_unique : ∀ {a b : α} {o : part α}, a ∈ o → b ∈ o → a = b
6969
| _ _ ⟨p, f⟩ ⟨h₁, rfl⟩ ⟨h₂, rfl⟩ := rfl
7070

7171
theorem mem.left_unique : relator.left_unique ((∈) : α → part α → Prop) :=
72-
λ a o b, mem_unique
72+
λ a o b, mem_unique
7373

7474
theorem get_eq_of_mem {o : part α} {a} (h : a ∈ o) (h') : get o h' = a :=
7575
mem_unique ⟨_, rfl⟩ h

src/data/seq/computation.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -240,7 +240,7 @@ theorem mem_unique {s : computation α} {a b : α} : a ∈ s → b ∈ s → a =
240240
(le_stable s (le_max_right m n) hb.symm)
241241

242242
theorem mem.left_unique : relator.left_unique ((∈) : α → computation α → Prop) :=
243-
λ a s b, mem_unique
243+
λ a s b, mem_unique
244244

245245
/-- `terminates s` asserts that the computation `s` eventually terminates with some value. -/
246246
class terminates (s : computation α) : Prop := (term : ∃ a, a ∈ s)

src/logic/relation.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -257,7 +257,7 @@ begin
257257
{ rcases IH with IH | IH,
258258
{ rcases cases_head IH with rfl | ⟨e, be, ec⟩,
259259
{ exact or.inr (single bd) },
260-
{ cases U.unique bd be, exact or.inl ec } },
260+
{ cases U bd be, exact or.inl ec } },
261261
{ exact or.inr (IH.tail bd) } }
262262
end
263263

src/logic/relator.lean

Lines changed: 25 additions & 39 deletions
Original file line numberDiff line numberDiff line change
@@ -6,7 +6,7 @@ Authors: Johannes Hölzl
66
Relator for functions, pairs, sums, and lists.
77
-/
88

9-
import tactic.reserved_notation
9+
import logic.basic
1010

1111
namespace relator
1212
universe variables u₁ u₂ v₁ v₂
@@ -30,50 +30,40 @@ infixr ⇒ := lift_fun
3030

3131
end
3232

33-
section
34-
variables {α : Type u₁} {β : out_param $ Type u₂} (R : out_param $ α → β → Prop)
35-
36-
class right_total : Prop := (right : ∀b, ∃a, R a b)
37-
class left_total : Prop := (left : ∀a, ∃b, R a b)
38-
class bi_total extends left_total R, right_total R : Prop
39-
40-
end
41-
4233
section
4334
variables {α : Type u₁} {β : Type u₂} (R : α → β → Prop)
4435

45-
class left_unique : Prop := (left : ∀{a b c}, R a b → R c b → a = c)
46-
class right_unique : Prop := (right : ∀{a b c}, R a b → R a c → b = c)
36+
def right_total : Prop := ∀ b, ∃ a, R a b
37+
def left_total : Prop := ∀ a, ∃ b, R a b
38+
def bi_total : Prop := left_total R ∧ right_total R
4739

48-
lemma left_unique.unique {R : α → β → Prop} (h : left_unique R) {a b c} :
49-
R a b → R c b → a = c := left_unique.left β
40+
def left_unique : Prop := ∀ ⦃a b c⦄, R a c → R b c → a = b
41+
def right_unique : Prop := ∀ ⦃a b c⦄, R a b → R a c → b = c
42+
def bi_unique : Prop := left_unique R ∧ right_unique R
5043

51-
lemma right_unique.unique {R : α → β → Prop} (h : right_unique R) {a b c} :
52-
R a b → R a c → b = c := right_unique.right α
44+
variable {R}
5345

54-
lemma rel_forall_of_right_total [right_total R] :
46+
lemma right_total.rel_forall (h : right_total R) :
5547
((R ⇒ implies) ⇒ implies) (λp, ∀i, p i) (λq, ∀i, q i) :=
56-
assume p q Hrel H b, exists.elim (right_total.right b) (assume a Rab, Hrel Rab (H _))
48+
assume p q Hrel H b, exists.elim (h b) (assume a Rab, Hrel Rab (H _))
5749

58-
lemma rel_exists_of_left_total [left_total R] :
50+
lemma left_total.rel_exists (h : left_total R) :
5951
((R ⇒ implies) ⇒ implies) (λp, ∃i, p i) (λq, ∃i, q i) :=
60-
assume p q Hrel ⟨a, pa⟩, let ⟨b, Rab⟩ := left_total.left a in ⟨b, Hrel Rab pa
52+
assume p q Hrel ⟨a, pa⟩, (h a).imp $ λ b Rab, Hrel Rab pa
6153

62-
lemma rel_forall_of_total [bi_total R] : ((R ⇒ iff) ⇒ iff) (λp, ∀i, p i) (λq, ∀i, q i) :=
54+
lemma bi_total.rel_forall (h : bi_total R) :
55+
((R ⇒ iff) ⇒ iff) (λp, ∀i, p i) (λq, ∀i, q i) :=
6356
assume p q Hrel,
64-
⟨assume H b, exists.elim (right_total.right b) (assume a Rab, (Hrel Rab).mp (H _)),
65-
assume H a, exists.elim (left_total.left a) (assume b Rab, (Hrel Rab).mpr (H _))⟩
57+
⟨assume H b, exists.elim (h.right b) (assume a Rab, (Hrel Rab).mp (H _)),
58+
assume H a, exists.elim (h.left a) (assume b Rab, (Hrel Rab).mpr (H _))⟩
6659

67-
lemma rel_exists_of_total [bi_total R] : ((R ⇒ iff) ⇒ iff) (λp, ∃i, p i) (λq, ∃i, q i) :=
60+
lemma bi_total.rel_exists (h : bi_total R) : ((R ⇒ iff) ⇒ iff) (λp, ∃i, p i) (λq, ∃i, q i) :=
6861
assume p q Hrel,
69-
⟨assume ⟨a, pa⟩, let ⟨b, Rab⟩ := left_total.left a in ⟨b, (Hrel Rab).1 pa,
70-
assume ⟨b, qb⟩, let ⟨a, Rab⟩ := right_total.right b in ⟨a, (Hrel Rab).2 qb
62+
⟨assume ⟨a, pa⟩, (h.left a).imp $ λ b Rab, (Hrel Rab).1 pa,
63+
assume ⟨b, qb⟩, (h.right b).imp $ λ a Rab, (Hrel Rab).2 qb⟩
7164

7265
lemma left_unique_of_rel_eq {eq' : β → β → Prop} (he : (R ⇒ (R ⇒ iff)) eq eq') : left_unique R :=
73-
⟨λ a b c (ab : R a b) (cb : R c b),
74-
have eq' b b,
75-
from iff.mp (he ab ab) rfl,
76-
iff.mpr (he ab cb) this
66+
λ a b c (ac : R a c) (bc : R b c), (he ac bc).mpr ((he bc bc).mp rfl)
7767

7868
end
7969

@@ -83,18 +73,14 @@ assume p q h r s l, imp_congr h l
8373
lemma rel_not : (iff ⇒ iff) not not :=
8474
assume p q h, not_congr h
8575

86-
@[priority 100] -- see Note [lower instance priority]
87-
-- (this is an instance is always applies, since the relation is an out-param)
88-
instance bi_total_eq {α : Type u₁} : relator.bi_total (@eq α) :=
76+
lemma bi_total_eq {α : Type u₁} : relator.bi_total (@eq α) :=
8977
{ left := λ a, ⟨a, rfl⟩, right := λ a, ⟨a, rfl⟩ }
9078

9179
variables {α : Type*} {β : Type*} {γ : Type*} {δ : Type*}
9280
variables {r : α → β → Prop} {p : β → γ → Prop} {q : γ → δ → Prop}
9381

94-
class bi_unique (r : α → β → Prop) extends left_unique r, right_unique r : Prop
95-
96-
lemma left_unique_flip : left_unique r → right_unique (flip r)
97-
| ⟨h⟩ := ⟨λ a b c h₁ h₂, h h₁ h₂⟩
82+
lemma left_unique.flip (h : left_unique r) : right_unique (flip r) :=
83+
λ a b c h₁ h₂, h h₁ h₂
9884

9985
lemma rel_and : ((↔) ⇒ (↔) ⇒ (↔)) (∧) (∧) :=
10086
assume a b h₁ c d h₂, and_congr h₁ h₂
@@ -108,7 +94,7 @@ assume a b h₁ c d h₂, iff_congr h₁ h₂
10894
lemma rel_eq {r : α → β → Prop} (hr : bi_unique r) : (r ⇒ r ⇒ (↔)) (=) (=) :=
10995
assume a b h₁ c d h₂,
11096
iff.intro
111-
begin intro h, subst h, exact right_unique.right α h₁ h₂ end
112-
begin intro h, subst h, exact left_unique.left β h₁ h₂ end
97+
begin intro h, subst h, exact hr.right h₁ h₂ end
98+
begin intro h, subst h, exact hr.left h₁ h₂ end
11399

114100
end relator

0 commit comments

Comments
 (0)