Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat: super factorial #6768

Closed
wants to merge 13 commits into from
1 change: 1 addition & 0 deletions Mathlib.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1607,6 +1607,7 @@ import Mathlib.Data.Nat.Factorial.Basic
import Mathlib.Data.Nat.Factorial.BigOperators
import Mathlib.Data.Nat.Factorial.Cast
import Mathlib.Data.Nat.Factorial.DoubleFactorial
import Mathlib.Data.Nat.Factorial.SuperFactorial
import Mathlib.Data.Nat.Factorization.Basic
import Mathlib.Data.Nat.Factorization.PrimePow
import Mathlib.Data.Nat.Factors
Expand Down
52 changes: 52 additions & 0 deletions Mathlib/Data/Nat/Factorial/SuperFactorial.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,52 @@
/-
Copyright (c) 2023 Moritz Firsching. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Moritz Firsching
-/
import Mathlib.Data.Nat.Factorial.Basic

/-!
# Superfactorial

This file defines the [superfactorial](https://en.wikipedia.org/wiki/Superfactorial)
`1! * 2! * 3! * ...* n!`.

## Main declarations

* `Nat.superFactorial`: The superfactorial.
-/
mo271 marked this conversation as resolved.
Show resolved Hide resolved


namespace Nat

/-- `Nat.superFactorial n` is the superfactorial of `n`. -/
@[simp]
mo271 marked this conversation as resolved.
Show resolved Hide resolved
def superFactorial : ℕ → ℕ
| 0 => 1
| succ n => factorial n.succ * superFactorial n

/-- `sf` notation for superfactorial -/
scoped notation "sf" n:60 => Nat.superFactorial n

section SuperFactorial

variable {n : ℕ}

@[simp]
theorem superFactorial_zero : sf 0 = 1 :=
rfl

theorem superFactorial_succ (n : ℕ) : (sf n.succ) = (n + 1)! * sf n :=
rfl

@[simp]
theorem superFactorial_one : sf 1 = 1 :=

Check failure on line 43 in Mathlib/Data/Nat/Factorial/SuperFactorial.lean

View workflow job for this annotation

GitHub Actions / Build

Nat.superFactorial_one Left-hand side simplifies from
rfl

@[simp]
theorem superFactorial_two : sf 2 = 2 :=

Check failure on line 47 in Mathlib/Data/Nat/Factorial/SuperFactorial.lean

View workflow job for this annotation

GitHub Actions / Build

Nat.superFactorial_two Left-hand side simplifies from
rfl

end SuperFactorial

end Nat