Consistently format lean files.

finite-set-exercises
Joshua Potter 2023-04-10 15:30:29 -06:00
parent cd3c98ee95
commit a2b138233a
5 changed files with 70 additions and 67 deletions

View File

@ -41,7 +41,7 @@ theorem term_recursive_closed (seq : Arithmetic) (n : Nat)
| succ n ih =>
calc
termRecursive seq (Nat.succ n)
= seq.Δ + seq.termRecursive n := rfl
_ = seq.Δ + seq.termRecursive n := rfl
_ = seq.Δ + seq.termClosed n := by rw [ih]
_ = seq.Δ + (seq.a₀ + seq.Δ * n) := rfl
_ = seq.a₀ + seq.Δ * (↑n + 1) := by ring

View File

@ -40,7 +40,7 @@ theorem term_recursive_closed (seq : Geometric) (n : Nat)
| zero => unfold termClosed termRecursive; norm_num
| succ n ih => calc
seq.termRecursive (n + 1)
= seq.r * (seq.termRecursive n) := rfl
_ = seq.r * (seq.termRecursive n) := rfl
_ = seq.r * (seq.termClosed n) := by rw [ih]
_ = seq.r * (seq.a₀ * seq.r ^ n) := rfl
_ = seq.a₀ * seq.r ^ (n + 1) := by ring

View File

@ -1,13 +1,12 @@
import Mathlib.Tactic.Ring
/--
As described in [1], `n`-tuples are defined recursively as such:
`n`-tuples are defined recursively as such:
`⟨x₁, ..., xₙ⟩ = ⟨⟨x₁, ..., xₙ₋₁⟩, xₙ⟩`
We allow for empty tuples; [2] expects this functionality.
For a `Tuple`-like type with opposite "endian", refer to `Vector`.
We allow empty tuples. For a `Tuple`-like type with opposite "endian", refer to
`Mathlib.Data.Vector`.
-/
inductive Tuple : (α : Type u) → (size : Nat) → Type u where
| nil : Tuple α 0
@ -67,7 +66,8 @@ theorem eq_iff_snoc {t₁ t₂ : Tuple α n}
exact And.intro h₂ h₁
/--
Implements decidable equality for `Tuple α m`, provided `a` has decidable equality.
Implements decidable equality for `Tuple α m`, provided `a` has decidable
equality.
-/
protected def hasDecEq [DecidableEq α] (t₁ t₂ : Tuple α n)
: Decidable (Eq t₁ t₂) :=
@ -138,22 +138,22 @@ theorem nil_concat_self_eq_self (t : Tuple α m) : concat t[] t = t := by
induction t with
| nil => unfold concat; simp
| @snoc n as a ih =>
unfold concat
rw [ih]
suffices HEq (snoc (cast (_ : Tuple α n = Tuple α (0 + n)) as) a) ↑(snoc as a)
from eq_of_heq this
have h₁ := Eq.recOn
(motive := fun x h => HEq
(snoc (cast (show Tuple α n = Tuple α x by rw [h]) as) a)
(snoc as a))
(show n = 0 + n by simp)
HEq.rfl
exact Eq.recOn
(motive := fun x h => HEq
(snoc (cast (_ : Tuple α n = Tuple α (0 + n)) as) a)
(cast h (snoc as a)))
(show Tuple α (n + 1) = Tuple α (0 + (n + 1)) by simp)
h₁
unfold concat
rw [ih]
suffices HEq (snoc (cast (_ : Tuple α n = Tuple α (0 + n)) as) a) ↑(snoc as a)
from eq_of_heq this
have h₁ := Eq.recOn
(motive := fun x h => HEq
(snoc (cast (show Tuple α n = Tuple α x by rw [h]) as) a)
(snoc as a))
(show n = 0 + n by simp)
HEq.rfl
exact Eq.recOn
(motive := fun x h => HEq
(snoc (cast (_ : Tuple α n = Tuple α (0 + n)) as) a)
(cast h (snoc as a)))
(show Tuple α (n + 1) = Tuple α (0 + (n + 1)) by simp)
h₁
/--
Concatenating a `Tuple` to a nonempty `Tuple` moves `concat` calls closer to
@ -227,21 +227,22 @@ theorem take_subst_last {as : Tuple α n} (a₁ a₂ : α)
/--
Taking `n` elements from a tuple of size `n + 1` is the same as invoking `init`.
-/
theorem init_eq_take_pred (t : Tuple α (n + 1)) : take t n = init t :=
match t with
| snoc as a => by
unfold init take
simp
rw [self_take_size_eq_self]
simp
theorem init_eq_take_pred (t : Tuple α (n + 1)) : take t n = init t := by
cases t with
| snoc as a =>
unfold init take
simp
rw [self_take_size_eq_self]
simp
/--
If two `Tuple`s are equal, then any initial sequences of those two `Tuple`s are
also equal.
-/
theorem eq_tuple_eq_take {t₁ t₂ : Tuple α n}
: (t₁ = t₂) → (t₁.take k = t₂.take k) :=
fun h => by rw [h]
: (t₁ = t₂) → (t₁.take k = t₂.take k) := by
intro h
rw [h]
/--
Given a `Tuple` of size `k`, concatenating an arbitrary `Tuple` and taking `k`
@ -250,13 +251,15 @@ elements yields the original `Tuple`.
theorem eq_take_concat {t₁ : Tuple α m} {t₂ : Tuple α n}
: take (concat t₁ t₂) m = t₁ := by
induction t₂ with
| nil => simp; rw [self_concat_nil_eq_self, self_take_size_eq_self]
| nil =>
simp
rw [self_concat_nil_eq_self, self_take_size_eq_self]
| @snoc n' as a ih =>
simp
rw [concat_snoc_snoc_concat]
unfold take
simp
rw [ih]
simp
simp
rw [concat_snoc_snoc_concat]
unfold take
simp
rw [ih]
simp
end Tuple

