Skip to content

Commit

Permalink
chore(combinatorics/configuration): Generalize universes (#18039)
Browse files Browse the repository at this point in the history
This PR generalizes the universes, allowing the `P` (points) and `L` (lines) to have different universe levels.

I also switched from forall to Pi for data fields, as suggested by @alreadydone.
  • Loading branch information
tb65536 committed Jan 2, 2023
1 parent c3442db commit 9830a30
Showing 1 changed file with 8 additions and 10 deletions.
18 changes: 8 additions & 10 deletions src/combinatorics/configuration.lean
Expand Up @@ -36,9 +36,7 @@ open_locale big_operators

namespace configuration

universe u

variables (P L : Type u) [has_mem P L]
variables (P L : Type*) [has_mem P L]

/-- A type synonym. -/
def dual := P
Expand All @@ -63,13 +61,13 @@ class nondegenerate : Prop :=
(eq_or_eq : ∀ {p₁ p₂ : P} {l₁ l₂ : L}, p₁ ∈ l₁ → p₂ ∈ l₁ → p₁ ∈ l₂ → p₂ ∈ l₂ → p₁ = p₂ ∨ l₁ = l₂)

/-- A nondegenerate configuration in which every pair of lines has an intersection point. -/
class has_points extends nondegenerate P L : Type u :=
(mk_point : {l₁ l₂ : L} (h : l₁ ≠ l₂), P)
class has_points extends nondegenerate P L :=
(mk_point : Π {l₁ l₂ : L} (h : l₁ ≠ l₂), P)
(mk_point_ax : ∀ {l₁ l₂ : L} (h : l₁ ≠ l₂), mk_point h ∈ l₁ ∧ mk_point h ∈ l₂)

/-- A nondegenerate configuration in which every pair of points has a line through them. -/
class has_lines extends nondegenerate P L : Type u :=
(mk_line : {p₁ p₂ : P} (h : p₁ ≠ p₂), L)
class has_lines extends nondegenerate P L :=
(mk_line : Π {p₁ p₂ : P} (h : p₁ ≠ p₂), L)
(mk_line_ax : ∀ {p₁ p₂ : P} (h : p₁ ≠ p₂), p₁ ∈ mk_line h ∧ p₂ ∈ mk_line h)

open nondegenerate has_points has_lines
Expand Down Expand Up @@ -309,10 +307,10 @@ variables (P L)
/-- A projective plane is a nondegenerate configuration in which every pair of lines has
an intersection point, every pair of points has a line through them,
and which has three points in general position. -/
class projective_plane extends nondegenerate P L : Type u :=
(mk_point : {l₁ l₂ : L} (h : l₁ ≠ l₂), P)
class projective_plane extends nondegenerate P L :=
(mk_point : Π {l₁ l₂ : L} (h : l₁ ≠ l₂), P)
(mk_point_ax : ∀ {l₁ l₂ : L} (h : l₁ ≠ l₂), mk_point h ∈ l₁ ∧ mk_point h ∈ l₂)
(mk_line : {p₁ p₂ : P} (h : p₁ ≠ p₂), L)
(mk_line : Π {p₁ p₂ : P} (h : p₁ ≠ p₂), L)
(mk_line_ax : ∀ {p₁ p₂ : P} (h : p₁ ≠ p₂), p₁ ∈ mk_line h ∧ p₂ ∈ mk_line h)
(exists_config : ∃ (p₁ p₂ p₃ : P) (l₁ l₂ l₃ : L), p₁ ∉ l₂ ∧ p₁ ∉ l₃ ∧
p₂ ∉ l₁ ∧ p₂ ∈ l₂ ∧ p₂ ∈ l₃ ∧ p₃ ∉ l₁ ∧ p₃ ∈ l₂ ∧ p₃ ∉ l₃)
Expand Down

0 comments on commit 9830a30

Please sign in to comment.