refactor(library/init): define prod as an inductive datatype
Motivation: prod is used internally in the definitional package. If we define prod as a structure, then Lean will tag pr1 and pr2 as projections. This creates problems when we add special support for projections in the elaborator. The heuristics avoid some case-splits that are currently performed, and without them some files break.
This commit is contained in:
parent
d2e64d30e8
commit
1b414d36e7
4 changed files with 22 additions and 7 deletions
|
@ -35,8 +35,14 @@ refl : eq a a
|
||||||
inductive heq {A : Type} (a : A) : Π {B : Type}, B → Prop :=
|
inductive heq {A : Type} (a : A) : Π {B : Type}, B → Prop :=
|
||||||
refl : heq a a
|
refl : heq a a
|
||||||
|
|
||||||
structure prod (A B : Type) :=
|
inductive prod (A B : Type) :=
|
||||||
mk :: (pr1 : A) (pr2 : B)
|
mk : A → B → prod A B
|
||||||
|
|
||||||
|
definition prod.pr1 [reducible] [unfold-c 3] {A B : Type} (p : prod A B) : A :=
|
||||||
|
prod.rec (λ a b, a) p
|
||||||
|
|
||||||
|
definition prod.pr2 [reducible] [unfold-c 3] {A B : Type} (p : prod A B) : B :=
|
||||||
|
prod.rec (λ a b, b) p
|
||||||
|
|
||||||
inductive and (a b : Prop) : Prop :=
|
inductive and (a b : Prop) : Prop :=
|
||||||
intro : a → b → and a b
|
intro : a → b → and a b
|
||||||
|
|
|
@ -27,6 +27,16 @@ namespace prod
|
||||||
postfix `.2`:(max+1) := pr2
|
postfix `.2`:(max+1) := pr2
|
||||||
end ops
|
end ops
|
||||||
|
|
||||||
|
definition destruct [reducible] := @prod.cases_on
|
||||||
|
|
||||||
|
section
|
||||||
|
variables {A B : Type}
|
||||||
|
lemma pr1.mk (a : A) (b : B) : pr1 (mk a b) = a := rfl
|
||||||
|
lemma pr2.mk (a : A) (b : B) : pr2 (mk a b) = b := rfl
|
||||||
|
lemma eta : ∀ (p : A × B), mk (pr1 p) (pr2 p) = p
|
||||||
|
| (a, b) := rfl
|
||||||
|
end
|
||||||
|
|
||||||
open well_founded
|
open well_founded
|
||||||
|
|
||||||
section
|
section
|
||||||
|
|
|
@ -1,7 +1,6 @@
|
||||||
inductive pnat.pnat : Type₁
|
inductive pnat.pnat : Type₁
|
||||||
constructors:
|
constructors:
|
||||||
pnat.pnat.pos : Π (n : ℕ), nat.gt n (nat.of_num 0) → ℕ+
|
pnat.pnat.pos : Π (n : ℕ), nat.gt n (nat.of_num 0) → ℕ+
|
||||||
structure prod : Type → Type → Type
|
inductive prod : Type → Type → Type
|
||||||
fields:
|
constructors:
|
||||||
prod.pr1 : Π {A : Type} {B : Type}, A × B → A
|
prod.mk : Π {A : Type} {B : Type}, A → B → A × B
|
||||||
prod.pr2 : Π {A : Type} {B : Type}, A × B → B
|
|
||||||
|
|
|
@ -2,7 +2,7 @@ open nat prod
|
||||||
|
|
||||||
set_option pp.coercions true
|
set_option pp.coercions true
|
||||||
|
|
||||||
definition s : nat × nat := {| prod, pr1 := 10, pr2 := 20 |}
|
definition s : nat × nat := pair 10 20
|
||||||
|
|
||||||
structure test :=
|
structure test :=
|
||||||
(A : Type) (a : A) (B : A → Type) (b : B a)
|
(A : Type) (a : A) (B : A → Type) (b : B a)
|
||||||
|
|
Loading…
Reference in a new issue