Skip to content

Commit

Permalink
feat: apply pp_using_anonymous_constructor attribute (#3735)
Browse files Browse the repository at this point in the history
This attribute, which was implemented in #3640, is applied to the
following structures: `Sigma`, `PSigma`, `PProd`, `And`, `Subtype`, and
`Fin`. These were given this attribute in Lean 3.
  • Loading branch information
kmill authored Mar 22, 2024
1 parent d884a94 commit acb188f
Show file tree
Hide file tree
Showing 14 changed files with 39 additions and 28 deletions.
1 change: 1 addition & 0 deletions RELEASES.md
Original file line number Diff line number Diff line change
Expand Up @@ -55,6 +55,7 @@ v4.8.0 (development in progress)

* Attribute `@[pp_using_anonymous_constructor]` to make structures pretty print like `⟨x, y, z⟩`
rather than `{a := x, b := y, c := z}`.
This attribute is applied to `Sigma`, `PSigma`, `PProd`, `Subtype`, `And`, and `Fin`.

Breaking changes:

Expand Down
2 changes: 2 additions & 0 deletions src/Init/Core.lean
Original file line number Diff line number Diff line change
Expand Up @@ -165,6 +165,7 @@ whose first component is `a : α` and whose second component is `b : β a`
It is sometimes known as the dependent sum type, since it is the type level version
of an indexed summation.
-/
@[pp_using_anonymous_constructor]
structure Sigma {α : Type u} (β : α → Type v) where
/-- Constructor for a dependent pair. If `a : α` and `b : β a` then `⟨a, b⟩ : Sigma β`.
(This will usually require a type ascription to determine `β`
Expand All @@ -190,6 +191,7 @@ which can cause problems for universe level unification,
because the equation `max 1 u v = ?u + 1` has no solution in level arithmetic.
`PSigma` is usually only used in automation that constructs pairs of arbitrary types.
-/
@[pp_using_anonymous_constructor]
structure PSigma {α : Sort u} (β : α → Sort v) where
/-- Constructor for a dependent pair. If `a : α` and `b : β a` then `⟨a, b⟩ : PSigma β`.
(This will usually require a type ascription to determine `β`
Expand Down
4 changes: 4 additions & 0 deletions src/Init/Prelude.lean
Original file line number Diff line number Diff line change
Expand Up @@ -488,6 +488,7 @@ attribute [unbox] Prod
Similar to `Prod`, but `α` and `β` can be propositions.
We use this type internally to automatically generate the `brecOn` recursor.
-/
@[pp_using_anonymous_constructor]
structure PProd (α : Sort u) (β : Sort v) where
/-- The first projection out of a pair. if `p : PProd α β` then `p.1 : α`. -/
fst : α
Expand All @@ -509,6 +510,7 @@ structure MProd (α β : Type u) where
constructed and destructed like a pair: if `ha : a` and `hb : b` then
`⟨ha, hb⟩ : a ∧ b`, and if `h : a ∧ b` then `h.left : a` and `h.right : b`.
-/
@[pp_using_anonymous_constructor]
structure And (a b : Prop) : Prop where
/-- `And.intro : a → b → a ∧ b` is the constructor for the And operation. -/
intro ::
Expand Down Expand Up @@ -575,6 +577,7 @@ a pair-like type, so if you have `x : α` and `h : p x` then
`⟨x, h⟩ : {x // p x}`. An element `s : {x // p x}` will coerce to `α` but
you can also make it explicit using `s.1` or `s.val`.
-/
@[pp_using_anonymous_constructor]
structure Subtype {α : Sort u} (p : α → Prop) where
/-- If `s : {x // p x}` then `s.val : α` is the underlying element in the base
type. You can also write this as `s.1`, or simply as `s` when the type is
Expand Down Expand Up @@ -1815,6 +1818,7 @@ theorem System.Platform.numBits_eq : Or (Eq numBits 32) (Eq numBits 64) :=
`Fin n` is a natural number `i` with the constraint that `0 ≤ i < n`.
It is the "canonical type with `n` elements".
-/
@[pp_using_anonymous_constructor]
structure Fin (n : Nat) where
/-- If `i : Fin n`, then `i.val : ℕ` is the described number. It can also be
written as `i.1` or just `i` when the target type is known. -/
Expand Down
4 changes: 2 additions & 2 deletions tests/lean/1081.lean.expected.out
Original file line number Diff line number Diff line change
Expand Up @@ -7,6 +7,6 @@ but is expected to have type
1081.lean:23:4-23:7: error: type mismatch
rfl
has type
insert a { val := 0, isLt := ⋯ } v = insert a { val := 0, isLt := ⋯ } v : Prop
insert a 0, ⋯⟩ v = insert a 0, ⋯⟩ v : Prop
but is expected to have type
insert a { val := 0, isLt := ⋯ } v = cons a v : Prop
insert a 0, ⋯⟩ v = cons a v : Prop
4 changes: 2 additions & 2 deletions tests/lean/243.lean.expected.out
Original file line number Diff line number Diff line change
@@ -1,13 +1,13 @@
243.lean:2:10-2:14: error: application type mismatch
{ fst := Bool, snd := true }
Bool, true
argument
true
has type
_root_.Bool : Type
but is expected to have type
Bool : Type
243.lean:13:7-13:8: error: application type mismatch
{ fst := A, snd := a }
A, a⟩
argument
a
has type
Expand Down
4 changes: 2 additions & 2 deletions tests/lean/congrThmIssue.lean.expected.out
Original file line number Diff line number Diff line change
Expand Up @@ -4,7 +4,7 @@ cap : Nat
backing : Fin cap → Option α
size : Nat
h_size : size ≤ cap
h_full : ∀ (i : Nat) (h : i < size), Option.isSome (backing { val := i, isLt := ⋯ }) = true
h_full : ∀ (i : Nat) (h : i < size), Option.isSome (backing i, ⋯⟩) = true
i : Nat
h : i < size
⊢ Option.isSome (if h_1 : i < cap then backing { val := i, isLt := ⋯ } else none) = true
⊢ Option.isSome (if h_1 : i < cap then backing i, ⋯⟩ else none) = true
15 changes: 5 additions & 10 deletions tests/lean/decreasing_by.lean.expected.out
Original file line number Diff line number Diff line change
Expand Up @@ -15,21 +15,17 @@ Please use `termination_by` to specify a decreasing measure.
decreasing_by.lean:81:13-83:3: error: unexpected token 'end'; expected '{' or tactic
decreasing_by.lean:81:0-81:13: error: unsolved goals
n m : Nat
⊢ (invImage (fun x => PSigma.casesOn x fun n m => (n, m)) Prod.instWellFoundedRelationProd).1
{ fst := n, snd := dec2 m } { fst := n, snd := m }
⊢ (invImage (fun x => PSigma.casesOn x fun n m => (n, m)) Prod.instWellFoundedRelationProd).1 ⟨n, dec2 m⟩ ⟨n, m⟩

n m : Nat
⊢ (invImage (fun x => PSigma.casesOn x fun n m => (n, m)) Prod.instWellFoundedRelationProd).1
{ fst := dec1 n, snd := 100 } { fst := n, snd := m }
⊢ (invImage (fun x => PSigma.casesOn x fun n m => (n, m)) Prod.instWellFoundedRelationProd).1 ⟨dec1 n, 100⟩ ⟨n, m⟩
decreasing_by.lean:91:0-91:22: error: unsolved goals
case a
n m : Nat
⊢ (invImage (fun x => PSigma.casesOn x fun n m => (n, m)) Prod.instWellFoundedRelationProd).1
{ fst := n, snd := dec2 m } { fst := n, snd := m }
⊢ (invImage (fun x => PSigma.casesOn x fun n m => (n, m)) Prod.instWellFoundedRelationProd).1 ⟨n, dec2 m⟩ ⟨n, m⟩

n m : Nat
⊢ (invImage (fun x => PSigma.casesOn x fun n m => (n, m)) Prod.instWellFoundedRelationProd).1
{ fst := dec1 n, snd := 100 } { fst := n, snd := m }
⊢ (invImage (fun x => PSigma.casesOn x fun n m => (n, m)) Prod.instWellFoundedRelationProd).1 ⟨dec1 n, 100⟩ ⟨n, m⟩
decreasing_by.lean:99:0-100:22: error: Could not find a decreasing measure.
The arguments relate at each recursive call as follows:
(<, ≤, =: relation proved, ? all proofs failed, _: no proof attempted)
Expand All @@ -39,8 +35,7 @@ The arguments relate at each recursive call as follows:
Please use `termination_by` to specify a decreasing measure.
decreasing_by.lean:110:0-113:17: error: unsolved goals
n m : Nat
⊢ (invImage (fun x => PSigma.casesOn x fun n m => (n, m)) Prod.instWellFoundedRelationProd).1
{ fst := dec1 n, snd := 100 } { fst := n, snd := m }
⊢ (invImage (fun x => PSigma.casesOn x fun n m => (n, m)) Prod.instWellFoundedRelationProd).1 ⟨dec1 n, 100⟩ ⟨n, m⟩
decreasing_by.lean:121:0-125:17: error: Could not find a decreasing measure.
The arguments relate at each recursive call as follows:
(<, ≤, =: relation proved, ? all proofs failed, _: no proof attempted)
Expand Down
4 changes: 1 addition & 3 deletions tests/lean/issue2981.lean.expected.out
Original file line number Diff line number Diff line change
Expand Up @@ -8,6 +8,4 @@ Tactic is run (ideally only twice, in most general context)
Tactic is run (ideally only twice, in most general context)
n : Nat
⊢ sizeOf n < sizeOf (Nat.succ n)
n m : Nat
⊢ (invImage (fun x => PSigma.casesOn x fun a a_1 => a) instWellFoundedRelation).1 { fst := n, snd := m + 1 }
{ fst := Nat.succ n, snd := m }
n m : Nat ⊢ (invImage (fun x => PSigma.casesOn x fun a a_1 => a) instWellFoundedRelation).1 ⟨n, m + 1⟩ ⟨Nat.succ n, m⟩
2 changes: 1 addition & 1 deletion tests/lean/letFun.lean.expected.out
Original file line number Diff line number Diff line change
Expand Up @@ -15,7 +15,7 @@ a b : Nat
h : a > b
⊢ b < a
let_fun n := 5;
{ val := [], property := ⋯ } : { as // List.length as ≤ 5 }
[], ⋯⟩ : { as // List.length as ≤ 5 }
rfl : (let_fun n := 5;
n) =
let_fun n := 5;
Expand Down
3 changes: 1 addition & 2 deletions tests/lean/run/986.lean
Original file line number Diff line number Diff line change
Expand Up @@ -12,8 +12,7 @@ info: Array.insertionSort.swapLoop.eq_2.{u_1} {α : Type u_1} (lt : α → α
(h : Nat.succ j' < Array.size a) :
Array.insertionSort.swapLoop lt a (Nat.succ j') h =
let_fun h' := ⋯;
if lt a[Nat.succ j'] a[j'] = true then
Array.insertionSort.swapLoop lt (Array.swap a { val := Nat.succ j', isLt := h } { val := j', isLt := h' }) j' ⋯
if lt a[Nat.succ j'] a[j'] = true then Array.insertionSort.swapLoop lt (Array.swap a ⟨Nat.succ j', h⟩ ⟨j', h'⟩) j' ⋯
else a
-/
#guard_msgs in
Expand Down
4 changes: 2 additions & 2 deletions tests/lean/run/PPTopDownAnalyze.lean
Original file line number Diff line number Diff line change
Expand Up @@ -196,7 +196,7 @@ set_option pp.analyze.trustSubst true in

set_option pp.analyze.trustId true in
#testDelab Sigma.mk (β := fun α => α) Bool true
expecting { fst := _, snd := true }
expecting _, true

set_option pp.analyze.trustId false in
#testDelab Sigma.mk (β := fun α => α) Bool true
Expand Down Expand Up @@ -330,7 +330,7 @@ set_option pp.analyze.explicitHoles false in

set_option pp.analyze.trustSubtypeMk true in
#testDelab fun (n : Nat) (val : List Nat) (property : List.length val = n) => List.length { val := val, property := property : { x : List Nat // List.length x = n } }.val = n
expecting fun n val property => List.length { val := val, property := property : { x : List Nat // List.length x = n } }.val = n
expecting fun n val property => List.length (⟨val, property⟩ : { x : List Nat // List.length x = n }).val = n

#testDelabN Nat.brecOn
#testDelabN Nat.below
Expand Down
2 changes: 1 addition & 1 deletion tests/lean/run/funind_tests.lean
Original file line number Diff line number Diff line change
Expand Up @@ -721,7 +721,7 @@ info: Nary.foo.induct (motive : Nat → Nat → (k : Nat) → Fin k → Prop)
(case4 : ∀ (x x_1 : Nat) (x_2 : Fin 1), (x = 0 → False) → (x_1 = 0 → False) → motive x x_1 1 x_2)
(case5 :
∀ (n m k : Nat) (x : Fin (k + 2)),
motive n m (k + 1) { val := 0, isLt := ⋯ } → motive (Nat.succ n) (Nat.succ m) (Nat.succ (Nat.succ k)) x) :
motive n m (k + 1) 0, ⋯⟩ → motive (Nat.succ n) (Nat.succ m) (Nat.succ (Nat.succ k)) x) :
∀ (a a_1 k : Nat) (a_2 : Fin k), motive a a_1 k a_2
-/
#guard_msgs in
Expand Down
6 changes: 3 additions & 3 deletions tests/lean/run/litToCtor.lean
Original file line number Diff line number Diff line change
Expand Up @@ -42,17 +42,17 @@ info: Int.ofNat 3
#guard_msgs in
#eval test (3 : Int)
/--
info: { val := 3, isLt := ⋯ }
info: 3, ⋯⟩
-/
#guard_msgs in
#eval test (3 : Fin 5)
/--
info: { val := 0, isLt := ⋯ }
info: 0, ⋯⟩
-/
#guard_msgs in
#eval test (0 : Fin 5)
/--
info: { val := 1, isLt := ⋯ }
info: 1, ⋯⟩
-/
#guard_msgs in
#eval test (6 : Fin 5)
12 changes: 12 additions & 0 deletions tests/lean/run/ppUsingAnonymousConstructor.lean
Original file line number Diff line number Diff line change
Expand Up @@ -15,3 +15,15 @@ structure S where
attribute [pp_using_anonymous_constructor] S
/-- info: ⟨1, 2⟩ : S -/
#guard_msgs in #check {x := 1, y := 2 : S}

/-!
`Fin`
-/
/-- info: ⟨2, ⋯⟩ : Fin 3 -/
#guard_msgs in #check Fin.mk 2 (by omega : 2 < 3)

/-!
`Subtype`
-/
/-- info: ⟨2, ⋯⟩ : { n // n < 3 } -/
#guard_msgs in #check (⟨2, by omega⟩ : {n : Nat // n < 3})

0 comments on commit acb188f

Please sign in to comment.