diff --git a/src/data/fin.lean b/src/data/fin.lean index e69c1240ee239..b7c88f3341e53 100644 --- a/src/data/fin.lean +++ b/src/data/fin.lean @@ -22,6 +22,9 @@ This file expands on the development in the core library. `0`-th element `C (n+1) 0` of an `(n+1)`-tuple, and `Hs n i` defines `(i+1)`-st element of `(n+1)`-tuple based on `n`, `i`, and `i`-th element of `n`-tuple. * `fin.succ_rec_on` : same as `fin.succ_rec` but `i : fin n` is the first argument; +* `fin.induction` : Define `C i` by induction on `i : fin (n + 1)`, separating into the + `nat`-like base cases of `C 0` and `C (i.succ)`. +* `fin.induction_on` : same as `fin.induction` but with `i : fin (n + 1)` as the first argument. ### Casts @@ -666,13 +669,43 @@ rfl @fin.succ_rec_on (succ n) i.succ C H0 Hs = Hs n i (fin.succ_rec_on i H0 Hs) := by cases i; refl +/-- +Define `C i` by induction on `i : fin (n + 1)` via induction on the underlying `nat` value. +This function has two arguments: `h0` handles the base case on `C 0`, +and `hs` defines the inductive step using `C i.cast_succ`. +-/ +@[elab_as_eliminator] def induction + {C : fin (n + 1) → Sort*} + (h0 : C 0) + (hs : ∀ i : fin n, C i.cast_succ → C i.succ) : + Π (i : fin (n + 1)), C i := +begin + rintro ⟨i, hi⟩, + induction i with i IH, + { rwa [fin.mk_zero] }, + { refine hs ⟨i, lt_of_succ_lt_succ hi⟩ _, + exact IH (lt_of_succ_lt hi) } +end + +/-- +Define `C i` by induction on `i : fin (n + 1)` via induction on the underlying `nat` value. +This function has two arguments: `h0` handles the base case on `C 0`, +and `hs` defines the inductive step using `C i.cast_succ`. + +A version of `fin.induction` taking `i : fin (n + 1)` as the first argument. +-/ +@[elab_as_eliminator] def induction_on (i : fin (n + 1)) + {C : fin (n + 1) → Sort*} + (h0 : C 0) + (hs : ∀ i : fin n, C i.cast_succ → C i.succ) : C i := +induction h0 hs i + /-- Define `f : Π i : fin n.succ, C i` by separately handling the cases `i = 0` and `i = j.succ`, `j : fin n`. -/ @[elab_as_eliminator] def cases {C : fin (succ n) → Sort*} (H0 : C 0) (Hs : Π i : fin n, C (i.succ)) : - Π (i : fin (succ n)), C i -| ⟨0, h⟩ := H0 -| ⟨succ i, h⟩ := Hs ⟨i, lt_of_succ_lt_succ h⟩ + Π (i : fin (succ n)), C i := +induction H0 (λ i _, Hs i) @[simp] theorem cases_zero {n} {C : fin (succ n) → Sort*} {H0 Hs} : @fin.cases n C H0 Hs 0 = H0 := rfl @@ -681,6 +714,10 @@ rfl @fin.cases n C H0 Hs i.succ = Hs i := by cases i; refl +@[simp] theorem cases_succ' {n} {C : fin (succ n) → Sort*} {H0 Hs} {i : ℕ} (h : i + 1 < n + 1) : + @fin.cases n C H0 Hs ⟨i.succ, h⟩ = Hs ⟨i, lt_of_succ_lt_succ h⟩ := +by cases i; refl + lemma forall_fin_succ {P : fin (n+1) → Prop} : (∀ i, P i) ↔ P 0 ∧ (∀ i:fin n, P i.succ) := ⟨λ H, ⟨H 0, λ i, H _⟩, λ ⟨H0, H1⟩ i, fin.cases H0 H1 i⟩ diff --git a/src/data/matrix/notation.lean b/src/data/matrix/notation.lean index 3a780444bdeb2..dc200503141db 100644 --- a/src/data/matrix/notation.lean +++ b/src/data/matrix/notation.lean @@ -88,7 +88,7 @@ by simp [vec_cons] @[simp] lemma cons_val_succ' {i : ℕ} (h : i.succ < m.succ) (x : α) (u : fin m → α) : vec_cons x u ⟨i.succ, h⟩ = u ⟨i, nat.lt_of_succ_lt_succ h⟩ := -by simp [vec_cons, fin.cons, fin.cases] +by simp only [vec_cons, fin.cons, fin.cases_succ'] @[simp] lemma head_cons (x : α) (u : fin m → α) : vec_head (vec_cons x u) = x :=