View File

@ -140,7 +140,10 @@ size `k` of the tuple in normal form.
-/
theorem self_fst_eq_norm_take (t : XTuple α (m, n)) : t.fst = t.norm.take m :=
match t with
| x[] => by unfold fst; rw [Tuple.self_take_zero_eq_nil]; simp
| x[] => by
unfold fst
rw [Tuple.self_take_zero_eq_nil]
simp
| snoc tf tl => by
unfold fst
conv => rhs; unfold norm
@ -152,8 +155,9 @@ If the normal form of an `XTuple` is equal to a `Tuple`, the `fst` component
must be a prefix of the `Tuple`.
-/
theorem norm_eq_fst_eq_take {t₁ : XTuple α (m, n)} {t₂ : Tuple α (m + n)}
: (t₁.norm = t₂) → (t₁.fst = t₂.take m) :=
fun h => by rw [self_fst_eq_norm_take, h]
: (t₁.norm = t₂) → (t₁.fst = t₂.take m) := by
intro h
rw [self_fst_eq_norm_take, h]
/--
Returns the first component of our `XTuple`. For example, the first component of
@ -173,9 +177,7 @@ variable {k m n : Nat}
variable (p : 1 ≤ m)
variable (q : n + (m - 1) = m + k)
namespace Lemma_0a
lemma n_eq_succ_k : n = k + 1 :=
private lemma n_eq_succ_k : n = k + 1 := by
let ⟨m', h⟩ := Nat.exists_eq_succ_of_ne_zero $ show m ≠ 0 by
intro h
have ff : 1 ≤ 0 := h ▸ p
@ -189,48 +191,46 @@ lemma n_eq_succ_k : n = k + 1 :=
_ = 1 + k := by simp
_ = k + 1 := by rw [Nat.add_comm]
lemma n_pred_eq_k : n - 1 = k := by
private lemma n_pred_eq_k : n - 1 = k := by
have h : k + 1 - 1 = k + 1 - 1 := rfl
conv at h => lhs; rw [←n_eq_succ_k p q]
simp at h
exact h
lemma n_geq_one : 1 ≤ n := by
private lemma n_geq_one : 1 ≤ n := by
rw [n_eq_succ_k p q]
simp
lemma min_comm_succ_eq : min (m + k) (k + 1) = k + 1 :=
private lemma min_comm_succ_eq : min (m + k) (k + 1) = k + 1 :=
Nat.recOn k
(by simp; exact p)
(fun k' ih => calc
min (m + (k' + 1)) (k' + 1 + 1)
= min (m + k' + 1) (k' + 1 + 1) := by conv => rw [Nat.add_assoc]
_ = min (m + k') (k' + 1) + 1 := Nat.min_succ_succ (m + k') (k' + 1)
_ = k' + 1 + 1 := by rw [ih])
(fun k' ih => calc min (m + (k' + 1)) (k' + 1 + 1)
_ = min (m + k' + 1) (k' + 1 + 1) := by conv => rw [Nat.add_assoc]
_ = min (m + k') (k' + 1) + 1 := Nat.min_succ_succ (m + k') (k' + 1)
_ = k' + 1 + 1 := by rw [ih])
lemma n_eq_min_comm_succ : n = min (m + k) (k + 1) := by
private lemma n_eq_min_comm_succ : n = min (m + k) (k + 1) := by
rw [min_comm_succ_eq p]
exact n_eq_succ_k p q
lemma n_pred_m_eq_m_k : n + (m - 1) = m + k := by
private lemma n_pred_m_eq_m_k : n + (m - 1) = m + k := by
rw [←Nat.add_sub_assoc p, Nat.add_comm, Nat.add_sub_assoc (n_geq_one p q)]
conv => lhs; rw [n_pred_eq_k p q]
def cast_norm : XTuple α (n, m - 1) → Tuple α (m + k)
private def cast_norm : XTuple α (n, m - 1) → Tuple α (m + k)
| xs => cast (by rw [q]) xs.norm
def cast_fst : XTuple α (n, m - 1) → Tuple α (k + 1)
private def cast_fst : XTuple α (n, m - 1) → Tuple α (k + 1)
| xs => cast (by rw [n_eq_succ_k p q]) xs.fst
def cast_take (ys : Tuple α (m + k)) :=
private def cast_take (ys : Tuple α (m + k)) :=
cast (by rw [min_comm_succ_eq p]) (ys.take (k + 1))
end Lemma_0a
/--
Lemma 0A
open Lemma_0a
/--[1]
Assume that ⟨x₁, ..., xₘ⟩ = ⟨y₁, ..., yₘ, ..., yₘ₊ₖ⟩. Then x₁ = ⟨y₁, ..., yₖ₊₁⟩.
Assume that `⟨x₁, ..., xₘ⟩ = ⟨y₁, ..., yₘ, ..., yₘ₊ₖ⟩`.
Then `x₁ = ⟨y₁, ..., yₖ₊₁⟩`.
-/
theorem lemma_0a (xs : XTuple α (n, m - 1)) (ys : Tuple α (m + k))
: (cast_norm q xs = ys) → (cast_fst p q xs = cast_take p ys) := by

View File

@ -88,10 +88,10 @@ only if `-x` is the greatest lower bound of `S`.
theorem is_lub_neg_set_iff_is_glb_set_neg (S : Set )
: IsLUB (-S) x = IsGLB S (-x) :=
calc IsLUB (-S) x
_ = IsLeast (upperBounds (-S)) x := rfl
_ = IsLeast (-lowerBounds S) x := by rw [upper_bounds_neg_eq_neg_lower_bounds S]
_ = IsGreatest (lowerBounds S) (-x) := by rw [is_least_neg_set_eq_is_greatest_set_neq]
_ = IsGLB S (-x) := rfl
_ = IsLeast (upperBounds (-S)) x := rfl
_ = IsLeast (-lowerBounds S) x := by rw [upper_bounds_neg_eq_neg_lower_bounds S]
_ = IsGreatest (lowerBounds S) (-x) := by rw [is_least_neg_set_eq_is_greatest_set_neq]
_ = IsGLB S (-x) := rfl
/--
Theorem I.27