From 17987b09743026f71b0508488d1904f49109d8a0 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 11 Jun 2024 06:52:19 +0100 Subject: [PATCH 001/177] chore: squeeze simps in HashMap.WF (#833) --- Batteries/Data/HashMap/WF.lean | 85 +++++++++++++++++++++------------- 1 file changed, 53 insertions(+), 32 deletions(-) diff --git a/Batteries/Data/HashMap/WF.lean b/Batteries/Data/HashMap/WF.lean index d535685691..fb2635b8fe 100644 --- a/Batteries/Data/HashMap/WF.lean +++ b/Batteries/Data/HashMap/WF.lean @@ -23,22 +23,25 @@ theorem update_data (self : Buckets α β) (i d h) : theorem exists_of_update (self : Buckets α β) (i d h) : ∃ l₁ l₂, self.1.data = l₁ ++ self.1[i] :: l₂ ∧ List.length l₁ = i.toNat ∧ (self.update i d h).1.data = l₁ ++ d :: l₂ := by - simp [Array.getElem_eq_data_get]; exact List.exists_of_set' h + simp only [Array.data_length, Array.ugetElem_eq_getElem, Array.getElem_eq_data_get] + exact List.exists_of_set' h theorem update_update (self : Buckets α β) (i d d' h h') : (self.update i d h).update i d' h' = self.update i d' h := by - simp [update]; congr 1; rw [Array.set_set] + simp only [update, Array.uset, Array.data_length] + congr 1 + rw [Array.set_set] theorem size_eq (data : Buckets α β) : size data = .sum (data.1.data.map (·.toList.length)) := rfl theorem mk_size (h) : (mk n h : Buckets α β).size = 0 := by - simp [Buckets.size_eq, Buckets.mk, mkArray]; clear h + simp only [mk, mkArray, size_eq]; clear h induction n <;> simp [*] theorem WF.mk' [BEq α] [Hashable α] (h) : (Buckets.mk n h : Buckets α β).WF := by refine ⟨fun _ h => ?_, fun i h => ?_⟩ - · simp [Buckets.mk, empty', mkArray, List.mem_replicate] at h + · simp only [Buckets.mk, mkArray, List.mem_replicate, ne_eq] at h simp [h, List.Pairwise.nil] · simp [Buckets.mk, empty', mkArray, Array.getElem_eq_data_get, AssocList.All] @@ -53,16 +56,19 @@ theorem WF.update [BEq α] [Hashable α] {buckets : Buckets α β} {i d h} (H : · exact match List.mem_or_eq_of_mem_set hl with | .inl hl => H.1 _ hl | .inr rfl => h₁ (H.1 _ (Array.getElem_mem_data ..)) - · revert hp; simp [update_data, Array.getElem_eq_data_get, List.get_set] + · revert hp + simp only [Array.getElem_eq_data_get, update_data, List.get_set, Array.data_length, update_size] split <;> intro hp · next eq => exact eq ▸ h₂ (H.2 _ _) _ hp - · simp at hi; exact H.2 i hi _ hp + · simp only [update_size, Array.data_length] at hi + exact H.2 i hi _ hp end Buckets theorem reinsertAux_size [Hashable α] (data : Buckets α β) (a : α) (b : β) : (reinsertAux data a b).size = data.size.succ := by - simp [Buckets.size_eq, reinsertAux] + simp only [reinsertAux, Array.data_length, Array.ugetElem_eq_getElem, Buckets.size_eq, + Nat.succ_eq_add_one] refine have ⟨l₁, l₂, h₁, _, eq⟩ := Buckets.exists_of_update ..; eq ▸ ?_ simp [h₁, Nat.succ_add]; rfl @@ -88,12 +94,13 @@ where · next H => refine (go (i+1) _ _ fun j hj => ?a).trans ?b <;> simp · case a => - simp [List.getD_eq_get?, List.get?_set]; split + simp only [List.getD_eq_get?, List.get?_set, Option.map_eq_map]; split · cases List.get? .. <;> rfl · next H => exact hs _ (Nat.lt_of_le_of_ne (Nat.le_of_lt_succ hj) (Ne.symm H)) · case b => refine have ⟨l₁, l₂, h₁, _, eq⟩ := List.exists_of_set' H; eq ▸ ?_ - simp [h₁, Buckets.size_eq] + simp only [Buckets.size_eq, h₁, List.map_append, List.map_cons, AssocList.toList, + List.length_nil, Nat.sum_append, Nat.sum_cons, Nat.zero_add, Array.data_length] rw [Nat.add_assoc, Nat.add_assoc, Nat.add_assoc]; congr 1 (conv => rhs; rw [Nat.add_left_comm]); congr 1 rw [← Array.getElem_eq_data_get] @@ -102,9 +109,10 @@ where · next H => rw [(_ : Nat.sum _ = 0), Nat.zero_add] rw [← (_ : source.data.map (fun _ => .nil) = source.data)] - · simp; induction source.data <;> simp [*] + · simp only [List.map_map] + induction source.data <;> simp [*] refine List.ext_get (by simp) fun j h₁ h₂ => ?_ - simp + simp only [List.get_map, Array.data_length] have := (hs j (Nat.lt_of_lt_of_le h₂ (Nat.not_lt.1 H))).symm rwa [List.getD_eq_get?, List.get?_eq_get, Option.getD_some] at this termination_by source.size - i @@ -126,7 +134,8 @@ theorem expand_WF.foldl [BEq α] [Hashable α] (rank : α → Nat) {l : List (α refine ih hl₁.2 hl₂.2 (reinsertAux_WF ht₁ fun _ h => (ht₂ _ (Array.getElem_mem_data ..) _ h).2.1) (fun _ h => ?_) - simp [reinsertAux, Buckets.update] at h + simp only [reinsertAux, Buckets.update, Array.uset, Array.data_length, + Array.ugetElem_eq_getElem, Array.data_set] at h match List.mem_or_eq_of_mem_set h with | .inl h => intro _ hf @@ -157,7 +166,9 @@ where · match List.mem_or_eq_of_mem_set hl with | .inl hl => exact hs₁ _ hl | .inr e => exact e ▸ .nil - · simp [Array.getElem_eq_data_get, List.get_set]; split + · simp only [Array.data_length, Array.size_set, Array.getElem_eq_data_get, Array.data_set, + List.get_set] + split · nofun · exact hs₂ _ (by simp_all) · let rank (k : α) := ((hash k).toUSize % source.size).toNat @@ -182,7 +193,7 @@ theorem insert_size [BEq α] [Hashable α] {m : Imp α β} {k v} · unfold Buckets.size refine have ⟨_, _, h₁, _, eq⟩ := Buckets.exists_of_update ..; eq ▸ ?_ simp [h, h₁, Buckets.size_eq, Nat.succ_add]; rfl - · rw [expand_size]; simp [h, expand, Buckets.size] + · rw [expand_size]; simp only [expand, h, Buckets.size, Array.data_length, Buckets.update_size] refine have ⟨_, _, h₁, _, eq⟩ := Buckets.exists_of_update ..; eq ▸ ?_ simp [h₁, Buckets.size_eq, Nat.succ_add]; rfl @@ -191,7 +202,8 @@ private theorem mem_replaceF {l : List (α × β)} {x : α × β} {p : α × β induction l with | nil => exact .inr | cons a l ih => - simp; generalize e : cond .. = z; revert e + simp only [List.replaceF, List.mem_cons] + generalize e : cond .. = z; revert e unfold cond; split <;> (intro h; subst h; simp) · intro | .inl eq => exact eq ▸ .inl rfl @@ -208,7 +220,7 @@ private theorem pairwise_replaceF [BEq α] [PartialEquivBEq α] induction l with | nil => simp [H] | cons a l ih => - simp at H ⊢ + simp only [List.pairwise_cons, List.replaceF] at H ⊢ generalize e : cond .. = z; unfold cond at e; revert e split <;> (intro h; subst h; simp) · next e => exact ⟨(H.1 · · ∘ PartialEquivBEq.trans e), H.2⟩ @@ -222,15 +234,17 @@ theorem insert_WF [BEq α] [Hashable α] {m : Imp α β} {k v} (h : m.buckets.WF) : (insert m k v).buckets.WF := by dsimp [insert, cond]; split · next h₁ => - simp at h₁; have ⟨x, hx₁, hx₂⟩ := h₁ + simp only [AssocList.contains_eq, List.any_eq_true] at h₁; have ⟨x, hx₁, hx₂⟩ := h₁ refine h.update (fun H => ?_) (fun H a h => ?_) - · simp; exact pairwise_replaceF H - · simp [AssocList.All] at H h ⊢ + · simp only [AssocList.toList_replace] + exact pairwise_replaceF H + · simp only [AssocList.All, Array.ugetElem_eq_getElem, AssocList.toList_replace] at H h ⊢ match mem_replaceF h with | .inl rfl => rfl | .inr h => exact H _ h · next h₁ => - rw [Bool.eq_false_iff] at h₁; simp at h₁ + rw [Bool.eq_false_iff] at h₁ + simp only [AssocList.contains_eq, ne_eq, List.any_eq_true, not_exists, not_and] at h₁ suffices _ by split <;> [exact this; refine expand_WF this] refine h.update (.cons ?_) (fun H a h => ?_) · exact fun a h h' => h₁ a h (PartialEquivBEq.symm h') @@ -243,12 +257,13 @@ theorem erase_size [BEq α] [Hashable α] {m : Imp α β} {k} (erase m k).size = (erase m k).buckets.size := by dsimp [erase, cond]; split · next H => - simp [h, Buckets.size] + simp only [h, Buckets.size] refine have ⟨_, _, h₁, _, eq⟩ := Buckets.exists_of_update ..; eq ▸ ?_ - simp [h, h₁, Buckets.size_eq] + simp only [h₁, Array.data_length, Array.ugetElem_eq_getElem, List.map_append, List.map_cons, + Nat.sum_append, Nat.sum_cons, AssocList.toList_erase] rw [(_ : List.length _ = _ + 1), Nat.add_right_comm]; {rfl} clear h₁ eq - simp [AssocList.contains_eq] at H + simp only [AssocList.contains_eq, List.any_eq_true] at H have ⟨a, h₁, h₂⟩ := H refine have ⟨_, _, _, _, _, h, eq⟩ := List.exists_of_eraseP h₁ h₂; eq ▸ ?_ simp [h]; rfl @@ -257,7 +272,7 @@ theorem erase_size [BEq α] [Hashable α] {m : Imp α β} {k} theorem erase_WF [BEq α] [Hashable α] {m : Imp α β} {k} (h : m.buckets.WF) : (erase m k).buckets.WF := by dsimp [erase, cond]; split - · refine h.update (fun H => ?_) (fun H a h => ?_) <;> simp at h ⊢ + · refine h.update (fun H => ?_) (fun H a h => ?_) <;> simp only [AssocList.toList_erase] at h ⊢ · exact H.sublist (List.eraseP_sublist _) · exact H _ (List.mem_of_mem_eraseP h) · exact h @@ -266,7 +281,7 @@ theorem modify_size [BEq α] [Hashable α] {m : Imp α β} {k} (h : m.size = m.buckets.size) : (modify m k f).size = (modify m k f).buckets.size := by dsimp [modify, cond]; rw [Buckets.update_update] - simp [h, Buckets.size] + simp only [h, Buckets.size] refine have ⟨_, _, h₁, _, eq⟩ := Buckets.exists_of_update ..; eq ▸ ?_ simp [h, h₁, Buckets.size_eq] @@ -275,7 +290,7 @@ theorem modify_WF [BEq α] [Hashable α] {m : Imp α β} {k} dsimp [modify, cond]; rw [Buckets.update_update] refine h.update (fun H => ?_) (fun H a h => ?_) <;> simp at h ⊢ · exact pairwise_replaceF H - · simp [AssocList.All] at H h ⊢ + · simp only [AssocList.All, Array.ugetElem_eq_getElem] at H h ⊢ match mem_replaceF h with | .inl rfl => rfl | .inr h => exact H _ h @@ -296,12 +311,13 @@ theorem WF_iff [BEq α] [Hashable α] {m : Imp α β} : theorem WF.mapVal {α β γ} {f : α → β → γ} [BEq α] [Hashable α] {m : Imp α β} (H : WF m) : WF (mapVal f m) := by have ⟨h₁, h₂⟩ := H.out - simp [Imp.mapVal, Buckets.mapVal, WF_iff, h₁]; refine ⟨?_, ?_, fun i h => ?_⟩ - · simp [Buckets.size]; congr; funext l; simp + simp only [Imp.mapVal, h₁, Buckets.mapVal, WF_iff]; refine ⟨?_, ?_, fun i h => ?_⟩ + · simp only [Buckets.size, Array.map_data, List.map_map]; congr; funext l; simp · simp only [Array.map_data, List.forall_mem_map_iff] simp only [AssocList.toList_mapVal, List.pairwise_map] exact fun _ => h₂.1 _ - · simp [AssocList.All] at h ⊢ + · simp only [Array.size_map, AssocList.All, Array.getElem_map, AssocList.toList_mapVal, + List.mem_map, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] at h ⊢ intro a m apply h₂.2 _ _ _ m @@ -313,7 +329,11 @@ theorem WF.filterMap {α β γ} {f : α → β → Option γ} [BEq α] [Hashable induction l generalizing n acc with simp [filterMap.go, g₁, *] | cons a b l => match f a b with | none => rfl - | some c => simp; rw [Nat.add_right_comm]; rfl + | some c => + simp only [Option.map_some', List.reverse_cons, List.append_assoc, List.singleton_append, + List.length_cons, Nat.succ_eq_add_one, Prod.mk.injEq, true_and] + rw [Nat.add_right_comm] + rfl let g l := (g₁ l).reverse.toAssocList let M := StateT (ULift Nat) Id have H2 (l : List (AssocList α β)) n : @@ -334,7 +354,7 @@ theorem WF.filterMap {α β γ} {f : α → β → Option γ} [BEq α] [Hashable suffices ∀ bk sz (h : 0 < bk.length), m.buckets.val.mapM (m := M) (filterMap.go f .nil) ⟨0⟩ = (⟨bk⟩, ⟨sz⟩) → WF ⟨sz, ⟨bk⟩, h⟩ from this _ _ _ rfl - simp [Array.mapM_eq_mapM_data, bind, StateT.bind, H2, g] + simp only [Array.mapM_eq_mapM_data, bind, StateT.bind, H2, List.map_map, Nat.zero_add, g] intro bk sz h e'; cases e' refine .mk (by simp [Buckets.size]) ⟨?_, fun i h => ?_⟩ · simp only [List.forall_mem_map_iff, List.toList_toAssocList] @@ -345,7 +365,8 @@ theorem WF.filterMap {α β γ} {f : α → β → Option γ} [BEq α] [Hashable · simp only [Array.size_mk, List.length_map, Array.data_length, Array.getElem_eq_data_get, List.get_map] at h ⊢ have := H.out.2.2 _ h - simp [AssocList.All, g₁] at this ⊢ + simp only [AssocList.All, List.toList_toAssocList, List.mem_reverse, List.mem_filterMap, + Option.map_eq_some', forall_exists_index, and_imp, g₁] at this ⊢ rintro _ _ h' _ _ rfl exact this _ h' From f96a34401de084c73c787ecb45b85d4fb47bb981 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 12 Jun 2024 02:26:30 +0100 Subject: [PATCH 002/177] chore(RBSet): consolidate API from mathlib (#834) --- Batteries/Data/RBMap/Basic.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Batteries/Data/RBMap/Basic.lean b/Batteries/Data/RBMap/Basic.lean index ae60899f26..a9ba215ad9 100644 --- a/Batteries/Data/RBMap/Basic.lean +++ b/Batteries/Data/RBMap/Basic.lean @@ -678,6 +678,16 @@ instance [Repr α] : Repr (RBSet α cmp) where /-- `O(log n)`. Insert element `v` into the tree. -/ @[inline] def insert (t : RBSet α cmp) (v : α) : RBSet α cmp := ⟨t.1.insert cmp v, t.2.insert⟩ +/-- +Insert all elements from a collection into a `RBSet α cmp`. +-/ +def insertMany [ForIn Id ρ α] (s : RBSet α cmp) (as : ρ) : + RBSet α cmp := Id.run do + let mut s := s + for a in as do + s := s.insert a + return s + /-- `O(log n)`. Remove an element from the tree using a cut function. The `cut` function is used to locate an element in the tree: From 42b5dddbd6b2658fcfede9dad26cc47737edec2d Mon Sep 17 00:00:00 2001 From: L Date: Tue, 11 Jun 2024 23:15:19 -0700 Subject: [PATCH 003/177] fix(linter): "(invalid MessageData.lazy, missing context)" (#838) --- Batteries/Tactic/Lint/Frontend.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/Batteries/Tactic/Lint/Frontend.lean b/Batteries/Tactic/Lint/Frontend.lean index 4eff8b1450..072e23743f 100644 --- a/Batteries/Tactic/Lint/Frontend.lean +++ b/Batteries/Tactic/Lint/Frontend.lean @@ -125,9 +125,11 @@ def printWarning (declName : Name) (warning : MessageData) (useErrorFormat : Boo (filePath : System.FilePath := default) : CoreM MessageData := do if useErrorFormat then if let some range ← findDeclarationRanges? declName then - return m!"{filePath}:{range.range.pos.line}:{range.range.pos.column + 1}: error: { + let msg ← addMessageContextPartial + m!"{filePath}:{range.range.pos.line}:{range.range.pos.column + 1}: error: { ← mkConstWithLevelParams declName} {warning}" - pure m!"#check {← mkConstWithLevelParams declName} /- {warning} -/" + return msg + addMessageContextPartial m!"#check {← mkConstWithLevelParams declName} /- {warning} -/" /-- Formats a map of linter warnings using `print_warning`, sorted by line number. -/ def printWarnings (results : HashMap Name MessageData) (filePath : System.FilePath := default) From 007a82fab9671c960ad260d4e4bf664b9fb884e3 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 12 Jun 2024 14:15:41 +0200 Subject: [PATCH 004/177] chore(Tactic/Lint/Simp): make simpNF linter dsimp-aware (#839) * add test demonstrating undesired behaviour on dsimp lemmas * chore(Tactic/Lint/Simp): make simpNF linter dsimp-aware * test demonstrates correct behaviour on dsimp lemmas * process review comments * make formatLemmas dsimp-aware --- Batteries/Tactic/Lint/Simp.lean | 26 +++++++++++++++++++------- test/lint_simpNF.lean | 14 ++++++++++++++ 2 files changed, 33 insertions(+), 7 deletions(-) diff --git a/Batteries/Tactic/Lint/Simp.lean b/Batteries/Tactic/Lint/Simp.lean index 44a858b340..3dc801fe0a 100644 --- a/Batteries/Tactic/Lint/Simp.lean +++ b/Batteries/Tactic/Lint/Simp.lean @@ -93,14 +93,14 @@ def decorateError (msg : MessageData) (k : MetaM α) : MetaM α := do try k catch e => throw (.error e.getRef m!"{msg}\n{e.toMessageData}") /-- Render the list of simp lemmas. -/ -def formatLemmas (usedSimps : Simp.UsedSimps) : MetaM MessageData := do +def formatLemmas (usedSimps : Simp.UsedSimps) (simpName : String) : MetaM MessageData := do let mut args := #[] let env ← getEnv for (thm, _) in usedSimps.toArray.qsort (·.2 < ·.2) do if let .decl declName := thm then if env.contains declName && declName != ``eq_self then args := args.push (← mkConstWithFreshMVarLevels declName) - return m!"simp only {args.toList}" + return m!"{simpName} only {args.toList}" /-- A linter for simp lemmas whose lhs is not in simp-normal form, and which hence never fire. -/ @[env_linter] def simpNF : Linter where @@ -112,17 +112,29 @@ https://leanprover-community.github.io/mathlib_docs/notes.html#simp-normal%20for unless ← isSimpTheorem declName do return none let ctx := { ← Simp.Context.mkDefault with config.decide := false } checkAllSimpTheoremInfos (← getConstInfo declName).type fun {lhs, rhs, isConditional, ..} => do + let isRfl ← isRflTheorem declName let ({ expr := lhs', proof? := prf1, .. }, prf1Stats) ← - decorateError "simplify fails on left-hand side:" <| simp lhs ctx + decorateError "simplify fails on left-hand side:" <| + if !isRfl then + simp lhs ctx + else do + let (e, s) ← dsimp lhs ctx + return (Simp.Result.mk e .none .true, s) if prf1Stats.usedTheorems.contains (.decl declName) then return none let ({ expr := rhs', .. }, stats) ← - decorateError "simplify fails on right-hand side:" <| simp rhs ctx (stats := prf1Stats) + decorateError "simplify fails on right-hand side:" <| + if !isRfl then + simp rhs ctx (stats := prf1Stats) + else do + let (e, s) ← dsimp rhs ctx (stats := prf1Stats) + return (Simp.Result.mk e .none .true, s) let lhs'EqRhs' ← isSimpEq lhs' rhs' (whnfFirst := false) let lhsInNF ← isSimpEq lhs' lhs + let simpName := if !isRfl then "simp" else "dsimp" if lhs'EqRhs' then if prf1.isNone then return none -- TODO: FP rewriting foo.eq_2 using `simp only [foo]` - return m!"simp can prove this: - by {← formatLemmas stats.usedTheorems} + return m!"{simpName} can prove this: + by {← formatLemmas stats.usedTheorems simpName} One of the lemmas above could be a duplicate. If that's not the case try reordering lemmas or adding @[priority]. " @@ -132,7 +144,7 @@ If that's not the case try reordering lemmas or adding @[priority]. to {lhs'} using - {← formatLemmas prf1Stats.usedTheorems} + {← formatLemmas prf1Stats.usedTheorems simpName} Try to change the left-hand side to the simplified term! " else if !isConditional && lhs == lhs' then diff --git a/test/lint_simpNF.lean b/test/lint_simpNF.lean index 8587e464bb..e06e067d18 100644 --- a/test/lint_simpNF.lean +++ b/test/lint_simpNF.lean @@ -27,4 +27,18 @@ theorem sumCompl_apply_symm_of_pos (p : α → Prop) [DecidablePred p] (a : α) (sumCompl p).symm a = Sum.inl ⟨a, h⟩ := dif_pos h +def foo (n : Nat) : Nat := if n = n then n else 0 + +@[simp] +theorem foo_eq_id : foo = id := by + funext n + simp [foo] + +-- The following `dsimp`-lemma is (correctly) not flagged by the linter +@[simp] +theorem foo_eq_ite (n : Nat) : foo n = if n = n then n else 0 := by + rfl + +end Equiv + #lint- only simpNF From 0f5b28745722ab5f5168d2a418cabf8bee98f97a Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Thu, 13 Jun 2024 06:37:19 +0200 Subject: [PATCH 005/177] Revert "chore(Tactic/Lint/Simp): make simpNF linter dsimp-aware (#839)" (#841) This reverts commit 007a82fab9671c960ad260d4e4bf664b9fb884e3. --- Batteries/Tactic/Lint/Simp.lean | 26 +++++++------------------- test/lint_simpNF.lean | 14 -------------- 2 files changed, 7 insertions(+), 33 deletions(-) diff --git a/Batteries/Tactic/Lint/Simp.lean b/Batteries/Tactic/Lint/Simp.lean index 3dc801fe0a..44a858b340 100644 --- a/Batteries/Tactic/Lint/Simp.lean +++ b/Batteries/Tactic/Lint/Simp.lean @@ -93,14 +93,14 @@ def decorateError (msg : MessageData) (k : MetaM α) : MetaM α := do try k catch e => throw (.error e.getRef m!"{msg}\n{e.toMessageData}") /-- Render the list of simp lemmas. -/ -def formatLemmas (usedSimps : Simp.UsedSimps) (simpName : String) : MetaM MessageData := do +def formatLemmas (usedSimps : Simp.UsedSimps) : MetaM MessageData := do let mut args := #[] let env ← getEnv for (thm, _) in usedSimps.toArray.qsort (·.2 < ·.2) do if let .decl declName := thm then if env.contains declName && declName != ``eq_self then args := args.push (← mkConstWithFreshMVarLevels declName) - return m!"{simpName} only {args.toList}" + return m!"simp only {args.toList}" /-- A linter for simp lemmas whose lhs is not in simp-normal form, and which hence never fire. -/ @[env_linter] def simpNF : Linter where @@ -112,29 +112,17 @@ https://leanprover-community.github.io/mathlib_docs/notes.html#simp-normal%20for unless ← isSimpTheorem declName do return none let ctx := { ← Simp.Context.mkDefault with config.decide := false } checkAllSimpTheoremInfos (← getConstInfo declName).type fun {lhs, rhs, isConditional, ..} => do - let isRfl ← isRflTheorem declName let ({ expr := lhs', proof? := prf1, .. }, prf1Stats) ← - decorateError "simplify fails on left-hand side:" <| - if !isRfl then - simp lhs ctx - else do - let (e, s) ← dsimp lhs ctx - return (Simp.Result.mk e .none .true, s) + decorateError "simplify fails on left-hand side:" <| simp lhs ctx if prf1Stats.usedTheorems.contains (.decl declName) then return none let ({ expr := rhs', .. }, stats) ← - decorateError "simplify fails on right-hand side:" <| - if !isRfl then - simp rhs ctx (stats := prf1Stats) - else do - let (e, s) ← dsimp rhs ctx (stats := prf1Stats) - return (Simp.Result.mk e .none .true, s) + decorateError "simplify fails on right-hand side:" <| simp rhs ctx (stats := prf1Stats) let lhs'EqRhs' ← isSimpEq lhs' rhs' (whnfFirst := false) let lhsInNF ← isSimpEq lhs' lhs - let simpName := if !isRfl then "simp" else "dsimp" if lhs'EqRhs' then if prf1.isNone then return none -- TODO: FP rewriting foo.eq_2 using `simp only [foo]` - return m!"{simpName} can prove this: - by {← formatLemmas stats.usedTheorems simpName} + return m!"simp can prove this: + by {← formatLemmas stats.usedTheorems} One of the lemmas above could be a duplicate. If that's not the case try reordering lemmas or adding @[priority]. " @@ -144,7 +132,7 @@ If that's not the case try reordering lemmas or adding @[priority]. to {lhs'} using - {← formatLemmas prf1Stats.usedTheorems simpName} + {← formatLemmas prf1Stats.usedTheorems} Try to change the left-hand side to the simplified term! " else if !isConditional && lhs == lhs' then diff --git a/test/lint_simpNF.lean b/test/lint_simpNF.lean index e06e067d18..8587e464bb 100644 --- a/test/lint_simpNF.lean +++ b/test/lint_simpNF.lean @@ -27,18 +27,4 @@ theorem sumCompl_apply_symm_of_pos (p : α → Prop) [DecidablePred p] (a : α) (sumCompl p).symm a = Sum.inl ⟨a, h⟩ := dif_pos h -def foo (n : Nat) : Nat := if n = n then n else 0 - -@[simp] -theorem foo_eq_id : foo = id := by - funext n - simp [foo] - --- The following `dsimp`-lemma is (correctly) not flagged by the linter -@[simp] -theorem foo_eq_ite (n : Nat) : foo n = if n = n then n else 0 := by - rfl - -end Equiv - #lint- only simpNF From 15d42e1a92a80d0db5ca1c12123866ba392b9d76 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Thu, 13 Jun 2024 20:50:52 -0400 Subject: [PATCH 006/177] feat: ext lemma for `Thunk` (#842) --- Batteries.lean | 1 + Batteries/Data/Thunk.lean | 10 ++++++++++ 2 files changed, 11 insertions(+) create mode 100644 Batteries/Data/Thunk.lean diff --git a/Batteries.lean b/Batteries.lean index f56659960b..dea5262d3b 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -35,6 +35,7 @@ import Batteries.Data.Range import Batteries.Data.Rat import Batteries.Data.String import Batteries.Data.Sum +import Batteries.Data.Thunk import Batteries.Data.UInt import Batteries.Data.UnionFind import Batteries.Lean.AttributeExtra diff --git a/Batteries/Data/Thunk.lean b/Batteries/Data/Thunk.lean new file mode 100644 index 0000000000..88c92c3784 --- /dev/null +++ b/Batteries/Data/Thunk.lean @@ -0,0 +1,10 @@ +/- +Copyright (c) 2024 François G. Dorais. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE +Authors: François G. Dorais, et al. +-/ + +namespace Thunk + +@[ext] protected theorem ext : {a b : Thunk α} → a.get = b.get → a = b + | {..}, {..}, heq => congrArg _ <| funext fun _ => heq From 2506c85498141bc8ab3ee745bcb06021c3ed9a80 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 15 Jun 2024 11:02:44 +0100 Subject: [PATCH 007/177] chore: remove docgen from the manifest (#846) --- lake-manifest.json | 49 ++-------------------------------------------- 1 file changed, 2 insertions(+), 47 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index e33f34b1fa..68d5a83837 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -1,50 +1,5 @@ -{"version": 7, +{"version": "1.0.0", "packagesDir": ".lake/packages", - "packages": - [{"url": "https://github.com/xubaiw/CMark.lean", - "type": "git", - "subDir": null, - "rev": "0077cbbaa92abf855fc1c0413e158ffd8195ec77", - "name": "CMark", - "manifestFile": "lake-manifest.json", - "inputRev": "main", - "inherited": true, - "configFile": "lakefile.lean"}, - {"url": "https://github.com/fgdorais/lean4-unicode-basic", - "type": "git", - "subDir": null, - "rev": "280d75fdfe7be8eb337be7f1bf8479b4aac09f71", - "name": "UnicodeBasic", - "manifestFile": "lake-manifest.json", - "inputRev": "main", - "inherited": true, - "configFile": "lakefile.lean"}, - {"url": "https://github.com/mhuisi/lean4-cli", - "type": "git", - "subDir": null, - "rev": "39229f3630d734af7d9cfb5937ddc6b41d3aa6aa", - "name": "Cli", - "manifestFile": "lake-manifest.json", - "inputRev": "nightly", - "inherited": true, - "configFile": "lakefile.lean"}, - {"url": "https://github.com/hargonix/LeanInk", - "type": "git", - "subDir": null, - "rev": "2447df5cc6e48eb965c3c3fba87e46d353b5e9f1", - "name": "leanInk", - "manifestFile": "lake-manifest.json", - "inputRev": "doc-gen", - "inherited": true, - "configFile": "lakefile.lean"}, - {"url": "https://github.com/leanprover/doc-gen4", - "type": "git", - "subDir": null, - "rev": "b7fad51b87a5f8fb3a238dc820ec40ebfa2a636e", - "name": "«doc-gen4»", - "manifestFile": "lake-manifest.json", - "inputRev": "main", - "inherited": false, - "configFile": "lakefile.lean"}], + "packages": [], "name": "batteries", "lakeDir": ".lake"} From 194b50cb40f6aaf5596e378ec612f24985d5a1a5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sat, 15 Jun 2024 20:23:47 -0400 Subject: [PATCH 008/177] chore: bump toolchain to v4.9.0-rc2 (#847) --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 0ba3faf807..29c0cea43e 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.9.0-rc1 +leanprover/lean4:v4.9.0-rc2 From 47e4cc5c5800c07d9bf232173c9941fa5bf68589 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Sun, 16 Jun 2024 00:31:19 +0000 Subject: [PATCH 009/177] Reapply "chore(Tactic/Lint/Simp): make simpNF linter dsimp-aware (#839)" (#841) (#844) This reverts commit 0f5b28745722ab5f5168d2a418cabf8bee98f97a. --- Batteries/Tactic/Lint/Simp.lean | 26 +++++++++++++++++++------- test/lint_simpNF.lean | 14 ++++++++++++++ 2 files changed, 33 insertions(+), 7 deletions(-) diff --git a/Batteries/Tactic/Lint/Simp.lean b/Batteries/Tactic/Lint/Simp.lean index 44a858b340..3dc801fe0a 100644 --- a/Batteries/Tactic/Lint/Simp.lean +++ b/Batteries/Tactic/Lint/Simp.lean @@ -93,14 +93,14 @@ def decorateError (msg : MessageData) (k : MetaM α) : MetaM α := do try k catch e => throw (.error e.getRef m!"{msg}\n{e.toMessageData}") /-- Render the list of simp lemmas. -/ -def formatLemmas (usedSimps : Simp.UsedSimps) : MetaM MessageData := do +def formatLemmas (usedSimps : Simp.UsedSimps) (simpName : String) : MetaM MessageData := do let mut args := #[] let env ← getEnv for (thm, _) in usedSimps.toArray.qsort (·.2 < ·.2) do if let .decl declName := thm then if env.contains declName && declName != ``eq_self then args := args.push (← mkConstWithFreshMVarLevels declName) - return m!"simp only {args.toList}" + return m!"{simpName} only {args.toList}" /-- A linter for simp lemmas whose lhs is not in simp-normal form, and which hence never fire. -/ @[env_linter] def simpNF : Linter where @@ -112,17 +112,29 @@ https://leanprover-community.github.io/mathlib_docs/notes.html#simp-normal%20for unless ← isSimpTheorem declName do return none let ctx := { ← Simp.Context.mkDefault with config.decide := false } checkAllSimpTheoremInfos (← getConstInfo declName).type fun {lhs, rhs, isConditional, ..} => do + let isRfl ← isRflTheorem declName let ({ expr := lhs', proof? := prf1, .. }, prf1Stats) ← - decorateError "simplify fails on left-hand side:" <| simp lhs ctx + decorateError "simplify fails on left-hand side:" <| + if !isRfl then + simp lhs ctx + else do + let (e, s) ← dsimp lhs ctx + return (Simp.Result.mk e .none .true, s) if prf1Stats.usedTheorems.contains (.decl declName) then return none let ({ expr := rhs', .. }, stats) ← - decorateError "simplify fails on right-hand side:" <| simp rhs ctx (stats := prf1Stats) + decorateError "simplify fails on right-hand side:" <| + if !isRfl then + simp rhs ctx (stats := prf1Stats) + else do + let (e, s) ← dsimp rhs ctx (stats := prf1Stats) + return (Simp.Result.mk e .none .true, s) let lhs'EqRhs' ← isSimpEq lhs' rhs' (whnfFirst := false) let lhsInNF ← isSimpEq lhs' lhs + let simpName := if !isRfl then "simp" else "dsimp" if lhs'EqRhs' then if prf1.isNone then return none -- TODO: FP rewriting foo.eq_2 using `simp only [foo]` - return m!"simp can prove this: - by {← formatLemmas stats.usedTheorems} + return m!"{simpName} can prove this: + by {← formatLemmas stats.usedTheorems simpName} One of the lemmas above could be a duplicate. If that's not the case try reordering lemmas or adding @[priority]. " @@ -132,7 +144,7 @@ If that's not the case try reordering lemmas or adding @[priority]. to {lhs'} using - {← formatLemmas prf1Stats.usedTheorems} + {← formatLemmas prf1Stats.usedTheorems simpName} Try to change the left-hand side to the simplified term! " else if !isConditional && lhs == lhs' then diff --git a/test/lint_simpNF.lean b/test/lint_simpNF.lean index 8587e464bb..e06e067d18 100644 --- a/test/lint_simpNF.lean +++ b/test/lint_simpNF.lean @@ -27,4 +27,18 @@ theorem sumCompl_apply_symm_of_pos (p : α → Prop) [DecidablePred p] (a : α) (sumCompl p).symm a = Sum.inl ⟨a, h⟩ := dif_pos h +def foo (n : Nat) : Nat := if n = n then n else 0 + +@[simp] +theorem foo_eq_id : foo = id := by + funext n + simp [foo] + +-- The following `dsimp`-lemma is (correctly) not flagged by the linter +@[simp] +theorem foo_eq_ite (n : Nat) : foo n = if n = n then n else 0 := by + rfl + +end Equiv + #lint- only simpNF From a962bdc0eea7d9ad6bfaf450ad0f94c3ebc69ab1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 17 Jun 2024 18:29:59 -0400 Subject: [PATCH 010/177] feat: add `BinaryHeap` (#849) Co-authored-by: Mario Carneiro --- Batteries.lean | 1 + Batteries/Data/BinaryHeap.lean | 176 +++++++++++++++++++++++++++++++++ 2 files changed, 177 insertions(+) create mode 100644 Batteries/Data/BinaryHeap.lean diff --git a/Batteries.lean b/Batteries.lean index dea5262d3b..706e967883 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -15,6 +15,7 @@ import Batteries.Control.Lemmas import Batteries.Control.Nondet.Basic import Batteries.Data.Array import Batteries.Data.AssocList +import Batteries.Data.BinaryHeap import Batteries.Data.BinomialHeap import Batteries.Data.BitVec import Batteries.Data.Bool diff --git a/Batteries/Data/BinaryHeap.lean b/Batteries/Data/BinaryHeap.lean new file mode 100644 index 0000000000..b36ed97fd3 --- /dev/null +++ b/Batteries/Data/BinaryHeap.lean @@ -0,0 +1,176 @@ +/- +Copyright (c) 2021 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ + +namespace Batteries + +/-- A max-heap data structure. -/ +structure BinaryHeap (α) (lt : α → α → Bool) where + /-- Backing array for `BinaryHeap`. -/ + arr : Array α + +namespace BinaryHeap + +/-- Core operation for binary heaps, expressed directly on arrays. +Given an array which is a max-heap, push item `i` down to restore the max-heap property. -/ +def heapifyDown (lt : α → α → Bool) (a : Array α) (i : Fin a.size) : + {a' : Array α // a'.size = a.size} := + let left := 2 * i.1 + 1 + let right := left + 1 + have left_le : i ≤ left := Nat.le_trans + (by rw [Nat.succ_mul, Nat.one_mul]; exact Nat.le_add_left i i) + (Nat.le_add_right ..) + have right_le : i ≤ right := Nat.le_trans left_le (Nat.le_add_right ..) + have i_le : i ≤ i := Nat.le_refl _ + have j : {j : Fin a.size // i ≤ j} := if h : left < a.size then + if lt (a.get i) (a.get ⟨left, h⟩) then ⟨⟨left, h⟩, left_le⟩ else ⟨i, i_le⟩ else ⟨i, i_le⟩ + have j := if h : right < a.size then + if lt (a.get j) (a.get ⟨right, h⟩) then ⟨⟨right, h⟩, right_le⟩ else j else j + if h : i.1 = j then ⟨a, rfl⟩ else + let a' := a.swap i j + let j' := ⟨j, by rw [a.size_swap i j]; exact j.1.2⟩ + have : a'.size - j < a.size - i := by + rw [a.size_swap i j]; exact Nat.sub_lt_sub_left i.2 <| Nat.lt_of_le_of_ne j.2 h + let ⟨a₂, h₂⟩ := heapifyDown lt a' j' + ⟨a₂, h₂.trans (a.size_swap i j)⟩ +termination_by a.size - i + +@[simp] theorem size_heapifyDown (lt : α → α → Bool) (a : Array α) (i : Fin a.size) : + (heapifyDown lt a i).1.size = a.size := (heapifyDown lt a i).2 + +/-- Core operation for binary heaps, expressed directly on arrays. +Construct a heap from an unsorted array, by heapifying all the elements. -/ +def mkHeap (lt : α → α → Bool) (a : Array α) : {a' : Array α // a'.size = a.size} := + loop (a.size / 2) a (Nat.div_le_self ..) +where + /-- Inner loop for `mkHeap`. -/ + loop : (i : Nat) → (a : Array α) → i ≤ a.size → {a' : Array α // a'.size = a.size} + | 0, a, _ => ⟨a, rfl⟩ + | i+1, a, h => + let h := Nat.lt_of_succ_le h + let a' := heapifyDown lt a ⟨i, h⟩ + let ⟨a₂, h₂⟩ := loop i a' ((heapifyDown ..).2.symm ▸ Nat.le_of_lt h) + ⟨a₂, h₂.trans a'.2⟩ + +@[simp] theorem size_mkHeap (lt : α → α → Bool) (a : Array α) : + (mkHeap lt a).1.size = a.size := (mkHeap lt a).2 + +/-- Core operation for binary heaps, expressed directly on arrays. +Given an array which is a max-heap, push item `i` up to restore the max-heap property. -/ +def heapifyUp (lt : α → α → Bool) (a : Array α) (i : Fin a.size) : + {a' : Array α // a'.size = a.size} := + if i0 : i.1 = 0 then ⟨a, rfl⟩ else + have : (i.1 - 1) / 2 < i := Nat.lt_of_le_of_lt (Nat.div_le_self ..) <| + Nat.sub_lt (Nat.pos_of_ne_zero i0) Nat.zero_lt_one + let j := ⟨(i.1 - 1) / 2, Nat.lt_trans this i.2⟩ + if lt (a.get j) (a.get i) then + let a' := a.swap i j + let ⟨a₂, h₂⟩ := heapifyUp lt a' ⟨j.1, by rw [a.size_swap i j]; exact j.2⟩ + ⟨a₂, h₂.trans (a.size_swap i j)⟩ + else ⟨a, rfl⟩ + +@[simp] theorem size_heapifyUp (lt : α → α → Bool) (a : Array α) (i : Fin a.size) : + (heapifyUp lt a i).1.size = a.size := (heapifyUp lt a i).2 + +/-- `O(1)`. Build a new empty heap. -/ +def empty (lt) : BinaryHeap α lt := ⟨#[]⟩ + +instance (lt) : Inhabited (BinaryHeap α lt) := ⟨empty _⟩ +instance (lt) : EmptyCollection (BinaryHeap α lt) := ⟨empty _⟩ + +/-- `O(1)`. Build a one-element heap. -/ +def singleton (lt) (x : α) : BinaryHeap α lt := ⟨#[x]⟩ + +/-- `O(1)`. Get the number of elements in a `BinaryHeap`. -/ +def size (self : BinaryHeap α lt) : Nat := self.1.size + +/-- `O(1)`. Get an element in the heap by index. -/ +def get (self : BinaryHeap α lt) (i : Fin self.size) : α := self.1.get i + +/-- `O(log n)`. Insert an element into a `BinaryHeap`, preserving the max-heap property. -/ +def insert (self : BinaryHeap α lt) (x : α) : BinaryHeap α lt where + arr := let n := self.size; + heapifyUp lt (self.1.push x) ⟨n, by rw [Array.size_push]; apply Nat.lt_succ_self⟩ + +@[simp] theorem size_insert (self : BinaryHeap α lt) (x : α) : + (self.insert x).size = self.size + 1 := by + simp [insert, size, size_heapifyUp] + +/-- `O(1)`. Get the maximum element in a `BinaryHeap`. -/ +def max (self : BinaryHeap α lt) : Option α := self.1.get? 0 + +/-- Auxiliary for `popMax`. -/ +def popMaxAux (self : BinaryHeap α lt) : {a' : BinaryHeap α lt // a'.size = self.size - 1} := + match e: self.1.size with + | 0 => ⟨self, by simp [size, e]⟩ + | n+1 => + have h0 := by rw [e]; apply Nat.succ_pos + have hn := by rw [e]; apply Nat.lt_succ_self + if hn0 : 0 < n then + let a := self.1.swap ⟨0, h0⟩ ⟨n, hn⟩ |>.pop + ⟨⟨heapifyDown lt a ⟨0, by rwa [Array.size_pop, Array.size_swap, e]⟩⟩, + by simp [size, a]⟩ + else + ⟨⟨self.1.pop⟩, by simp [size]⟩ + +/-- `O(log n)`. Remove the maximum element from a `BinaryHeap`. +Call `max` first to actually retrieve the maximum element. -/ +def popMax (self : BinaryHeap α lt) : BinaryHeap α lt := self.popMaxAux + +@[simp] theorem size_popMax (self : BinaryHeap α lt) : + self.popMax.size = self.size - 1 := self.popMaxAux.2 + +/-- `O(log n)`. Return and remove the maximum element from a `BinaryHeap`. -/ +def extractMax (self : BinaryHeap α lt) : Option α × BinaryHeap α lt := + (self.max, self.popMax) + +theorem size_pos_of_max {self : BinaryHeap α lt} (e : self.max = some x) : 0 < self.size := + Decidable.of_not_not fun h : ¬ 0 < self.1.size => by simp [BinaryHeap.max, Array.get?, h] at e + +/-- `O(log n)`. Equivalent to `extractMax (self.insert x)`, except that extraction cannot fail. -/ +def insertExtractMax (self : BinaryHeap α lt) (x : α) : α × BinaryHeap α lt := + match e: self.max with + | none => (x, self) + | some m => + if lt x m then + let a := self.1.set ⟨0, size_pos_of_max e⟩ x + (m, ⟨heapifyDown lt a ⟨0, by simp only [Array.size_set, a]; exact size_pos_of_max e⟩⟩) + else (x, self) + +/-- `O(log n)`. Equivalent to `(self.max, self.popMax.insert x)`. -/ +def replaceMax (self : BinaryHeap α lt) (x : α) : Option α × BinaryHeap α lt := + match e: self.max with + | none => (none, ⟨self.1.push x⟩) + | some m => + let a := self.1.set ⟨0, size_pos_of_max e⟩ x + (some m, ⟨heapifyDown lt a ⟨0, by simp only [Array.size_set, a]; exact size_pos_of_max e⟩⟩) + +/-- `O(log n)`. Replace the value at index `i` by `x`. Assumes that `x ≤ self.get i`. -/ +def decreaseKey (self : BinaryHeap α lt) (i : Fin self.size) (x : α) : BinaryHeap α lt where + arr := heapifyDown lt (self.1.set i x) ⟨i, by rw [self.1.size_set]; exact i.2⟩ + +/-- `O(log n)`. Replace the value at index `i` by `x`. Assumes that `self.get i ≤ x`. -/ +def increaseKey (self : BinaryHeap α lt) (i : Fin self.size) (x : α) : BinaryHeap α lt where + arr := heapifyUp lt (self.1.set i x) ⟨i, by rw [self.1.size_set]; exact i.2⟩ + +end Batteries.BinaryHeap + +/-- `O(n)`. Convert an unsorted array to a `BinaryHeap`. -/ +def Array.toBinaryHeap (lt : α → α → Bool) (a : Array α) : Batteries.BinaryHeap α lt where + arr := Batteries.BinaryHeap.mkHeap lt a + +/-- `O(n log n)`. Sort an array using a `BinaryHeap`. -/ +@[specialize] def Array.heapSort (a : Array α) (lt : α → α → Bool) : Array α := + loop (a.toBinaryHeap (flip lt)) #[] +where + /-- Inner loop for `heapSort`. -/ + loop (a : Batteries.BinaryHeap α (flip lt)) (out : Array α) : Array α := + match e: a.max with + | none => out + | some x => + have : a.popMax.size < a.size := by + simp; exact Nat.sub_lt (Batteries.BinaryHeap.size_pos_of_max e) Nat.zero_lt_one + loop a.popMax (out.push x) + termination_by a.size From f34392984409936947306b26898091d7221cb1bf Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 18 Jun 2024 22:08:11 +0100 Subject: [PATCH 011/177] chore: robust proof in HashMap.WF.filterMap (#852) --- Batteries/Data/HashMap/WF.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/Batteries/Data/HashMap/WF.lean b/Batteries/Data/HashMap/WF.lean index fb2635b8fe..57b121c83c 100644 --- a/Batteries/Data/HashMap/WF.lean +++ b/Batteries/Data/HashMap/WF.lean @@ -326,7 +326,9 @@ theorem WF.filterMap {α β γ} {f : α → β → Option γ} [BEq α] [Hashable let g₁ (l : AssocList α β) := l.toList.filterMap (fun x => (f x.1 x.2).map (x.1, ·)) have H1 (l n acc) : filterMap.go f acc l n = (((g₁ l).reverse ++ acc.toList).toAssocList, ⟨n.1 + (g₁ l).length⟩) := by - induction l generalizing n acc with simp [filterMap.go, g₁, *] + induction l generalizing n acc with simp only [filterMap.go, AssocList.toList, + List.filterMap_nil, List.reverse_nil, List.nil_append, AssocList.toList_toAssocList, + List.length_nil, Nat.add_zero, List.filterMap_cons, g₁, *] | cons a b l => match f a b with | none => rfl | some c => @@ -348,7 +350,7 @@ theorem WF.filterMap {α β γ} {f : α → β → Option γ} [BEq α] [Hashable induction l with | nil => exact .slnil | cons a l ih => - simp; exact match f a.1 a.2 with + simp only [List.filterMap_cons, List.map_cons]; exact match f a.1 a.2 with | none => .cons _ ih | some b => .cons₂ _ ih suffices ∀ bk sz (h : 0 < bk.length), From 73c7f46f4c0eaa48e9538c9fd108d3ee5524868a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Thu, 20 Jun 2024 21:06:34 -0400 Subject: [PATCH 012/177] feat: add simps and `toNat` lemmas for `UIntX` types (#853) --- Batteries/Data/UInt.lean | 109 ++++++++++++++++++++++++++++++++++++++- 1 file changed, 107 insertions(+), 2 deletions(-) diff --git a/Batteries/Data/UInt.lean b/Batteries/Data/UInt.lean index 3151a25852..46fcb26b59 100644 --- a/Batteries/Data/UInt.lean +++ b/Batteries/Data/UInt.lean @@ -14,6 +14,12 @@ theorem UInt8.ext_iff {x y : UInt8} : x = y ↔ x.toNat = y.toNat := ⟨congrArg @[simp] theorem UInt8.val_val_eq_toNat (x : UInt8) : x.val.val = x.toNat := rfl +@[simp] theorem UInt8.val_ofNat (n) : + (no_index (OfNat.ofNat n) : UInt8).val = OfNat.ofNat n := rfl + +@[simp] theorem UInt8.toNat_ofNat (n) : + (no_index (OfNat.ofNat n) : UInt8).toNat = n % UInt8.size := rfl + theorem UInt8.toNat_lt (x : UInt8) : x.toNat < 2 ^ 8 := x.val.isLt @[simp] theorem UInt8.toUInt16_toNat (x : UInt8) : x.toUInt16.toNat = x.toNat := rfl @@ -22,6 +28,21 @@ theorem UInt8.toNat_lt (x : UInt8) : x.toNat < 2 ^ 8 := x.val.isLt @[simp] theorem UInt8.toUInt64_toNat (x : UInt8) : x.toUInt64.toNat = x.toNat := rfl +theorem UInt8.toNat_zero : (0 : UInt8).toNat = 0 := rfl + +theorem UInt8.toNat_add (x y : UInt8) : (x + y).toNat = (x.toNat + y.toNat) % UInt8.size := rfl + +theorem UInt8.toNat_sub (x y : UInt8) : + (x - y).toNat = (x.toNat + (UInt8.size - y.toNat)) % UInt8.size := rfl + +theorem UInt8.toNat_mul (x y : UInt8) : (x * y).toNat = (x.toNat * y.toNat) % UInt8.size := rfl + +theorem UInt8.toNat_div (x y : UInt8) : (x / y).toNat = x.toNat / y.toNat := rfl + +theorem UInt8.toNat_mod (x y : UInt8) : (x % y).toNat = x.toNat % y.toNat := rfl + +theorem UInt8.toNat_modn (x : UInt8) (n) : (x.modn n).toNat = x.toNat % n := rfl + theorem UInt8.le_antisymm_iff {x y : UInt8} : x = y ↔ x ≤ y ∧ y ≤ x := UInt8.ext_iff.trans Nat.le_antisymm_iff @@ -38,16 +59,37 @@ instance : Batteries.LawfulOrd UInt8 := .compareOfLessAndEq theorem UInt16.ext_iff {x y : UInt16} : x = y ↔ x.toNat = y.toNat := ⟨congrArg _, UInt16.ext⟩ -theorem UInt16.toNat_lt (x : UInt16) : x.toNat < 2 ^ 16 := x.val.isLt - @[simp] theorem UInt16.val_val_eq_toNat (x : UInt16) : x.val.val = x.toNat := rfl +@[simp] theorem UInt16.val_ofNat (n) : + (no_index (OfNat.ofNat n) : UInt16).val = OfNat.ofNat n := rfl + +@[simp] theorem UInt16.toNat_ofNat (n) : + (no_index (OfNat.ofNat n) : UInt16).toNat = n % UInt16.size := rfl + +theorem UInt16.toNat_lt (x : UInt16) : x.toNat < 2 ^ 16 := x.val.isLt + @[simp] theorem UInt16.toUInt8_toNat (x : UInt16) : x.toUInt8.toNat = x.toNat % 2 ^ 8 := rfl @[simp] theorem UInt16.toUInt32_toNat (x : UInt16) : x.toUInt32.toNat = x.toNat := rfl @[simp] theorem UInt16.toUInt64_toNat (x : UInt16) : x.toUInt64.toNat = x.toNat := rfl +theorem UInt16.toNat_zero : (0 : UInt16).toNat = 0 := rfl + +theorem UInt16.toNat_add (x y : UInt16) : (x + y).toNat = (x.toNat + y.toNat) % UInt16.size := rfl + +theorem UInt16.toNat_sub (x y : UInt16) : + (x - y).toNat = (x.toNat + (UInt16.size - y.toNat)) % UInt16.size := rfl + +theorem UInt16.toNat_mul (x y : UInt16) : (x * y).toNat = (x.toNat * y.toNat) % UInt16.size := rfl + +theorem UInt16.toNat_div (x y : UInt16) : (x / y).toNat = x.toNat / y.toNat := rfl + +theorem UInt16.toNat_mod (x y : UInt16) : (x % y).toNat = x.toNat % y.toNat := rfl + +theorem UInt16.toNat_modn (x : UInt16) (n) : (x.modn n).toNat = x.toNat % n := rfl + theorem UInt16.le_antisymm_iff {x y : UInt16} : x = y ↔ x ≤ y ∧ y ≤ x := UInt16.ext_iff.trans Nat.le_antisymm_iff @@ -66,6 +108,12 @@ theorem UInt32.ext_iff {x y : UInt32} : x = y ↔ x.toNat = y.toNat := ⟨congrA @[simp] theorem UInt32.val_val_eq_toNat (x : UInt32) : x.val.val = x.toNat := rfl +@[simp] theorem UInt32.val_ofNat (n) : + (no_index (OfNat.ofNat n) : UInt32).val = OfNat.ofNat n := rfl + +@[simp] theorem UInt32.toNat_ofNat (n) : + (no_index (OfNat.ofNat n) : UInt32).toNat = n % UInt32.size := rfl + theorem UInt32.toNat_lt (x : UInt32) : x.toNat < 2 ^ 32 := x.val.isLt @[simp] theorem UInt32.toUInt8_toNat (x : UInt32) : x.toUInt8.toNat = x.toNat % 2 ^ 8 := rfl @@ -74,6 +122,21 @@ theorem UInt32.toNat_lt (x : UInt32) : x.toNat < 2 ^ 32 := x.val.isLt @[simp] theorem UInt32.toUInt64_toNat (x : UInt32) : x.toUInt64.toNat = x.toNat := rfl +theorem UInt32.toNat_zero : (0 : UInt32).toNat = 0 := rfl + +theorem UInt32.toNat_add (x y : UInt32) : (x + y).toNat = (x.toNat + y.toNat) % UInt32.size := rfl + +theorem UInt32.toNat_sub (x y : UInt32) : + (x - y).toNat = (x.toNat + (UInt32.size - y.toNat)) % UInt32.size := rfl + +theorem UInt32.toNat_mul (x y : UInt32) : (x * y).toNat = (x.toNat * y.toNat) % UInt32.size := rfl + +theorem UInt32.toNat_div (x y : UInt32) : (x / y).toNat = x.toNat / y.toNat := rfl + +theorem UInt32.toNat_mod (x y : UInt32) : (x % y).toNat = x.toNat % y.toNat := rfl + +theorem UInt32.toNat_modn (x : UInt32) (n) : (x.modn n).toNat = x.toNat % n := rfl + theorem UInt32.le_antisymm_iff {x y : UInt32} : x = y ↔ x ≤ y ∧ y ≤ x := UInt32.ext_iff.trans Nat.le_antisymm_iff @@ -92,6 +155,12 @@ theorem UInt64.ext_iff {x y : UInt64} : x = y ↔ x.toNat = y.toNat := ⟨congrA @[simp] theorem UInt64.val_val_eq_toNat (x : UInt64) : x.val.val = x.toNat := rfl +@[simp] theorem UInt64.val_ofNat (n) : + (no_index (OfNat.ofNat n) : UInt64).val = OfNat.ofNat n := rfl + +@[simp] theorem UInt64.toNat_ofNat (n) : + (no_index (OfNat.ofNat n) : UInt64).toNat = n % UInt64.size := rfl + theorem UInt64.toNat_lt (x : UInt64) : x.toNat < 2 ^ 64 := x.val.isLt @[simp] theorem UInt64.toUInt8_toNat (x : UInt64) : x.toUInt8.toNat = x.toNat % 2 ^ 8 := rfl @@ -100,6 +169,21 @@ theorem UInt64.toNat_lt (x : UInt64) : x.toNat < 2 ^ 64 := x.val.isLt @[simp] theorem UInt64.toUInt32_toNat (x : UInt64) : x.toUInt32.toNat = x.toNat % 2 ^ 32 := rfl +theorem UInt64.toNat_zero : (0 : UInt64).toNat = 0 := rfl + +theorem UInt64.toNat_add (x y : UInt64) : (x + y).toNat = (x.toNat + y.toNat) % UInt64.size := rfl + +theorem UInt64.toNat_sub (x y : UInt64) : + (x - y).toNat = (x.toNat + (UInt64.size - y.toNat)) % UInt64.size := rfl + +theorem UInt64.toNat_mul (x y : UInt64) : (x * y).toNat = (x.toNat * y.toNat) % UInt64.size := rfl + +theorem UInt64.toNat_div (x y : UInt64) : (x / y).toNat = x.toNat / y.toNat := rfl + +theorem UInt64.toNat_mod (x y : UInt64) : (x % y).toNat = x.toNat % y.toNat := rfl + +theorem UInt64.toNat_modn (x : UInt64) (n) : (x.modn n).toNat = x.toNat % n := rfl + theorem UInt64.le_antisymm_iff {x y : UInt64} : x = y ↔ x ≤ y ∧ y ≤ x := UInt64.ext_iff.trans Nat.le_antisymm_iff @@ -118,6 +202,12 @@ theorem USize.ext_iff {x y : USize} : x = y ↔ x.toNat = y.toNat := ⟨congrArg @[simp] theorem USize.val_val_eq_toNat (x : USize) : x.val.val = x.toNat := rfl +@[simp] theorem USize.val_ofNat (n) : + (no_index (OfNat.ofNat n) : USize).val = OfNat.ofNat n := rfl + +@[simp] theorem USize.toNat_ofNat (n) : + (no_index (OfNat.ofNat n) : USize).toNat = n % USize.size := rfl + theorem USize.size_eq : USize.size = 2 ^ System.Platform.numBits := by have : 1 ≤ 2 ^ System.Platform.numBits := Nat.succ_le_of_lt (Nat.two_pow_pos _) rw [USize.size, Nat.sub_add_cancel this] @@ -140,6 +230,21 @@ theorem USize.toNat_lt (x : USize) : x.toNat < 2 ^ System.Platform.numBits := by @[simp] theorem UInt32.toUSize_toNat (x : UInt32) : x.toUSize.toNat = x.toNat := rfl +theorem USize.toNat_zero : (0 : USize).toNat = 0 := rfl + +theorem USize.toNat_add (x y : USize) : (x + y).toNat = (x.toNat + y.toNat) % USize.size := rfl + +theorem USize.toNat_sub (x y : USize) : + (x - y).toNat = (x.toNat + (USize.size - y.toNat)) % USize.size := rfl + +theorem USize.toNat_mul (x y : USize) : (x * y).toNat = (x.toNat * y.toNat) % USize.size := rfl + +theorem USize.toNat_div (x y : USize) : (x / y).toNat = x.toNat / y.toNat := rfl + +theorem USize.toNat_mod (x y : USize) : (x % y).toNat = x.toNat % y.toNat := rfl + +theorem USize.toNat_modn (x : USize) (n) : (x.modn n).toNat = x.toNat % n := rfl + theorem USize.le_antisymm_iff {x y : USize} : x = y ↔ x ≤ y ∧ y ≤ x := USize.ext_iff.trans Nat.le_antisymm_iff From da2d1d464eb2e99db4f39b850e89cee4d5f72fc8 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 25 Jun 2024 07:59:01 +0100 Subject: [PATCH 013/177] chore: bump toolchain to v4.9.0-rc3 (#858) --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 29c0cea43e..e5ea66000f 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.9.0-rc2 +leanprover/lean4:v4.9.0-rc3 From 93343d4debf6d7f07ec467e1e60801072e81f0ab Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 25 Jun 2024 14:05:27 +0100 Subject: [PATCH 014/177] chore: replace `Std` with `Batteries` and recycle a doc-string (#859) --- Batteries/Linter/UnnecessarySeqFocus.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Batteries/Linter/UnnecessarySeqFocus.lean b/Batteries/Linter/UnnecessarySeqFocus.lean index 3ebe54dcb4..566da1d490 100644 --- a/Batteries/Linter/UnnecessarySeqFocus.lean +++ b/Batteries/Linter/UnnecessarySeqFocus.lean @@ -7,7 +7,7 @@ import Lean.Elab.Command import Lean.Linter.Util import Batteries.Lean.AttributeExtra -namespace Std.Linter +namespace Batteries.Linter open Lean Elab Command Linter /-- @@ -147,7 +147,7 @@ partial def markUsedTactics : InfoTree → M ω Unit end -/-- The main entry point to the unused tactic linter. -/ +@[inherit_doc Batteries.Linter.linter.unnecessarySeqFocus] def unnecessarySeqFocusLinter : Linter where run := withSetOptionIn fun stx => do unless getLinterUnnecessarySeqFocus (← getOptions) && (← getInfoState).enabled do return From 555ec79bc6effe7009036184a5f73f7d8d4850ed Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 26 Jun 2024 00:21:42 +0100 Subject: [PATCH 015/177] chore: `unreachableTactic` -- replace a `Std` namespace with `Batteries` and add an `inherit_doc`. (#861) --- Batteries/Linter/UnreachableTactic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Batteries/Linter/UnreachableTactic.lean b/Batteries/Linter/UnreachableTactic.lean index c4ea4dd10c..a4bedbe10c 100644 --- a/Batteries/Linter/UnreachableTactic.lean +++ b/Batteries/Linter/UnreachableTactic.lean @@ -7,7 +7,7 @@ import Lean.Elab.Command import Lean.Linter.Util import Batteries.Tactic.Unreachable -namespace Std.Linter +namespace Batteries.Linter open Lean Elab Command Linter /-- @@ -86,7 +86,7 @@ partial def eraseUsedTactics : InfoTree → M Unit end -/-- The main entry point to the unreachable tactic linter. -/ +@[inherit_doc Batteries.Linter.linter.unreachableTactic] def unreachableTacticLinter : Linter where run := withSetOptionIn fun stx => do unless getLinterUnreachableTactic (← getOptions) && (← getInfoState).enabled do return From 54bb04c3119f24fde14b9068c4b2e69db52a1450 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 1 Jul 2024 10:45:02 +1000 Subject: [PATCH 016/177] chore: bump toolchain to v4.9.0 (#865) --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index e5ea66000f..4ef27c472e 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.9.0-rc3 +leanprover/lean4:v4.9.0 From 4115b227caa6b6433c5ccf28c859a63607e5a29a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 1 Jul 2024 11:07:57 +1000 Subject: [PATCH 017/177] chore: deal with nonterminal simps in String/Lemmas (#857) * chore: deal with nonterminal simps in String/Lemmas * fix * squeeze more simps * chore: avoid nonterminal simps in String/Lemmas (#864) * chore: bump toolchain to v4.9.0-rc3 (#858) * chore: replace `Std` with `Batteries` and recycle a doc-string (#859) * chore: `unreachableTactic` -- replace a `Std` namespace with `Batteries` and add an `inherit_doc`. (#861) * chore: avoid nonterminal simps in String/Lemmas Co-authored-by: Mario Carneiro Co-authored-by: Bulhwi Cha --------- Co-authored-by: Kim Morrison Co-authored-by: damiano Co-authored-by: Mario Carneiro --------- Co-authored-by: Mario Carneiro Co-authored-by: Bulhwi Cha Co-authored-by: damiano --- Batteries/Data/String/Lemmas.lean | 203 ++++++++++++++++++------------ 1 file changed, 126 insertions(+), 77 deletions(-) diff --git a/Batteries/Data/String/Lemmas.lean b/Batteries/Data/String/Lemmas.lean index 8a9376fe05..ad15234f51 100644 --- a/Batteries/Data/String/Lemmas.lean +++ b/Batteries/Data/String/Lemmas.lean @@ -8,6 +8,7 @@ import Batteries.Data.List.Lemmas import Batteries.Data.String.Basic import Batteries.Tactic.Lint.Misc import Batteries.Tactic.SeqFocus +import Batteries.Tactic.SqueezeScope namespace String @@ -135,10 +136,10 @@ theorem utf8GetAux_add_right_cancel (s : List Char) (i p n : Nat) : utf8GetAux s ⟨i + n⟩ ⟨p + n⟩ = utf8GetAux s ⟨i⟩ ⟨p⟩ := by apply utf8InductionOn s ⟨i⟩ ⟨p⟩ (motive := fun s i => utf8GetAux s ⟨i.byteIdx + n⟩ ⟨p + n⟩ = utf8GetAux s i ⟨p⟩) <;> - simp [utf8GetAux] + simp only [utf8GetAux, Char.reduceDefault, implies_true, ↓reduceIte, ne_eq, pos_add_char] intro c cs ⟨i⟩ h ih - simp [Pos.ext_iff, Pos.addChar_eq] at h ⊢ - simp [Nat.add_right_cancel_iff, h] + simp only [Pos.ext_iff, Pos.addChar_eq] at h ⊢ + simp only [Nat.add_right_cancel_iff, h, ↓reduceIte] rw [Nat.add_right_comm] exact ih @@ -151,8 +152,9 @@ theorem utf8GetAux_of_valid (cs cs' : List Char) {i p : Nat} (hp : i + utf8Len c | [], [] => rfl | [], c::cs' => simp [← hp, utf8GetAux] | c::cs, cs' => - simp [utf8GetAux, -List.headD_eq_head?]; rw [if_neg] - case hnc => simp [← hp, Pos.ext_iff]; exact ne_self_add_add_csize + simp only [utf8GetAux, List.append_eq, Char.reduceDefault, ↓Char.isValue] + rw [if_neg] + case hnc => simp only [← hp, utf8Len_cons, Pos.ext_iff]; exact ne_self_add_add_csize refine utf8GetAux_of_valid cs cs' ?_ simpa [Nat.add_assoc, Nat.add_comm] using hp @@ -169,7 +171,8 @@ theorem utf8GetAux?_of_valid (cs cs' : List Char) {i p : Nat} (hp : i + utf8Len | [], [] => rfl | [], c::cs' => simp [← hp, utf8GetAux?] | c::cs, cs' => - simp [utf8GetAux?]; rw [if_neg] + simp only [utf8GetAux?, List.append_eq] + rw [if_neg] case hnc => simp [← hp, Pos.ext_iff]; exact ne_self_add_add_csize refine utf8GetAux?_of_valid cs cs' ?_ simpa [Nat.add_assoc, Nat.add_comm] using hp @@ -183,7 +186,8 @@ theorem utf8SetAux_of_valid (c' : Char) (cs cs' : List Char) {i p : Nat} (hp : i | [], [] => rfl | [], c::cs' => simp [← hp, utf8SetAux] | c::cs, cs' => - simp [utf8SetAux]; rw [if_neg] + simp only [utf8SetAux, List.append_eq, List.cons_append] + rw [if_neg] case hnc => simp [← hp, Pos.ext_iff]; exact ne_self_add_add_csize refine congrArg (c::·) (utf8SetAux_of_valid c' cs cs' ?_) simpa [Nat.add_assoc, Nat.add_comm] using hp @@ -219,15 +223,19 @@ theorem utf8PrevAux_of_valid {cs cs' : List Char} {c : Char} {i p : Nat} match cs with | [] => simp [utf8PrevAux, ← hp, Pos.addChar_eq] | c'::cs => - simp [utf8PrevAux, Pos.addChar_eq, ← hp]; rw [if_neg] + simp only [utf8PrevAux, Pos.addChar_eq, ← hp, utf8Len_cons, List.append_eq] + rw [if_neg] case hnc => - simp [Pos.ext_iff]; rw [Nat.add_right_comm, Nat.add_left_comm]; apply ne_add_csize_add_self + simp only [Pos.ext_iff] + rw [Nat.add_right_comm, Nat.add_left_comm] + apply ne_add_csize_add_self refine (utf8PrevAux_of_valid (by simp [Nat.add_assoc, Nat.add_left_comm])).trans ?_ simp [Nat.add_assoc, Nat.add_comm] theorem prev_of_valid (cs : List Char) (c : Char) (cs' : List Char) : prev ⟨cs ++ c :: cs'⟩ ⟨utf8Len cs + csize c⟩ = ⟨utf8Len cs⟩ := by - simp [prev]; refine (if_neg (Pos.ne_of_gt add_csize_pos)).trans ?_ + simp only [prev] + refine (if_neg (Pos.ne_of_gt add_csize_pos)).trans ?_ rw [utf8PrevAux_of_valid] <;> simp theorem prev_of_valid' (cs cs' : List Char) : @@ -270,11 +278,16 @@ theorem findAux_of_valid (p) : ∀ l m r, unfold findAux List.takeWhile rw [dif_pos (by exact Nat.lt_add_of_pos_right add_csize_pos)] have h1 := get_of_valid l (c::m++r); have h2 := next_of_valid l c (m++r) - simp at h1 h2; simp [h1, h2] - cases p c <;> simp - have foo := findAux_of_valid p (l++[c]) m r; simp at foo - rw [Nat.add_right_comm, Nat.add_assoc] at foo - rw [foo, Nat.add_right_comm, Nat.add_assoc] + simp only [List.cons_append, Char.reduceDefault, List.headD_cons] at h1 h2 + simp only [List.append_assoc, List.cons_append, h1, utf8Len_cons, h2] + cases p c + · simp only [Bool.false_eq_true, ↓reduceIte, Bool.not_false, utf8Len_cons] + have foo := findAux_of_valid p (l++[c]) m r + simp only [List.append_assoc, List.singleton_append, List.cons_append, utf8Len_append, + utf8Len_cons, utf8Len_nil, Nat.zero_add, List.nil_append] at foo + rw [Nat.add_right_comm, Nat.add_assoc] at foo + rw [foo, Nat.add_right_comm, Nat.add_assoc] + · simp theorem find_of_valid (p s) : find s p = ⟨utf8Len (s.1.takeWhile (!p ·))⟩ := by simpa using findAux_of_valid p [] s.1 [] @@ -287,8 +300,10 @@ theorem revFindAux_of_valid (p) : ∀ l r, unfold revFindAux List.dropWhile rw [dif_neg (by exact Pos.ne_of_gt add_csize_pos)] have h1 := get_of_valid l.reverse (c::r); have h2 := prev_of_valid l.reverse c r - simp at h1 h2; simp [h1, h2] - cases p c <;> simp + simp only [utf8Len_reverse, Char.reduceDefault, List.headD_cons] at h1 h2 + simp only [List.reverse_cons, List.append_assoc, List.singleton_append, utf8Len_cons, h2, h1] + cases p c <;> simp only [Bool.false_eq_true, ↓reduceIte, Bool.not_false, Bool.not_true, + List.tail?_cons, Option.map_some'] exact revFindAux_of_valid p l (c::r) theorem revFind_of_valid (p s) : @@ -308,12 +323,15 @@ theorem firstDiffPos_loop_eq (l₁ l₂ r₁ r₂ stop p) refine Nat.lt_min.2 ⟨?_, ?_⟩ <;> exact Nat.lt_add_of_pos_right add_csize_pos, show get ⟨l₁ ++ a :: r₁⟩ ⟨p⟩ = a by simp [hl₁, get_of_valid], show get ⟨l₂ ++ b :: r₂⟩ ⟨p⟩ = b by simp [hl₂, get_of_valid]] - simp; split <;> simp - subst b - rw [show next ⟨l₁ ++ a :: r₁⟩ ⟨p⟩ = ⟨utf8Len l₁ + csize a⟩ by simp [hl₁, next_of_valid]] - simpa [← hl₁, ← Nat.add_assoc, Nat.add_right_comm] using - firstDiffPos_loop_eq (l₁ ++ [a]) (l₂ ++ [a]) r₁ r₂ stop (p + csize a) - (by simp [hl₁]) (by simp [hl₂]) (by simp [hstop, ← Nat.add_assoc, Nat.add_right_comm]) + simp only [bne_iff_ne, ne_eq, ite_not, decide_eq_true_eq] + split + · simp only [utf8Len_cons] + subst b + rw [show next ⟨l₁ ++ a :: r₁⟩ ⟨p⟩ = ⟨utf8Len l₁ + csize a⟩ by simp [hl₁, next_of_valid]] + simpa [← hl₁, ← Nat.add_assoc, Nat.add_right_comm] using + firstDiffPos_loop_eq (l₁ ++ [a]) (l₂ ++ [a]) r₁ r₂ stop (p + csize a) + (by simp [hl₁]) (by simp [hl₂]) (by simp [hstop, ← Nat.add_assoc, Nat.add_right_comm]) + · simp · next h => rw [dif_neg] <;> simp [hstop, ← hl₁, ← hl₂, -Nat.not_lt, Nat.lt_min] intro h₁ h₂ @@ -330,10 +348,11 @@ theorem firstDiffPos_eq (a b : String) : theorem extract.go₂_add_right_cancel (s : List Char) (i e n : Nat) : go₂ s ⟨i + n⟩ ⟨e + n⟩ = go₂ s ⟨i⟩ ⟨e⟩ := by apply utf8InductionOn s ⟨i⟩ ⟨e⟩ (motive := fun s i => - go₂ s ⟨i.byteIdx + n⟩ ⟨e + n⟩ = go₂ s i ⟨e⟩) <;> simp [go₂] + go₂ s ⟨i.byteIdx + n⟩ ⟨e + n⟩ = go₂ s i ⟨e⟩) + <;> simp only [ne_eq, go₂, pos_add_char, implies_true, ↓reduceIte] intro c cs ⟨i⟩ h ih - simp [Pos.ext_iff, Pos.addChar_eq] at h ⊢ - simp [Nat.add_right_cancel_iff, h] + simp only [Pos.ext_iff, Pos.addChar_eq] at h ⊢ + simp only [Nat.add_right_cancel_iff, h, ↓reduceIte, List.cons.injEq, true_and] rw [Nat.add_right_comm] exact ih @@ -341,32 +360,35 @@ theorem extract.go₂_append_left : ∀ (s t : List Char) (i e : Nat), e = utf8Len s + i → go₂ (s ++ t) ⟨i⟩ ⟨e⟩ = s | [], t, i, _, rfl => by cases t <;> simp [go₂] | c :: cs, t, i, _, rfl => by - simp [go₂, Pos.ext_iff, ne_add_csize_add_self, Pos.addChar_eq] + simp only [go₂, utf8Len_cons, Pos.ext_iff, ne_add_csize_add_self, ↓reduceIte, List.append_eq, + Pos.addChar_eq, List.cons.injEq, true_and] apply go₂_append_left; rw [Nat.add_right_comm, Nat.add_assoc] theorem extract.go₁_add_right_cancel (s : List Char) (i b e n : Nat) : go₁ s ⟨i + n⟩ ⟨b + n⟩ ⟨e + n⟩ = go₁ s ⟨i⟩ ⟨b⟩ ⟨e⟩ := by apply utf8InductionOn s ⟨i⟩ ⟨b⟩ (motive := fun s i => - go₁ s ⟨i.byteIdx + n⟩ ⟨b + n⟩ ⟨e + n⟩ = go₁ s i ⟨b⟩ ⟨e⟩) <;> - simp [go₁] + go₁ s ⟨i.byteIdx + n⟩ ⟨b + n⟩ ⟨e + n⟩ = go₁ s i ⟨b⟩ ⟨e⟩) + <;> simp only [ne_eq, go₁, pos_add_char, implies_true, ↓reduceIte] · intro c cs apply go₂_add_right_cancel · intro c cs ⟨i⟩ h ih - simp [Pos.ext_iff, Pos.addChar_eq] at h ih ⊢ - simp [Nat.add_right_cancel_iff, h] + simp only [Pos.ext_iff, Pos.addChar_eq] at h ih ⊢ + simp only [Nat.add_right_cancel_iff, h, ↓reduceIte] rw [Nat.add_right_comm] exact ih theorem extract.go₁_cons_addChar (c : Char) (cs : List Char) (b e : Pos) : go₁ (c :: cs) 0 (b + c) (e + c) = go₁ cs 0 b e := by - simp [go₁, Pos.ext_iff, Nat.ne_of_lt add_csize_pos] + simp only [go₁, Pos.ext_iff, Pos.byteIdx_zero, pos_add_char, Nat.ne_of_lt add_csize_pos, + ↓reduceIte] apply go₁_add_right_cancel theorem extract.go₁_append_right : ∀ (s t : List Char) (i b : Nat) (e : Pos), b = utf8Len s + i → go₁ (s ++ t) ⟨i⟩ ⟨b⟩ e = go₂ t ⟨b⟩ e | [], t, i, _, e, rfl => by cases t <;> simp [go₁, go₂] | c :: cs, t, i, _, e, rfl => by - simp [go₁, Pos.ext_iff, ne_add_csize_add_self, Pos.addChar_eq] + simp only [go₁, utf8Len_cons, Pos.ext_iff, ne_add_csize_add_self, ↓reduceIte, List.append_eq, + Pos.addChar_eq] apply go₁_append_right; rw [Nat.add_right_comm, Nat.add_assoc] theorem extract.go₁_zero_utf8Len (s : List Char) : go₁ s 0 0 ⟨utf8Len s⟩ = s := @@ -375,13 +397,15 @@ theorem extract.go₁_zero_utf8Len (s : List Char) : go₁ s 0 0 ⟨utf8Len s⟩ theorem extract_cons_addChar (c : Char) (cs : List Char) (b e : Pos) : extract ⟨c :: cs⟩ (b + c) (e + c) = extract ⟨cs⟩ b e := by - simp [extract, Nat.add_le_add_iff_right] + simp only [extract, pos_add_char, ge_iff_le, Nat.add_le_add_iff_right] split <;> [rfl; rw [extract.go₁_cons_addChar]] theorem extract_zero_endPos : ∀ (s : String), s.extract 0 (endPos s) = s | ⟨[]⟩ => rfl | ⟨c :: cs⟩ => by - simp [extract, Nat.ne_of_gt add_csize_pos]; congr + simp only [extract, Pos.byteIdx_zero, endPos_eq, utf8Len_cons, ge_iff_le, Nat.le_zero_eq, + Nat.ne_of_gt add_csize_pos, ↓reduceIte] + congr apply extract.go₁_zero_utf8Len theorem extract_of_valid (l m r : List Char) : @@ -396,11 +420,17 @@ theorem splitAux_of_valid (p l m r acc) : splitAux ⟨l ++ m ++ r⟩ p ⟨utf8Len l⟩ ⟨utf8Len l + utf8Len m⟩ acc = acc.reverse ++ (List.splitOnP.go p r m.reverse).map mk := by unfold splitAux - simp [by simpa using atEnd_of_valid (l ++ m) r]; split + simp only [List.append_assoc, atEnd_iff, endPos_eq, utf8Len_append, Pos.mk_le_mk, by + simpa using atEnd_of_valid (l ++ m) r, List.reverse_cons, dite_eq_ite] + split · subst r; simpa [List.splitOnP.go] using extract_of_valid l m [] · obtain ⟨c, r, rfl⟩ := r.exists_cons_of_ne_nil ‹_› - simp [by simpa using (⟨get_of_valid (l++m) (c::r), next_of_valid (l++m) c r, - extract_of_valid l m (c::r)⟩ : _∧_∧_), List.splitOnP.go] + simp only [by + simpa using + (⟨get_of_valid (l ++ m) (c :: r), next_of_valid (l ++ m) c r, + extract_of_valid l m (c :: r)⟩ : + _ ∧ _ ∧ _), + List.splitOnP.go, List.reverse_reverse] split · simpa [Nat.add_assoc] using splitAux_of_valid p (l++m++[c]) [] r (⟨m⟩::acc) · simpa [Nat.add_assoc] using splitAux_of_valid p l (m++[c]) r acc @@ -478,7 +508,7 @@ theorem pos_eq_zero {l r it} (h : ValidFor l r it) : it.2 = 0 ↔ l = [] := by simp [h.pos, Pos.ext_iff] theorem pos_eq_endPos {l r it} (h : ValidFor l r it) : it.2 = it.1.endPos ↔ r = [] := by - simp [h.pos, h.toString, Pos.ext_iff] + simp only [h.pos, h.toString, endPos_eq, utf8Len_reverseAux, Pos.ext_iff] exact (Nat.add_left_cancel_iff (m := 0)).trans <| eq_comm.trans utf8Len_eq_zero theorem curr : ∀ {it}, ValidFor l r it → it.curr = r.headD default @@ -494,15 +524,21 @@ theorem prev : ∀ {it}, ValidFor (c :: l) r it → ValidFor l (c :: r) it.prev | it, h => by cases h.out' have := prev_of_valid l.reverse c r - simp at this; simp [Iterator.prev, this] + simp only [utf8Len_reverse] at this + simp only [Iterator.prev, List.reverse_cons, List.append_assoc, List.singleton_append, + utf8Len_append, utf8Len_reverse, utf8Len_cons, utf8Len_nil, Nat.zero_add, this] exact .of_eq _ (by simp [List.reverseAux_eq]) (by simp) theorem prev_nil : ∀ {it}, ValidFor [] r it → ValidFor [] r it.prev - | it, h => by simp [Iterator.prev, h.toString, h.pos]; constructor + | it, h => by + simp only [Iterator.prev, h.toString, List.reverseAux_nil, h.pos, utf8Len_nil, + Pos.mk_zero, prev_zero] + constructor theorem atEnd : ∀ {it}, ValidFor l r it → (it.atEnd ↔ r = []) | it, h => by - simp [Iterator.atEnd, h.pos, h.toString] + simp only [Iterator.atEnd, h.pos, h.toString, endPos_eq, utf8Len_reverseAux, ge_iff_le, + decide_eq_true_eq] exact Nat.add_le_add_iff_left.trans <| Nat.le_zero.trans utf8Len_eq_zero theorem hasNext : ∀ {it}, ValidFor l r it → (it.hasNext ↔ r ≠ []) @@ -515,26 +551,29 @@ theorem setCurr' : ∀ {it}, ValidFor l r it → ValidFor l (r.modifyHead fun _ => c) (it.setCurr c) | it, h => by cases h.out' - simp [Iterator.setCurr] + simp only [setCurr, utf8Len_reverse] refine .of_eq _ ?_ (by simp) have := set_of_valid l.reverse r c - simp at this; simp [List.reverseAux_eq, this] + simp only [utf8Len_reverse] at this; simp [List.reverseAux_eq, this] theorem setCurr (h : ValidFor l (c :: r) it) : ValidFor l (c :: r) (it.setCurr c) := h.setCurr' theorem toEnd (h : ValidFor l r it) : ValidFor (r.reverse ++ l) [] it.toEnd := by - simp [Iterator.toEnd, h.toString] + simp only [Iterator.toEnd, h.toString, endPos_eq, utf8Len_reverseAux] exact .of_eq _ (by simp [List.reverseAux_eq]) (by simp [Nat.add_comm]) theorem toEnd' (it : Iterator) : ValidFor it.s.1.reverse [] it.toEnd := by - simp [Iterator.toEnd] + simp only [Iterator.toEnd] exact .of_eq _ (by simp [List.reverseAux_eq]) (by simp [endPos, utf8ByteSize]) theorem extract (h₁ : ValidFor l (m ++ r) it₁) (h₂ : ValidFor (m.reverse ++ l) r it₂) : it₁.extract it₂ = ⟨m⟩ := by cases h₁.out; cases h₂.out - simp [Iterator.extract, List.reverseAux_eq, Nat.not_lt.2 (Nat.le_add_left ..)] + simp only [Iterator.extract, List.reverseAux_eq, List.reverse_append, List.reverse_reverse, + List.append_assoc, ne_eq, not_true_eq_false, decide_False, utf8Len_append, utf8Len_reverse, + gt_iff_lt, pos_lt_eq, Nat.not_lt.2 (Nat.le_add_left ..), Bool.or_self, Bool.false_eq_true, + ↓reduceIte] simpa [Nat.add_comm] using extract_of_valid l.reverse m r theorem remainingToString {it} (h : ValidFor l r it) : it.remainingToString = ⟨r⟩ := by @@ -545,7 +584,7 @@ theorem nextn : ∀ {it}, ValidFor l r it → ∀ n, n ≤ r.length → ValidFor ((r.take n).reverse ++ l) (r.drop n) (it.nextn n) | it, h, 0, _ => by simp [h, Iterator.nextn] | it, h, n+1, hn => by - simp [h, Iterator.nextn] + simp only [Iterator.nextn] have a::r := r simpa using h.next.nextn _ (Nat.le_of_succ_le_succ hn) @@ -553,7 +592,7 @@ theorem prevn : ∀ {it}, ValidFor l r it → ∀ n, n ≤ l.length → ValidFor (l.drop n) ((l.take n).reverse ++ r) (it.prevn n) | it, h, 0, _ => by simp [h, Iterator.prevn] | it, h, n+1, hn => by - simp [h, Iterator.prevn] + simp only [Iterator.prevn] have a::l := l simpa using h.prev.prevn _ (Nat.le_of_succ_le_succ hn) @@ -608,7 +647,8 @@ theorem offsetOfPosAux_of_valid : ∀ l m r n, unfold offsetOfPosAux rw [if_neg (by exact Nat.not_le.2 (Nat.lt_add_of_pos_right add_csize_pos))] simp only [List.append_assoc, atEnd_of_valid l (c::m++r)] - simp [next_of_valid l c (m++r)] + simp only [List.cons_append, ↓reduceDite, utf8Len_cons, next_of_valid l c (m ++ r), + List.length_cons, Nat.succ_eq_add_one] simpa [← Nat.add_assoc, Nat.add_right_comm, Nat.succ_eq_add_one] using offsetOfPosAux_of_valid (l++[c]) m r (n + 1) @@ -622,7 +662,8 @@ theorem foldlAux_of_valid (f : α → Char → α) : ∀ l m r a, | l, c::m, r, a => by unfold foldlAux rw [dif_pos (by exact Nat.lt_add_of_pos_right add_csize_pos)] - simp [get_of_valid l (c::(m++r)), next_of_valid l c (m++r)] + simp only [List.append_assoc, List.cons_append, utf8Len_cons, next_of_valid l c (m ++ r), + get_of_valid l (c :: (m ++ r)), Char.reduceDefault, List.headD_cons, List.foldl_cons] simpa [← Nat.add_assoc, Nat.add_right_comm] using foldlAux_of_valid f (l++[c]) m r (f a c) theorem foldl_eq (f : α → Char → α) (s a) : foldl f a s = s.1.foldl f a := by @@ -635,8 +676,8 @@ theorem foldrAux_of_valid (f : Char → α → α) (l m r a) : induction m.reverse generalizing r a with (unfold foldrAux; simp) | cons c m IH => rw [if_pos (by exact Nat.lt_add_of_pos_right add_csize_pos)] - simp [← Nat.add_assoc, by simpa using prev_of_valid (l++m.reverse) c r] - simp [by simpa using get_of_valid (l++m.reverse) (c::r)] + simp only [← Nat.add_assoc, by simpa using prev_of_valid (l ++ m.reverse) c r] + simp only [by simpa using get_of_valid (l ++ m.reverse) (c :: r)] simpa using IH (c::r) (f c a) theorem foldr_eq (f : Char → α → α) (s a) : foldr f a s = s.1.foldr f a := by @@ -649,7 +690,9 @@ theorem anyAux_of_valid (p : Char → Bool) : ∀ l m r, | l, c::m, r => by unfold anyAux rw [dif_pos (by exact Nat.lt_add_of_pos_right add_csize_pos)] - simp [get_of_valid l (c::(m++r)), next_of_valid l c (m++r)] + simp only [List.append_assoc, List.cons_append, get_of_valid l (c :: (m ++ r)), + Char.reduceDefault, List.headD_cons, utf8Len_cons, next_of_valid l c (m ++ r), + Bool.if_true_left, Bool.decide_eq_true, List.any_cons] cases p c <;> simp simpa [← Nat.add_assoc, Nat.add_right_comm] using anyAux_of_valid p (l++[c]) m r @@ -672,7 +715,8 @@ theorem mapAux_of_valid (f : Char → Char) : ∀ l r, mapAux f ⟨utf8Len l⟩ | l, c::r => by unfold mapAux rw [dif_neg (by rw [atEnd_of_valid]; simp)] - simp [set_of_valid l (c::r), get_of_valid l (c::r), next_of_valid l (f c) r] + simp only [get_of_valid l (c :: r), Char.reduceDefault, List.headD_cons, + set_of_valid l (c :: r), List.modifyHead_cons, next_of_valid l (f c) r, List.map_cons] simpa using mapAux_of_valid f (l++[f c]) r theorem map_eq (f : Char → Char) (s) : map f s = ⟨s.1.map f⟩ := by @@ -690,7 +734,8 @@ theorem takeWhileAux_of_valid (p : Char → Bool) : ∀ l m r, | l, c::m, r => by unfold Substring.takeWhileAux List.takeWhile rw [dif_pos (by exact Nat.lt_add_of_pos_right add_csize_pos)] - simp [get_of_valid l (c::(m++r)), next_of_valid l c (m++r)] + simp only [List.append_assoc, List.cons_append, get_of_valid l (c :: (m ++ r)), + Char.reduceDefault, List.headD_cons, utf8Len_cons, next_of_valid l c (m ++ r)] cases p c <;> simp simpa [← Nat.add_assoc, Nat.add_right_comm] using takeWhileAux_of_valid p (l++[c]) m r @@ -752,7 +797,7 @@ theorem toString : ∀ {s}, ValidFor l m r s → s.toString = ⟨m⟩ theorem toIterator : ∀ {s}, ValidFor l m r s → s.toIterator.ValidFor l.reverse (m ++ r) | _, h => by - simp [Substring.toIterator] + simp only [Substring.toIterator] exact .of_eq _ (by simp [h.str, List.reverseAux_eq]) (by simp [h.startPos]) theorem get : ∀ {s}, ValidFor l (m₁ ++ c :: m₂) r s → s.get ⟨utf8Len m₁⟩ = c @@ -760,13 +805,13 @@ theorem get : ∀ {s}, ValidFor l (m₁ ++ c :: m₂) r s → s.get ⟨utf8Len m theorem next : ∀ {s}, ValidFor l (m₁ ++ c :: m₂) r s → s.next ⟨utf8Len m₁⟩ = ⟨utf8Len m₁ + csize c⟩ | _, ⟨⟩ => by - simp [Substring.next] + simp only [Substring.next, utf8Len_append, utf8Len_cons, List.append_assoc, List.cons_append] rw [if_neg (mt Pos.ext_iff.1 ?a)] case a => simpa [Nat.add_assoc, Nat.add_comm, Nat.add_left_comm] using @ne_add_csize_add_self (utf8Len l + utf8Len m₁) (utf8Len m₂) c have := next_of_valid (l ++ m₁) c (m₂ ++ r) - simp [Pos.add_eq] at this ⊢; rw [this] + simp only [List.append_assoc, utf8Len_append, Pos.add_eq] at this ⊢; rw [this] simp [Nat.add_assoc, Nat.add_sub_cancel_left] theorem next_stop : ∀ {s}, ValidFor l m r s → s.next ⟨utf8Len m⟩ = ⟨utf8Len m⟩ @@ -774,11 +819,11 @@ theorem next_stop : ∀ {s}, ValidFor l m r s → s.next ⟨utf8Len m⟩ = ⟨ut theorem prev : ∀ {s}, ValidFor l (m₁ ++ c :: m₂) r s → s.prev ⟨utf8Len m₁ + csize c⟩ = ⟨utf8Len m₁⟩ | _, ⟨⟩ => by - simp [Substring.prev] + simp only [Substring.prev, List.append_assoc, List.cons_append] rw [if_neg (mt Pos.ext_iff.1 <| Ne.symm ?a)] case a => simpa [Nat.add_comm] using @ne_add_csize_add_self (utf8Len l) (utf8Len m₁) c have := prev_of_valid (l ++ m₁) c (m₂ ++ r) - simp [Pos.add_eq, Nat.add_assoc] at this ⊢; rw [this] + simp only [List.append_assoc, utf8Len_append, Nat.add_assoc, Pos.add_eq] at this ⊢; rw [this] simp [Nat.add_sub_cancel_left] theorem nextn_stop : ∀ {s}, ValidFor l m r s → ∀ n, s.nextn n ⟨utf8Len m⟩ = ⟨utf8Len m⟩ @@ -789,7 +834,7 @@ theorem nextn : ∀ {s}, ValidFor l (m₁ ++ m₂) r s → ∀ n, s.nextn n ⟨utf8Len m₁⟩ = ⟨utf8Len m₁ + utf8Len (m₂.take n)⟩ | _, _, 0 => by simp [Substring.nextn] | s, h, n+1 => by - simp [Substring.nextn] + simp only [Substring.nextn] match m₂ with | [] => simp at h; simp [h.next_stop, h.nextn_stop] | c::m₂ => @@ -801,7 +846,7 @@ theorem prevn : ∀ {s}, ValidFor l (m₁.reverse ++ m₂) r s → ∀ n, s.prevn n ⟨utf8Len m₁⟩ = ⟨utf8Len (m₁.drop n)⟩ | _, _, 0 => by simp [Substring.prevn] | s, h, n+1 => by - simp [Substring.prevn] + simp only [Substring.prevn] match m₁ with | [] => simp | c::m₁ => @@ -814,8 +859,9 @@ theorem front : ∀ {s}, ValidFor l (c :: m) r s → s.front = c theorem drop : ∀ {s}, ValidFor l m r s → ∀ n, ValidFor (l ++ m.take n) (m.drop n) r (s.drop n) | s, h, n => by have : Substring.nextn {..} .. = _ := h.nextn (m₁ := []) n - simp at this; simp [Substring.drop, this] - simp [h.str, h.startPos, h.stopPos] + simp only [utf8Len_nil, Pos.mk_zero, Nat.zero_add] at this + simp only [Substring.drop, this] + simp only [h.str, List.append_assoc, h.startPos, h.stopPos] rw [← List.take_append_drop n m] at h refine .of_eq _ (by simp) (by simp) ?_ conv => lhs; rw [← List.take_append_drop n m] @@ -824,8 +870,9 @@ theorem drop : ∀ {s}, ValidFor l m r s → ∀ n, ValidFor (l ++ m.take n) (m. theorem take : ∀ {s}, ValidFor l m r s → ∀ n, ValidFor l (m.take n) (m.drop n ++ r) (s.take n) | s, h, n => by have : Substring.nextn {..} .. = _ := h.nextn (m₁ := []) n - simp at this; simp [Substring.take, this] - simp [h.str, h.startPos, h.stopPos] + simp at this + simp only [Substring.take, this] + simp only [h.str, List.append_assoc, h.startPos] rw [← List.take_append_drop n m] at h refine .of_eq _ ?_ (by simp) (by simp) conv => lhs; rw [← List.take_append_drop n m] @@ -839,14 +886,16 @@ theorem atEnd : ∀ {s}, ValidFor l m r s → (s.atEnd ⟨p⟩ ↔ p = utf8Len m theorem extract : ∀ {s}, ValidFor l m r s → ValidFor ml mm mr ⟨⟨m⟩, b, e⟩ → ∃ l' r', ValidFor l' mm r' (s.extract b e) | _, ⟨⟩, ⟨⟩ => by - simp [Substring.extract]; split + simp only [Substring.extract, ge_iff_le, Pos.mk_le_mk, List.append_assoc, utf8Len_append] + split · next h => rw [utf8Len_eq_zero.1 <| Nat.le_zero.1 <| Nat.add_le_add_iff_left.1 h] exact ⟨[], [], ⟨⟩⟩ · next h => refine ⟨l ++ ml, mr ++ r, .of_eq _ (by simp) ?_ ?_⟩ <;> - simp [Nat.min_eq_min] <;> rw [Nat.min_eq_right] <;> - try simp [Nat.add_le_add_iff_left, Nat.le_add_right] + simp only [Pos.add_byteIdx, Nat.min_eq_min, utf8Len_append] + <;> rw [Nat.min_eq_right] + <;> try simp [Nat.add_le_add_iff_left, Nat.le_add_right] rw [Nat.add_assoc] -- TODO: splitOn @@ -888,11 +937,11 @@ namespace Valid theorem validFor : ∀ {s}, Valid s → ∃ l m r, ValidFor l m r s | ⟨⟨_⟩, ⟨_⟩, ⟨_⟩⟩, ⟨⟨l, mr, rfl, rfl⟩, ⟨lm, r, e, rfl⟩, h⟩ => by - simp at * + simp only [utf8ByteSize.go_eq, Pos.mk_le_mk] at * have := (or_iff_right_iff_imp.2 fun h => ?x).1 (List.append_eq_append_iff.1 e) case x => match l, r, h with | _, _, ⟨m, rfl, rfl⟩ => ?_ - simp at h + simp only [utf8Len_append] at h cases utf8Len_eq_zero.1 <| Nat.le_zero.1 (Nat.le_of_add_le_add_left (c := 0) h) exact ⟨[], by simp⟩ match lm, mr, this with @@ -913,13 +962,13 @@ theorem isEmpty : ∀ {s}, Valid s → (s.isEmpty ↔ s.toString = "") theorem get : ∀ {s}, Valid s → s.toString.1 = m₁ ++ c :: m₂ → s.get ⟨utf8Len m₁⟩ = c | _, h, e => by let ⟨l, m, r, h⟩ := h.validFor - simp [h.toString] at e; subst e; simp [h.get] + simp only [h.toString] at e; subst e; simp [h.get] theorem next : ∀ {s}, Valid s → s.toString.1 = m₁ ++ c :: m₂ → s.next ⟨utf8Len m₁⟩ = ⟨utf8Len m₁ + csize c⟩ | _, h, e => by let ⟨l, m, r, h⟩ := h.validFor - simp [h.toString] at e; subst e; simp [h.next] + simp only [h.toString] at e; subst e; simp [h.next] theorem next_stop : ∀ {s}, Valid s → s.next ⟨s.bsize⟩ = ⟨s.bsize⟩ | _, h => let ⟨l, m, r, h⟩ := h.validFor; by simp [h.bsize, h.next_stop] @@ -928,7 +977,7 @@ theorem prev : ∀ {s}, Valid s → s.toString.1 = m₁ ++ c :: m₂ → s.prev ⟨utf8Len m₁ + csize c⟩ = ⟨utf8Len m₁⟩ | _, h, e => by let ⟨l, m, r, h⟩ := h.validFor - simp [h.toString] at e; subst e; simp [h.prev] + simp only [h.toString] at e; subst e; simp [h.prev] theorem nextn_stop : ∀ {s}, Valid s → ∀ n, s.nextn n ⟨s.bsize⟩ = ⟨s.bsize⟩ | _, h, n => let ⟨l, m, r, h⟩ := h.validFor; by simp [h.bsize, h.nextn_stop] @@ -937,13 +986,13 @@ theorem nextn : ∀ {s}, Valid s → s.toString.1 = m₁ ++ m₂ → ∀ n, s.nextn n ⟨utf8Len m₁⟩ = ⟨utf8Len m₁ + utf8Len (m₂.take n)⟩ | _, h, e => by let ⟨l, m, r, h⟩ := h.validFor - simp [h.toString] at e; subst e; simp [h.nextn] + simp only [h.toString] at e; subst e; simp [h.nextn] theorem prevn : ∀ {s}, Valid s → s.toString.1 = m₁.reverse ++ m₂ → ∀ n, s.prevn n ⟨utf8Len m₁⟩ = ⟨utf8Len (m₁.drop n)⟩ | _, h, e => by let ⟨l, m, r, h⟩ := h.validFor - simp [h.toString] at e; subst e; simp [h.prevn] + simp only [h.toString] at e; subst e; simp [h.prevn] theorem front : ∀ {s}, Valid s → s.toString.1 = c :: m → s.front = c | _, h => h.get (m₁ := []) From 2ead90d24b4fac3a05c9c4294daa39bd8686fb98 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 1 Jul 2024 11:11:41 +1000 Subject: [PATCH 018/177] chore: bump toolchain to v4.10.0-rc1, merging bump/v4.10.0 (#867) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * chore: adaptations for nightly-2024-06-11 (#837) * chore: adaptation for nightly-2024-06-14 (#845) * fix(linter): "(invalid MessageData.lazy, missing context)" (#838) * chore(Tactic/Lint/Simp): make simpNF linter dsimp-aware (#839) * add test demonstrating undesired behaviour on dsimp lemmas * chore(Tactic/Lint/Simp): make simpNF linter dsimp-aware * test demonstrates correct behaviour on dsimp lemmas * process review comments * make formatLemmas dsimp-aware * Revert "chore(Tactic/Lint/Simp): make simpNF linter dsimp-aware (#839)" (#841) This reverts commit 007a82fab9671c960ad260d4e4bf664b9fb884e3. * feat: ext lemma for `Thunk` (#842) * chore: adaptations for nightly-2024-06-14 --------- Co-authored-by: L Co-authored-by: Johan Commelin Co-authored-by: François G. Dorais * chore: adaptations for nightly-2024-06-16 (#848) * chore: adaptations for nightly-2024-06-19 (#855) * chore: adaptations for nightly-2024-06-22 (#856) * chore: bump toolchain to v4.10.0-rc1 * chore: adaptations for nightly-2024-06-30 (#866) * chore: bump toolchain to v4.10.0-rc1 * chore: adaptations for nightly-2024-06-30 * fix --------- Co-authored-by: L Co-authored-by: Johan Commelin Co-authored-by: François G. Dorais --- Batteries.lean | 1 - Batteries/CodeAction/Misc.lean | 1 - Batteries/Data/Array/Lemmas.lean | 9 +- Batteries/Data/Char.lean | 14 +- Batteries/Data/Fin/Lemmas.lean | 11 +- Batteries/Data/HashMap/WF.lean | 36 ++--- Batteries/Data/List/Basic.lean | 7 +- Batteries/Data/List/EraseIdx.lean | 28 ++-- Batteries/Data/List/Lemmas.lean | 197 +++++++++++++++----------- Batteries/Data/List/Pairwise.lean | 22 ++- Batteries/Data/List/Perm.lean | 6 +- Batteries/Data/String/Lemmas.lean | 39 ++--- Batteries/Data/UInt.lean | 10 +- Batteries/Lean/AttributeExtra.lean | 2 +- Batteries/Lean/Name.lean | 28 ---- Batteries/Lean/PersistentHashMap.lean | 9 -- Batteries/Lean/PersistentHashSet.lean | 10 -- Batteries/Tactic/Alias.lean | 1 + Batteries/Tactic/Lint/Simp.lean | 4 +- Batteries/Tactic/NoMatch.lean | 1 + Batteries/Tactic/PrintPrefix.lean | 1 - Batteries/Tactic/SqueezeScope.lean | 3 +- lean-toolchain | 2 +- test/absurd.lean | 2 +- test/congr.lean | 2 +- 25 files changed, 229 insertions(+), 217 deletions(-) delete mode 100644 Batteries/Lean/Name.lean diff --git a/Batteries.lean b/Batteries.lean index 706e967883..2fac3bf121 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -59,7 +59,6 @@ import Batteries.Lean.Meta.SavedState import Batteries.Lean.Meta.Simp import Batteries.Lean.Meta.UnusedNames import Batteries.Lean.MonadBacktrack -import Batteries.Lean.Name import Batteries.Lean.NameMap import Batteries.Lean.NameMapAttribute import Batteries.Lean.PersistentHashMap diff --git a/Batteries/CodeAction/Misc.lean b/Batteries/CodeAction/Misc.lean index 1b5be55ce6..0bea77ff55 100644 --- a/Batteries/CodeAction/Misc.lean +++ b/Batteries/CodeAction/Misc.lean @@ -5,7 +5,6 @@ Authors: Mario Carneiro -/ import Lean.Elab.BuiltinTerm import Lean.Elab.BuiltinNotation -import Batteries.Lean.Name import Batteries.Lean.Position import Batteries.CodeAction.Attr import Lean.Meta.Tactic.TryThis diff --git a/Batteries/Data/Array/Lemmas.lean b/Batteries/Data/Array/Lemmas.lean index 5a16a355e1..cf069eb703 100644 --- a/Batteries/Data/Array/Lemmas.lean +++ b/Batteries/Data/Array/Lemmas.lean @@ -22,7 +22,7 @@ theorem forIn_eq_data_forIn [Monad m] have j_eq : j = size as - 1 - i := by simp [← ij, ← Nat.add_assoc] have : as.size - 1 - i < as.size := j_eq ▸ ij ▸ Nat.lt_succ_of_le (Nat.le_add_right ..) have : as[size as - 1 - i] :: as.data.drop (j + 1) = as.data.drop j := by - rw [j_eq]; exact List.get_cons_drop _ ⟨_, this⟩ + rw [j_eq]; exact List.getElem_cons_drop _ _ this simp only [← this, List.forIn_cons]; congr; funext x; congr; funext b rw [loop (i := i)]; rw [← ij, Nat.succ_add]; rfl conv => lhs; simp only [forIn, Array.forIn] @@ -64,12 +64,11 @@ theorem zipWith_eq_zipWith_data (f : α → β → γ) (as : Array α) (bs : Arr let i_bs : Fin bs.data.length := ⟨i, hbs⟩ rw [h₁, List.append_assoc] congr - rw [← List.zipWith_append (h := by simp), getElem_eq_data_get, getElem_eq_data_get] - show List.zipWith f ((List.get as.data i_as) :: List.drop (i_as + 1) as.data) + rw [← List.zipWith_append (h := by simp), getElem_eq_data_getElem, getElem_eq_data_getElem] + show List.zipWith f (as.data[i_as] :: List.drop (i_as + 1) as.data) ((List.get bs.data i_bs) :: List.drop (i_bs + 1) bs.data) = List.zipWith f (List.drop i as.data) (List.drop i bs.data) - simp only [List.get_cons_drop] - termination_by as.size - i + simp only [data_length, Fin.getElem_fin, List.getElem_cons_drop, List.get_eq_getElem] simp [zipWith, loop 0 #[] (by simp) (by simp)] theorem size_zipWith (as : Array α) (bs : Array β) (f : α → β → γ) : diff --git a/Batteries/Data/Char.lean b/Batteries/Data/Char.lean index c94b35ac44..5a9110ae12 100644 --- a/Batteries/Data/Char.lean +++ b/Batteries/Data/Char.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jannis Limperg -/ import Batteries.Data.UInt +import Batteries.Tactic.Alias @[ext] theorem Char.ext : {a b : Char} → a.val = b.val → a = b | ⟨_,_⟩, ⟨_,_⟩, rfl => rfl @@ -21,16 +22,7 @@ instance : Batteries.LawfulOrd Char := .compareOfLessAndEq namespace String -private theorem csize_eq (c) : - csize c = 1 ∨ csize c = 2 ∨ csize c = 3 ∨ - csize c = 4 := by - simp only [csize, Char.utf8Size] - repeat (first | split | (solve | simp (config := {decide := true}))) - -theorem csize_pos (c) : 0 < csize c := by - rcases csize_eq c with _|_|_|_ <;> simp_all (config := {decide := true}) - -theorem csize_le_4 (c) : csize c ≤ 4 := by - rcases csize_eq c with _|_|_|_ <;> simp_all (config := {decide := true}) +@[deprecated (since := "2024-06-11")] alias csize_pos := Char.utf8Size_pos +@[deprecated (since := "2024-06-11")] alias csize_le_4 := Char.utf8Size_le_four end String diff --git a/Batteries/Data/Fin/Lemmas.lean b/Batteries/Data/Fin/Lemmas.lean index a310aa6f93..4d4eb59a1a 100644 --- a/Batteries/Data/Fin/Lemmas.lean +++ b/Batteries/Data/Fin/Lemmas.lean @@ -30,8 +30,13 @@ protected theorem le_antisymm {x y : Fin n} (h1 : x ≤ y) (h2 : y ≤ x) : x = @[simp] theorem length_list (n) : (list n).length = n := by simp [list] -@[simp] theorem get_list (i : Fin (list n).length) : (list n).get i = i.cast (length_list n) := by - cases i; simp only [list]; rw [← Array.getElem_eq_data_get, getElem_enum, cast_mk] +@[simp] theorem getElem_list (i : Nat) (h : i < (list n).length) : + (list n)[i] = cast (length_list n) ⟨i, h⟩ := by + simp only [list]; rw [← Array.getElem_eq_data_getElem, getElem_enum, cast_mk] + +@[deprecated getElem_list (since := "2024-06-12")] +theorem get_list (i : Fin (list n).length) : (list n).get i = i.cast (length_list n) := by + simp [getElem_list] @[simp] theorem list_zero : list 0 = [] := by simp [list] @@ -52,7 +57,7 @@ theorem list_reverse (n) : (list n).reverse = (list n).map rev := by | succ n ih => conv => lhs; rw [list_succ_last] conv => rhs; rw [list_succ] - simp [List.reverse_map, ih, Function.comp_def, rev_succ] + simp [← List.map_reverse, ih, Function.comp_def, rev_succ] /-! ### foldl -/ diff --git a/Batteries/Data/HashMap/WF.lean b/Batteries/Data/HashMap/WF.lean index 57b121c83c..f6298d3322 100644 --- a/Batteries/Data/HashMap/WF.lean +++ b/Batteries/Data/HashMap/WF.lean @@ -23,7 +23,7 @@ theorem update_data (self : Buckets α β) (i d h) : theorem exists_of_update (self : Buckets α β) (i d h) : ∃ l₁ l₂, self.1.data = l₁ ++ self.1[i] :: l₂ ∧ List.length l₁ = i.toNat ∧ (self.update i d h).1.data = l₁ ++ d :: l₂ := by - simp only [Array.data_length, Array.ugetElem_eq_getElem, Array.getElem_eq_data_get] + simp only [Array.data_length, Array.ugetElem_eq_getElem, Array.getElem_eq_data_getElem] exact List.exists_of_set' h theorem update_update (self : Buckets α β) (i d d' h h') : @@ -37,13 +37,13 @@ theorem size_eq (data : Buckets α β) : theorem mk_size (h) : (mk n h : Buckets α β).size = 0 := by simp only [mk, mkArray, size_eq]; clear h - induction n <;> simp [*] + induction n <;> simp_all [List.replicate_succ] theorem WF.mk' [BEq α] [Hashable α] (h) : (Buckets.mk n h : Buckets α β).WF := by refine ⟨fun _ h => ?_, fun i h => ?_⟩ · simp only [Buckets.mk, mkArray, List.mem_replicate, ne_eq] at h simp [h, List.Pairwise.nil] - · simp [Buckets.mk, empty', mkArray, Array.getElem_eq_data_get, AssocList.All] + · simp [Buckets.mk, empty', mkArray, Array.getElem_eq_data_getElem, AssocList.All] theorem WF.update [BEq α] [Hashable α] {buckets : Buckets α β} {i d h} (H : buckets.WF) (h₁ : ∀ [PartialEquivBEq α] [LawfulHashable α], @@ -57,7 +57,8 @@ theorem WF.update [BEq α] [Hashable α] {buckets : Buckets α β} {i d h} (H : | .inl hl => H.1 _ hl | .inr rfl => h₁ (H.1 _ (Array.getElem_mem_data ..)) · revert hp - simp only [Array.getElem_eq_data_get, update_data, List.get_set, Array.data_length, update_size] + simp only [Array.getElem_eq_data_getElem, update_data, List.getElem_set, Array.data_length, + update_size] split <;> intro hp · next eq => exact eq ▸ h₂ (H.2 _ _) _ hp · simp only [update_size, Array.data_length] at hi @@ -87,23 +88,24 @@ theorem expand_size [Hashable α] {buckets : Buckets α β} : · rw [Buckets.mk_size]; simp [Buckets.size] · nofun where - go (i source) (target : Buckets α β) (hs : ∀ j < i, source.data.getD j .nil = .nil) : + go (i source) (target : Buckets α β) (hs : ∀ j < i, source.data[j]?.getD .nil = .nil) : (expand.go i source target).size = .sum (source.data.map (·.toList.length)) + target.size := by unfold expand.go; split · next H => refine (go (i+1) _ _ fun j hj => ?a).trans ?b <;> simp · case a => - simp only [List.getD_eq_get?, List.get?_set, Option.map_eq_map]; split - · cases List.get? .. <;> rfl + simp [List.getD_eq_getElem?, List.getElem?_set, Option.map_eq_map]; split + · cases source.data[j]? <;> rfl · next H => exact hs _ (Nat.lt_of_le_of_ne (Nat.le_of_lt_succ hj) (Ne.symm H)) · case b => refine have ⟨l₁, l₂, h₁, _, eq⟩ := List.exists_of_set' H; eq ▸ ?_ - simp only [Buckets.size_eq, h₁, List.map_append, List.map_cons, AssocList.toList, - List.length_nil, Nat.sum_append, Nat.sum_cons, Nat.zero_add, Array.data_length] + rw [h₁] + simp only [Buckets.size_eq, List.map_append, List.map_cons, AssocList.toList, + List.length_nil, Nat.sum_append, Nat.sum_cons, Nat.zero_add, Array.data_length] rw [Nat.add_assoc, Nat.add_assoc, Nat.add_assoc]; congr 1 (conv => rhs; rw [Nat.add_left_comm]); congr 1 - rw [← Array.getElem_eq_data_get] + rw [← Array.getElem_eq_data_getElem] have := @reinsertAux_size α β _; simp [Buckets.size] at this induction source[i].toList generalizing target <;> simp [*, Nat.succ_add]; rfl · next H => @@ -111,10 +113,10 @@ where rw [← (_ : source.data.map (fun _ => .nil) = source.data)] · simp only [List.map_map] induction source.data <;> simp [*] - refine List.ext_get (by simp) fun j h₁ h₂ => ?_ - simp only [List.get_map, Array.data_length] + refine List.ext_getElem (by simp) fun j h₁ h₂ => ?_ + simp only [List.getElem_map, Array.data_length] have := (hs j (Nat.lt_of_lt_of_le h₂ (Nat.not_lt.1 H))).symm - rwa [List.getD_eq_get?, List.get?_eq_get, Option.getD_some] at this + rwa [List.getElem?_eq_getElem] at this termination_by source.size - i theorem expand_WF.foldl [BEq α] [Hashable α] (rank : α → Nat) {l : List (α × β)} {i : Nat} @@ -166,8 +168,8 @@ where · match List.mem_or_eq_of_mem_set hl with | .inl hl => exact hs₁ _ hl | .inr e => exact e ▸ .nil - · simp only [Array.data_length, Array.size_set, Array.getElem_eq_data_get, Array.data_set, - List.get_set] + · simp only [Array.data_length, Array.size_set, Array.getElem_eq_data_getElem, Array.data_set, + List.getElem_set] split · nofun · exact hs₂ _ (by simp_all) @@ -364,8 +366,8 @@ theorem WF.filterMap {α β γ} {f : α → β → Option γ} [BEq α] [Hashable have := H.out.2.1 _ h rw [← List.pairwise_map (R := (¬ · == ·))] at this ⊢ exact this.sublist (H3 l.toList) - · simp only [Array.size_mk, List.length_map, Array.data_length, Array.getElem_eq_data_get, - List.get_map] at h ⊢ + · simp only [Array.size_mk, List.length_map, Array.data_length, Array.getElem_eq_data_getElem, + List.getElem_map] at h ⊢ have := H.out.2.2 _ h simp only [AssocList.All, List.toList_toAssocList, List.mem_reverse, List.mem_filterMap, Option.map_eq_some', forall_exists_index, and_imp, g₁] at this ⊢ diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index aacd1c9f2c..a8c3fef85a 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -313,7 +313,8 @@ def takeD : Nat → List α → α → List α @[simp] theorem takeD_succ (l) (a : α) : takeD (n+1) l a = l.head?.getD a :: takeD n l.tail a := by simp [takeD] -@[simp] theorem takeD_nil (n) (a : α) : takeD n [] a = replicate n a := by induction n <;> simp [*] +@[simp] theorem takeD_nil (n) (a : α) : takeD n [] a = replicate n a := by + induction n <;> simp [*, replicate_succ] /-- Tail-recursive version of `takeD`. -/ def takeDTR (n : Nat) (l : List α) (dflt : α) : List α := go n l #[] where @@ -325,7 +326,7 @@ def takeDTR (n : Nat) (l : List α) (dflt : α) : List α := go n l #[] where theorem takeDTR_go_eq : ∀ n l, takeDTR.go dflt n l acc = acc.data ++ takeD n l dflt | 0, _ => by simp [takeDTR.go] - | _+1, [] => by simp [takeDTR.go] + | _+1, [] => by simp [takeDTR.go, replicate_succ] | _+1, _::l => by simp [takeDTR.go, takeDTR_go_eq _ l] @[csimp] theorem takeD_eq_takeDTR : @takeD = @takeDTR := by @@ -485,7 +486,7 @@ def initsTR (l : List α) : List (List α) := l.foldr (fun a arrs => (arrs.map fun t => a :: t).push []) #[[]] |>.toListRev @[csimp] theorem inits_eq_initsTR : @inits = @initsTR := by - funext α l; simp [initsTR]; induction l <;> simp [*, reverse_map] + funext α l; simp [initsTR]; induction l <;> simp [*, map_reverse] /-- `tails l` is the list of terminal segments of `l`. diff --git a/Batteries/Data/List/EraseIdx.lean b/Batteries/Data/List/EraseIdx.lean index 31a595be57..ca3cf2e716 100644 --- a/Batteries/Data/List/EraseIdx.lean +++ b/Batteries/Data/List/EraseIdx.lean @@ -63,20 +63,32 @@ protected theorem IsPrefix.eraseIdx {l l' : List α} (h : l <+: l') (k : Nat) : rw [Nat.not_lt] at hkl simp [eraseIdx_append_of_length_le hkl, eraseIdx_of_length_le hkl] -theorem mem_eraseIdx_iff_get {x : α} : - ∀ {l} {k}, x ∈ eraseIdx l k ↔ ∃ i : Fin l.length, ↑i ≠ k ∧ l.get i = x +theorem mem_eraseIdx_iff_getElem {x : α} : + ∀ {l} {k}, x ∈ eraseIdx l k ↔ ∃ i : Fin l.length, ↑i ≠ k ∧ l[i.1] = x | [], _ => by simp only [eraseIdx, Fin.exists_iff, not_mem_nil, false_iff] rintro ⟨i, h, -⟩ exact Nat.not_lt_zero _ h | a::l, 0 => by simp [Fin.exists_fin_succ, mem_iff_get] | a::l, k+1 => by - simp [Fin.exists_fin_succ, mem_eraseIdx_iff_get, @eq_comm _ a, k.succ_ne_zero.symm] - + simp [Fin.exists_fin_succ, mem_eraseIdx_iff_getElem, @eq_comm _ a, k.succ_ne_zero.symm] + +@[deprecated mem_eraseIdx_iff_getElem (since := "2024-06-12")] +theorem mem_eraseIdx_iff_get {x : α} {l} {k} : + x ∈ eraseIdx l k ↔ ∃ i : Fin l.length, ↑i ≠ k ∧ l.get i = x := by + simp [mem_eraseIdx_iff_getElem] + +theorem mem_eraseIdx_iff_getElem? {x : α} {l} {k} : x ∈ eraseIdx l k ↔ ∃ i ≠ k, l[i]? = some x := by + simp only [mem_eraseIdx_iff_getElem, getElem_eq_iff, Fin.exists_iff, exists_and_left] + refine exists_congr fun i => and_congr_right' ?_ + constructor + · rintro ⟨_, h⟩; exact h + · rintro h; + obtain ⟨h', -⟩ := getElem?_eq_some.1 h + exact ⟨h', h⟩ + +@[deprecated mem_eraseIdx_iff_getElem? (since := "2024-06-12")] theorem mem_eraseIdx_iff_get? {x : α} {l} {k} : x ∈ eraseIdx l k ↔ ∃ i ≠ k, l.get? i = x := by - simp only [mem_eraseIdx_iff_get, Fin.exists_iff, exists_and_left, get_eq_iff, exists_prop] - refine exists_congr fun i => and_congr_right' <| and_iff_right_of_imp fun h => ?_ - obtain ⟨h, -⟩ := get?_eq_some.1 h - exact h + simp [mem_eraseIdx_iff_getElem?] end List diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index cd010a8150..44d9a5c18d 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -243,25 +243,41 @@ theorem tail_eq_tail? (l) : @tail α l = (tail? l).getD [] := by simp [tail_eq_t /-! ### get? -/ -theorem get_eq_iff : List.get l n = x ↔ l.get? n.1 = some x := by simp [get?_eq_some] +theorem getElem_eq_iff {l : List α} {n : Nat} {h : n < l.length} : l[n] = x ↔ l[n]? = some x := by + simp only [get_eq_getElem, get?_eq_getElem?, getElem?_eq_some] + exact ⟨fun w => ⟨h, w⟩, fun h => h.2⟩ -theorem get?_inj - (h₀ : i < xs.length) (h₁ : Nodup xs) (h₂ : xs.get? i = xs.get? j) : i = j := by +@[deprecated getElem_eq_iff (since := "2024-06-12")] +theorem get_eq_iff : List.get l n = x ↔ l.get? n.1 = some x := by + simp + +theorem getElem?_inj + (h₀ : i < xs.length) (h₁ : Nodup xs) (h₂ : xs[i]? = xs[j]?) : i = j := by induction xs generalizing i j with | nil => cases h₀ | cons x xs ih => match i, j with | 0, 0 => rfl | i+1, j+1 => simp; cases h₁ with - | cons ha h₁ => exact ih (Nat.lt_of_succ_lt_succ h₀) h₁ h₂ - | i+1, 0 => ?_ | 0, j+1 => ?_ + | cons ha h₁ => + simp only [getElem?_cons_succ] at h₂ + exact ih (Nat.lt_of_succ_lt_succ h₀) h₁ h₂ + | i+1, 0 => ?_ + | 0, j+1 => ?_ all_goals - simp at h₂ + simp only [get?_eq_getElem?, getElem?_cons_zero, getElem?_cons_succ] at h₂ cases h₁; rename_i h' h have := h x ?_ rfl; cases this rw [mem_iff_get?] + simp only [get?_eq_getElem?] exact ⟨_, h₂⟩; exact ⟨_ , h₂.symm⟩ +@[deprecated getElem?_inj (since := "2024-06-12")] +theorem get?_inj + (h₀ : i < xs.length) (h₁ : Nodup xs) (h₂ : xs.get? i = xs.get? j) : i = j := by + apply getElem?_inj h₀ h₁ + simp_all + /-! ### drop -/ theorem tail_drop (l : List α) (n : Nat) : (l.drop n).tail = l.drop (n + 1) := by @@ -294,15 +310,21 @@ theorem eraseIdx_eq_modifyNthTail : ∀ n (l : List α), eraseIdx l n = modifyNt @[deprecated] alias removeNth_eq_nth_tail := eraseIdx_eq_modifyNthTail -theorem get?_modifyNth (f : α → α) : - ∀ n (l : List α) m, (modifyNth f n l).get? m = (fun a => if n = m then f a else a) <$> l.get? m - | n, l, 0 => by cases l <;> cases n <;> rfl +theorem getElem?_modifyNth (f : α → α) : + ∀ n (l : List α) m, (modifyNth f n l)[m]? = (fun a => if n = m then f a else a) <$> l[m]? + | n, l, 0 => by cases l <;> cases n <;> simp | n, [], _+1 => by cases n <;> rfl - | 0, _ :: l, m+1 => by cases h : l.get? m <;> simp [h, modifyNth, m.succ_ne_zero.symm] - | n+1, a :: l, m+1 => - (get?_modifyNth f n l m).trans <| by - cases h' : l.get? m <;> by_cases h : n = m <;> - simp [h, if_pos, if_neg, Option.map, mt Nat.succ.inj, not_false_iff, h'] + | 0, _ :: l, m+1 => by cases h : l[m]? <;> simp [h, modifyNth, m.succ_ne_zero.symm] + | n+1, a :: l, m+1 => by + simp only [modifyNth_succ_cons, getElem?_cons_succ, Nat.reduceEqDiff, Option.map_eq_map] + refine (getElem?_modifyNth f n l m).trans ?_ + cases h' : l[m]? <;> by_cases h : n = m <;> + simp [h, if_pos, if_neg, Option.map, mt Nat.succ.inj, not_false_iff, h'] + +@[deprecated getElem?_modifyNth (since := "2024-06-12")] +theorem get?_modifyNth (f : α → α) (n) (l : List α) (m) : + (modifyNth f n l).get? m = (fun a => if n = m then f a else a) <$> l.get? m := by + simp [getElem?_modifyNth] theorem modifyNthTail_length (f : List α → List α) (H : ∀ l, length (f l) = length l) : ∀ n l, length (modifyNthTail f n l) = length l @@ -323,13 +345,23 @@ theorem exists_of_modifyNthTail (f : List α → List α) {n} {l : List α} (h : @[simp] theorem modify_get?_length (f : α → α) : ∀ n l, length (modifyNth f n l) = length l := modifyNthTail_length _ fun l => by cases l <;> rfl -@[simp] theorem get?_modifyNth_eq (f : α → α) (n) (l : List α) : - (modifyNth f n l).get? n = f <$> l.get? n := by - simp only [get?_modifyNth, if_pos] +@[simp] theorem getElem?_modifyNth_eq (f : α → α) (n) (l : List α) : + (modifyNth f n l)[n]? = f <$> l[n]? := by + simp only [getElem?_modifyNth, if_pos] + +@[deprecated getElem?_modifyNth_eq (since := "2024-06-12")] +theorem get?_modifyNth_eq (f : α → α) (n) (l : List α) : + (modifyNth f n l).get? n = f <$> l.get? n := by + simp [getElem?_modifyNth_eq] -@[simp] theorem get?_modifyNth_ne (f : α → α) {m n} (l : List α) (h : m ≠ n) : +@[simp] theorem getElem?_modifyNth_ne (f : α → α) {m n} (l : List α) (h : m ≠ n) : + (modifyNth f m l)[n]? = l[n]? := by + simp only [getElem?_modifyNth, if_neg h, id_map'] + +@[deprecated getElem?_modifyNth_ne (since := "2024-06-12")] +theorem get?_modifyNth_ne (f : α → α) {m n} (l : List α) (h : m ≠ n) : (modifyNth f m l).get? n = l.get? n := by - simp only [get?_modifyNth, if_neg h, id_map'] + simp [h] theorem exists_of_modifyNth (f : α → α) {n} {l : List α} (h : n < l.length) : ∃ l₁ a l₂, l = l₁ ++ a :: l₂ ∧ l₁.length = n ∧ modifyNth f n l = l₁ ++ f a :: l₂ := @@ -348,8 +380,8 @@ theorem modifyNth_eq_take_drop (f : α → α) : modifyNthTail_eq_take_drop _ rfl theorem modifyNth_eq_take_cons_drop (f : α → α) {n l} (h) : - modifyNth f n l = take n l ++ f (get l ⟨n, h⟩) :: drop (n + 1) l := by - rw [modifyNth_eq_take_drop, drop_eq_get_cons h]; rfl + modifyNth f n l = take n l ++ f l[n] :: drop (n + 1) l := by + rw [modifyNth_eq_take_drop, drop_eq_getElem_cons h]; rfl /-! ### set -/ @@ -367,7 +399,7 @@ theorem modifyNth_eq_set_get? (f : α → α) : | 0, l => by cases l <;> rfl | n+1, [] => rfl | n+1, b :: l => - (congrArg (cons _) (modifyNth_eq_set_get? ..)).trans <| by cases h : l.get? n <;> simp [h] + (congrArg (cons _) (modifyNth_eq_set_get? ..)).trans <| by cases h : l[n]? <;> simp [h] theorem modifyNth_eq_set_get (f : α → α) {n} {l : List α} (h) : l.modifyNth f n = l.set n (f (l.get ⟨n, h⟩)) := by @@ -378,42 +410,56 @@ theorem exists_of_set {l : List α} (h : n < l.length) : rw [set_eq_modifyNth]; exact exists_of_modifyNth _ h theorem exists_of_set' {l : List α} (h : n < l.length) : - ∃ l₁ l₂, l = l₁ ++ l.get ⟨n, h⟩ :: l₂ ∧ l₁.length = n ∧ l.set n a' = l₁ ++ a' :: l₂ := - have ⟨_, _, _, h₁, h₂, h₃⟩ := exists_of_set h; ⟨_, _, get_of_append h₁ h₂ ▸ h₁, h₂, h₃⟩ + ∃ l₁ l₂, l = l₁ ++ l[n] :: l₂ ∧ l₁.length = n ∧ l.set n a' = l₁ ++ a' :: l₂ := + have ⟨_, _, _, h₁, h₂, h₃⟩ := exists_of_set h; ⟨_, _, getElem_of_append h₁ h₂ ▸ h₁, h₂, h₃⟩ @[simp] +theorem getElem?_set_eq' (a : α) (n) (l : List α) : (set l n a)[n]? = (fun _ => a) <$> l[n]? := by + simp only [set_eq_modifyNth, getElem?_modifyNth_eq] + +@[deprecated getElem?_set_eq' (since := "2024-06-12")] theorem get?_set_eq (a : α) (n) (l : List α) : (set l n a).get? n = (fun _ => a) <$> l.get? n := by - simp only [set_eq_modifyNth, get?_modifyNth_eq] + simp +theorem getElem?_set_eq_of_lt (a : α) {n} {l : List α} (h : n < length l) : + (set l n a)[n]? = some a := by rw [getElem?_set_eq', getElem?_eq_getElem h]; rfl + +@[deprecated getElem?_set_eq_of_lt (since := "2024-06-12")] theorem get?_set_eq_of_lt (a : α) {n} {l : List α} (h : n < length l) : - (set l n a).get? n = some a := by rw [get?_set_eq, get?_eq_get h]; rfl + (set l n a).get? n = some a := by + rw [get?_eq_getElem?, getElem?_set_eq', getElem?_eq_getElem h]; rfl -@[simp] +@[deprecated getElem?_set_ne (since := "2024-06-12")] theorem get?_set_ne (a : α) {m n} (l : List α) (h : m ≠ n) : (set l m a).get? n = l.get? n := by - simp only [set_eq_modifyNth, get?_modifyNth_ne _ _ h] + simp [h] +theorem getElem?_set' (a : α) {m n} (l : List α) : + (set l m a)[n]? = if m = n then (fun _ => a) <$> l[n]? else l[n]? := by + by_cases m = n <;> simp [*] + +@[deprecated getElem?_set (since := "2024-06-12")] theorem get?_set (a : α) {m n} (l : List α) : (set l m a).get? n = if m = n then (fun _ => a) <$> l.get? n else l.get? n := by - by_cases m = n <;> simp [*, get?_set_eq, get?_set_ne] + simp [getElem?_set'] theorem get?_set_of_lt (a : α) {m n} (l : List α) (h : n < length l) : (set l m a).get? n = if m = n then some a else l.get? n := by - simp [get?_set, get?_eq_get h] + simp [getElem?_set', getElem?_eq_getElem h] theorem get?_set_of_lt' (a : α) {m n} (l : List α) (h : m < length l) : (set l m a).get? n = if m = n then some a else l.get? n := by - simp [get?_set]; split <;> subst_vars <;> simp [*, get?_eq_get h] + simp [getElem?_set]; split <;> subst_vars <;> simp [*, getElem?_eq_getElem h] theorem drop_set_of_lt (a : α) {n m : Nat} (l : List α) (h : n < m) : (l.set n a).drop m = l.drop m := - List.ext fun i => by rw [get?_drop, get?_drop, get?_set_ne _ _ (by omega)] + List.ext_getElem? fun i => by rw [getElem?_drop, getElem?_drop, getElem?_set_ne (by omega)] theorem take_set_of_lt (a : α) {n m : Nat} (l : List α) (h : m < n) : (l.set n a).take m = l.take m := - List.ext fun i => by - rw [get?_take_eq_if, get?_take_eq_if] + List.ext_getElem? fun i => by + rw [getElem?_take_eq_if, getElem?_take_eq_if] split - · next h' => rw [get?_set_ne _ _ (by omega)] + · next h' => rw [getElem?_set_ne (by omega)] · rfl /-! ### removeNth -/ @@ -607,32 +653,13 @@ end erase /-! ### filterMap -/ -theorem length_filter_le (p : α → Bool) (l : List α) : - (l.filter p).length ≤ l.length := (filter_sublist _).length_le - -theorem length_filterMap_le (f : α → Option β) (l : List α) : - (filterMap f l).length ≤ l.length := by - rw [← length_map _ some, map_filterMap_some_eq_filter_map_is_some, ← length_map _ f] - apply length_filter_le - protected theorem Sublist.filterMap (f : α → Option β) (s : l₁ <+ l₂) : filterMap f l₁ <+ filterMap f l₂ := by - induction s <;> simp <;> split <;> simp [*, cons, cons₂] + induction s <;> simp [filterMap_cons] <;> split <;> simp [*, cons, cons₂] theorem Sublist.filter (p : α → Bool) {l₁ l₂} (s : l₁ <+ l₂) : filter p l₁ <+ filter p l₂ := by rw [← filterMap_eq_filter]; apply s.filterMap -@[simp] -theorem filter_eq_self {l} : filter p l = l ↔ ∀ a ∈ l, p a := by - induction l with simp - | cons a l ih => - cases h : p a <;> simp [*] - intro h; exact Nat.lt_irrefl _ (h ▸ length_filter_le p l) - -@[simp] -theorem filter_length_eq_length {l} : (filter p l).length = l.length ↔ ∀ a ∈ l, p a := - Iff.trans ⟨l.filter_sublist.eq_of_length, congrArg length⟩ filter_eq_self - /-! ### findIdx -/ @[simp] theorem findIdx_nil {α : Type _} (p : α → Bool) : [].findIdx p = 0 := rfl @@ -697,7 +724,7 @@ theorem findIdx?_eq_some_iff (xs : List α) (p : α → Bool) : | nil => simp | cons x xs ih => simp only [findIdx?_cons, Nat.zero_add, findIdx?_succ, take_succ_cons, map_cons] - split <;> cases i <;> simp_all + split <;> cases i <;> simp_all [replicate_succ] theorem findIdx?_of_eq_some {xs : List α} {p : α → Bool} (w : xs.findIdx? p = some i) : match xs.get? i with | some a => p a | none => false := by @@ -1269,6 +1296,9 @@ protected theorem Pairwise.chain (p : Pairwise R (a :: l)) : Chain R a l := by /-! ### range', range -/ +theorem range'_succ (s n step) : range' s (n + 1) step = s :: range' (s + step) n step := by + simp [range', Nat.add_succ, Nat.mul_succ] + @[simp] theorem length_range' (s step) : ∀ n : Nat, length (range' s n step) = n | 0 => rfl | _ + 1 => congrArg succ (length_range' _ _ _) @@ -1331,15 +1361,27 @@ theorem range'_subset_right {s m n : Nat} (step0 : 0 < step) : theorem range'_subset_right_1 {s m n : Nat} : range' s m ⊆ range' s n ↔ m ≤ n := range'_subset_right (by decide) -theorem get?_range' (s step) : ∀ {m n : Nat}, m < n → get? (range' s n step) m = some (s + step * m) - | 0, n + 1, _ => rfl - | m + 1, n + 1, h => - (get?_range' (s + step) step (Nat.lt_of_add_lt_add_right h)).trans <| by +theorem getElem?_range' (s step) : + ∀ {m n : Nat}, m < n → (range' s n step)[m]? = some (s + step * m) + | 0, n + 1, _ => by simp [range'_succ] + | m + 1, n + 1, h => by + simp only [range'_succ, getElem?_cons_succ] + exact (getElem?_range' (s + step) step (Nat.lt_of_add_lt_add_right h)).trans <| by simp [Nat.mul_succ, Nat.add_assoc, Nat.add_comm] -@[simp] theorem get_range' {n m step} (i) (H : i < (range' n m step).length) : - get (range' n m step) ⟨i, H⟩ = n + step * i := - (get?_eq_some.1 <| get?_range' n step (by simpa using H)).2 +@[simp] theorem getElem_range' {n m step} (i) (H : i < (range' n m step).length) : + (range' n m step)[i] = n + step * i := + (getElem?_eq_some.1 <| getElem?_range' n step (by simpa using H)).2 + +@[deprecated getElem?_range' (since := "2024-06-12")] +theorem get?_range' (s step) {m n : Nat} (h : m < n) : + get? (range' s n step) m = some (s + step * m) := by + simp [h] + +@[deprecated getElem_range' (since := "2024-06-12")] +theorem get_range' {n m step} (i) (H : i < (range' n m step).length) : + get (range' n m step) ⟨i, H⟩ = n + step * i := by + simp theorem range'_concat (s n : Nat) : range' s (n + 1) step = range' s n step ++ [s + step * n] := by rw [Nat.add_comm n 1]; exact (range'_append s n 1 step).symm @@ -1383,14 +1425,23 @@ theorem not_mem_range_self {n : Nat} : n ∉ range n := by simp theorem self_mem_range_succ (n : Nat) : n ∈ range (n + 1) := by simp +theorem getElem?_range {m n : Nat} (h : m < n) : (range n)[m]? = some m := by + simp [range_eq_range', getElem?_range' _ _ h] + +@[simp] theorem getElem_range {n : Nat} (m) (h : m < (range n).length) : (range n)[m] = m := by + simp [range_eq_range'] + +@[deprecated getElem?_range (since := "2024-06-12")] theorem get?_range {m n : Nat} (h : m < n) : get? (range n) m = some m := by - simp [range_eq_range', get?_range' _ _ h] + simp [getElem?_range, h] + +@[deprecated getElem_range (since := "2024-06-12")] +theorem get_range {n} (i) (H : i < (range n).length) : get (range n) ⟨i, H⟩ = i := by + simp theorem range_succ (n : Nat) : range (succ n) = range n ++ [n] := by simp only [range_eq_range', range'_1_concat, Nat.zero_add] -@[simp] theorem range_zero : range 0 = [] := rfl - theorem range_add (a b : Nat) : range (a + b) = range a ++ (range b).map (a + ·) := by rw [← range'_eq_map_range] simpa [range_eq_range', Nat.add_comm] using (range'_append_1 0 a b).symm @@ -1410,10 +1461,8 @@ theorem reverse_range' : ∀ s n : Nat, reverse (range' s n) = map (s + n - 1 - | s, n + 1 => by rw [range'_1_concat, reverse_append, range_succ_eq_map, show s + (n + 1) - 1 = s + n from rfl, map, map_map] - simp [reverse_range', Nat.sub_right_comm]; rfl + simp [reverse_range', Nat.sub_right_comm, Nat.sub_sub] -@[simp] theorem get_range {n} (i) (H : i < (range n).length) : get (range n) ⟨i, H⟩ = i := - Option.some.inj <| by rw [← get?_eq_get _, get?_range (by simpa using H)] /-! ### enum, enumFrom -/ @@ -1425,16 +1474,6 @@ theorem reverse_range' : ∀ s n : Nat, reverse (range' s n) = map (s + n - 1 - @[simp] theorem enum_map_fst (l : List α) : map Prod.fst (enum l) = range l.length := by simp only [enum, enumFrom_map_fst, range_eq_range'] -/-! ### maximum? -/ - --- A specialization of `maximum?_eq_some_iff` to Nat. -theorem maximum?_eq_some_iff' {xs : List Nat} : - xs.maximum? = some a ↔ (a ∈ xs ∧ ∀ b ∈ xs, b ≤ a) := - maximum?_eq_some_iff - (le_refl := Nat.le_refl) - (max_eq_or := fun _ _ => Nat.max_def .. ▸ by split <;> simp) - (max_le_iff := fun _ _ _ => Nat.max_le) - /-! ### indexOf and indexesOf -/ theorem foldrIdx_start : diff --git a/Batteries/Data/List/Pairwise.lean b/Batteries/Data/List/Pairwise.lean index 392f261bf2..8583a41f2b 100644 --- a/Batteries/Data/List/Pairwise.lean +++ b/Batteries/Data/List/Pairwise.lean @@ -213,7 +213,7 @@ theorem map_get_sublist {l : List α} {is : List (Fin l.length)} (h : is.Pairwis simp; cases hl' have := IH h.of_cons (hd+1) _ rfl (pairwise_cons.mp h).1 specialize his hd (.head _) - have := (drop_eq_get_cons ..).symm ▸ this.cons₂ (get l hd) + have := (drop_eq_getElem_cons ..).symm ▸ this.cons₂ (get l hd) have := Sublist.append (nil_sublist (take hd l |>.drop n)) this rwa [nil_append, ← (drop_append_of_le_length ?_), take_append_drop] at this simp [Nat.min_eq_left (Nat.le_of_lt hd.isLt), his] @@ -232,25 +232,33 @@ theorem sublist_eq_map_get (h : l' <+ l) : ∃ is : List (Fin l.length), | cons₂ _ _ IH => rcases IH with ⟨is,IH⟩ refine ⟨⟨0, by simp [Nat.zero_lt_succ]⟩ :: is.map (·.succ), ?_⟩ - simp [comp_def, pairwise_map, IH] + simp [comp_def, pairwise_map, IH, ← get_eq_getElem] -theorem pairwise_iff_get : Pairwise R l ↔ ∀ (i j) (_hij : i < j), R (get l i) (get l j) := by +theorem pairwise_iff_getElem : Pairwise R l ↔ + ∀ (i j : Nat) (_hi : i < l.length) (_hj : j < l.length) (_hij : i < j), R l[i] l[j] := by rw [pairwise_iff_forall_sublist] constructor <;> intro h - · intros i j h' + · intros i j hi hj h' apply h - apply map_get_sublist (is := [i, j]) - rw [Fin.lt_def] at h'; simp [h'] + simpa [h'] using map_get_sublist (is := [⟨i, hi⟩, ⟨j, hj⟩]) · intros a b h' have ⟨is, h', hij⟩ := sublist_eq_map_get h' rcases is with ⟨⟩ | ⟨a', ⟨⟩ | ⟨b', ⟨⟩⟩⟩ <;> simp at h' rcases h' with ⟨rfl, rfl⟩ apply h; simpa using hij +theorem pairwise_iff_get : Pairwise R l ↔ ∀ (i j) (_hij : i < j), R (get l i) (get l j) := by + rw [pairwise_iff_getElem] + constructor <;> intro h + · intros i j h' + exact h _ _ _ _ h' + · intros i j hi hj h' + exact h ⟨i, hi⟩ ⟨j, hj⟩ h' + theorem pairwise_replicate {α : Type _} {r : α → α → Prop} {x : α} (hx : r x x) : ∀ n : Nat, Pairwise r (List.replicate n x) | 0 => by simp - | n + 1 => by simp [mem_replicate, hx, pairwise_replicate hx n] + | n + 1 => by simp [mem_replicate, hx, pairwise_replicate hx n, replicate_succ] /-! ### Pairwise filtering -/ diff --git a/Batteries/Data/List/Perm.lean b/Batteries/Data/List/Perm.lean index 49299089a5..ab12db10b1 100644 --- a/Batteries/Data/List/Perm.lean +++ b/Batteries/Data/List/Perm.lean @@ -155,8 +155,8 @@ theorem Perm.filterMap (f : α → Option β) {l₁ l₂ : List α} (p : l₁ ~ filterMap f l₁ ~ filterMap f l₂ := by induction p with | nil => simp - | cons x _p IH => cases h : f x <;> simp [h, filterMap, IH, Perm.cons] - | swap x y l₂ => cases hx : f x <;> cases hy : f y <;> simp [hx, hy, filterMap, swap] + | cons x _p IH => cases h : f x <;> simp [h, filterMap_cons, IH, Perm.cons] + | swap x y l₂ => cases hx : f x <;> cases hy : f y <;> simp [hx, hy, filterMap_cons, swap] | trans _p₁ _p₂ IH₁ IH₂ => exact IH₁.trans IH₂ theorem Perm.map (f : α → β) {l₁ l₂ : List α} (p : l₁ ~ l₂) : map f l₁ ~ map f l₂ := @@ -643,7 +643,7 @@ theorem Perm.union {l₁ l₂ t₁ t₂ : List α} (p₁ : l₁ ~ l₂) (p₂ : theorem Perm.inter_right {l₁ l₂ : List α} (t₁ : List α) : l₁ ~ l₂ → l₁ ∩ t₁ ~ l₂ ∩ t₁ := .filter _ theorem Perm.inter_left (l : List α) {t₁ t₂ : List α} (p : t₁ ~ t₂) : l ∩ t₁ = l ∩ t₂ := - filter_congr' fun a _ => by simpa using p.mem_iff (a := a) + filter_congr fun a _ => by simpa using p.mem_iff (a := a) theorem Perm.inter {l₁ l₂ t₁ t₂ : List α} (p₁ : l₁ ~ l₂) (p₂ : t₁ ~ t₂) : l₁ ∩ t₁ ~ l₂ ∩ t₂ := p₂.inter_left l₂ ▸ p₁.inter_right t₁ diff --git a/Batteries/Data/String/Lemmas.lean b/Batteries/Data/String/Lemmas.lean index ad15234f51..1ea4e4f788 100644 --- a/Batteries/Data/String/Lemmas.lean +++ b/Batteries/Data/String/Lemmas.lean @@ -34,13 +34,13 @@ instance : Batteries.BEqOrd String := .compareOfLessAndEq String.lt_irrefl attribute [simp] toList -- prefer `String.data` over `String.toList` in lemmas -private theorem add_csize_pos : 0 < i + csize c := - Nat.add_pos_right _ (csize_pos c) +private theorem add_csize_pos : 0 < i + Char.utf8Size c := + Nat.add_pos_right _ (Char.utf8Size_pos c) -private theorem ne_add_csize_add_self : i ≠ n + csize c + i := +private theorem ne_add_csize_add_self : i ≠ n + Char.utf8Size c + i := Nat.ne_of_lt (Nat.lt_add_of_pos_left add_csize_pos) -private theorem ne_self_add_add_csize : i ≠ i + (n + csize c) := +private theorem ne_self_add_add_csize : i ≠ i + (n + Char.utf8Size c) := Nat.ne_of_lt (Nat.lt_add_of_pos_right add_csize_pos) /-- The UTF-8 byte length of a list of characters. (This is intended for specification purposes.) -/ @@ -52,7 +52,7 @@ private theorem ne_self_add_add_csize : i ≠ i + (n + csize c) := @[simp] theorem utf8Len_nil : utf8Len [] = 0 := rfl -@[simp] theorem utf8Len_cons (c cs) : utf8Len (c :: cs) = utf8Len cs + csize c := rfl +@[simp] theorem utf8Len_cons (c cs) : utf8Len (c :: cs) = utf8Len cs + c.utf8Size := rfl @[simp] theorem utf8Len_append (cs₁ cs₂) : utf8Len (cs₁ ++ cs₂) = utf8Len cs₁ + utf8Len cs₂ := by induction cs₁ <;> simp [*, Nat.add_right_comm] @@ -89,7 +89,7 @@ namespace Pos attribute [ext] ext -theorem lt_addChar (p : Pos) (c : Char) : p < p + c := Nat.lt_add_of_pos_right (csize_pos _) +theorem lt_addChar (p : Pos) (c : Char) : p < p + c := Nat.lt_add_of_pos_right (Char.utf8Size_pos _) private theorem zero_ne_addChar {i : Pos} {c : Char} : 0 ≠ i + c := ne_of_lt add_csize_pos @@ -201,11 +201,11 @@ theorem modify_of_valid (cs cs' : List Char) : rw [modify, set_of_valid, get_of_valid]; cases cs' <;> rfl theorem next_of_valid' (cs cs' : List Char) : - next ⟨cs ++ cs'⟩ ⟨utf8Len cs⟩ = ⟨utf8Len cs + csize (cs'.headD default)⟩ := by + next ⟨cs ++ cs'⟩ ⟨utf8Len cs⟩ = ⟨utf8Len cs + (cs'.headD default).utf8Size⟩ := by simp only [next, get_of_valid]; rfl theorem next_of_valid (cs : List Char) (c : Char) (cs' : List Char) : - next ⟨cs ++ c :: cs'⟩ ⟨utf8Len cs⟩ = ⟨utf8Len cs + csize c⟩ := next_of_valid' .. + next ⟨cs ++ c :: cs'⟩ ⟨utf8Len cs⟩ = ⟨utf8Len cs + c.utf8Size⟩ := next_of_valid' .. @[simp] theorem atEnd_iff (s : String) (p : Pos) : atEnd s p ↔ s.endPos ≤ p := decide_eq_true_iff _ @@ -218,7 +218,7 @@ theorem valid_next {p : Pos} (h : p.Valid s) (h₂ : p < s.endPos) : (next s p). simpa using Pos.Valid.mk (cs ++ [c]) cs' theorem utf8PrevAux_of_valid {cs cs' : List Char} {c : Char} {i p : Nat} - (hp : i + (utf8Len cs + csize c) = p) : + (hp : i + (utf8Len cs + c.utf8Size) = p) : utf8PrevAux (cs ++ c :: cs') ⟨i⟩ ⟨p⟩ = ⟨i + utf8Len cs⟩ := by match cs with | [] => simp [utf8PrevAux, ← hp, Pos.addChar_eq] @@ -233,7 +233,7 @@ theorem utf8PrevAux_of_valid {cs cs' : List Char} {c : Char} {i p : Nat} simp [Nat.add_assoc, Nat.add_comm] theorem prev_of_valid (cs : List Char) (c : Char) (cs' : List Char) : - prev ⟨cs ++ c :: cs'⟩ ⟨utf8Len cs + csize c⟩ = ⟨utf8Len cs⟩ := by + prev ⟨cs ++ c :: cs'⟩ ⟨utf8Len cs + c.utf8Size⟩ = ⟨utf8Len cs⟩ := by simp only [prev] refine (if_neg (Pos.ne_of_gt add_csize_pos)).trans ?_ rw [utf8PrevAux_of_valid] <;> simp @@ -327,9 +327,9 @@ theorem firstDiffPos_loop_eq (l₁ l₂ r₁ r₂ stop p) split · simp only [utf8Len_cons] subst b - rw [show next ⟨l₁ ++ a :: r₁⟩ ⟨p⟩ = ⟨utf8Len l₁ + csize a⟩ by simp [hl₁, next_of_valid]] + rw [show next ⟨l₁ ++ a :: r₁⟩ ⟨p⟩ = ⟨utf8Len l₁ + a.utf8Size⟩ by simp [hl₁, next_of_valid]] simpa [← hl₁, ← Nat.add_assoc, Nat.add_right_comm] using - firstDiffPos_loop_eq (l₁ ++ [a]) (l₂ ++ [a]) r₁ r₂ stop (p + csize a) + firstDiffPos_loop_eq (l₁ ++ [a]) (l₂ ++ [a]) r₁ r₂ stop (p + a.utf8Size) (by simp [hl₁]) (by simp [hl₂]) (by simp [hstop, ← Nat.add_assoc, Nat.add_right_comm]) · simp · next h => @@ -647,9 +647,8 @@ theorem offsetOfPosAux_of_valid : ∀ l m r n, unfold offsetOfPosAux rw [if_neg (by exact Nat.not_le.2 (Nat.lt_add_of_pos_right add_csize_pos))] simp only [List.append_assoc, atEnd_of_valid l (c::m++r)] - simp only [List.cons_append, ↓reduceDite, utf8Len_cons, next_of_valid l c (m ++ r), - List.length_cons, Nat.succ_eq_add_one] - simpa [← Nat.add_assoc, Nat.add_right_comm, Nat.succ_eq_add_one] using + simp only [List.cons_append, utf8Len_cons, next_of_valid l c (m ++ r)] + simpa [← Nat.add_assoc, Nat.add_right_comm] using offsetOfPosAux_of_valid (l++[c]) m r (n + 1) theorem offsetOfPos_of_valid (l r) : offsetOfPos ⟨l ++ r⟩ ⟨utf8Len l⟩ = l.length := by @@ -803,7 +802,8 @@ theorem toIterator : ∀ {s}, ValidFor l m r s → s.toIterator.ValidFor l.rever theorem get : ∀ {s}, ValidFor l (m₁ ++ c :: m₂) r s → s.get ⟨utf8Len m₁⟩ = c | _, ⟨⟩ => by simpa using get_of_valid (l ++ m₁) (c :: m₂ ++ r) -theorem next : ∀ {s}, ValidFor l (m₁ ++ c :: m₂) r s → s.next ⟨utf8Len m₁⟩ = ⟨utf8Len m₁ + csize c⟩ +theorem next : ∀ {s}, ValidFor l (m₁ ++ c :: m₂) r s → + s.next ⟨utf8Len m₁⟩ = ⟨utf8Len m₁ + c.utf8Size⟩ | _, ⟨⟩ => by simp only [Substring.next, utf8Len_append, utf8Len_cons, List.append_assoc, List.cons_append] rw [if_neg (mt Pos.ext_iff.1 ?a)] @@ -817,7 +817,8 @@ theorem next : ∀ {s}, ValidFor l (m₁ ++ c :: m₂) r s → s.next ⟨utf8Len theorem next_stop : ∀ {s}, ValidFor l m r s → s.next ⟨utf8Len m⟩ = ⟨utf8Len m⟩ | _, ⟨⟩ => by simp [Substring.next, Pos.add_eq] -theorem prev : ∀ {s}, ValidFor l (m₁ ++ c :: m₂) r s → s.prev ⟨utf8Len m₁ + csize c⟩ = ⟨utf8Len m₁⟩ +theorem prev : ∀ {s}, ValidFor l (m₁ ++ c :: m₂) r s → + s.prev ⟨utf8Len m₁ + c.utf8Size⟩ = ⟨utf8Len m₁⟩ | _, ⟨⟩ => by simp only [Substring.prev, List.append_assoc, List.cons_append] rw [if_neg (mt Pos.ext_iff.1 <| Ne.symm ?a)] @@ -965,7 +966,7 @@ theorem get : ∀ {s}, Valid s → s.toString.1 = m₁ ++ c :: m₂ → s.get simp only [h.toString] at e; subst e; simp [h.get] theorem next : ∀ {s}, Valid s → s.toString.1 = m₁ ++ c :: m₂ → - s.next ⟨utf8Len m₁⟩ = ⟨utf8Len m₁ + csize c⟩ + s.next ⟨utf8Len m₁⟩ = ⟨utf8Len m₁ + c.utf8Size⟩ | _, h, e => by let ⟨l, m, r, h⟩ := h.validFor simp only [h.toString] at e; subst e; simp [h.next] @@ -974,7 +975,7 @@ theorem next_stop : ∀ {s}, Valid s → s.next ⟨s.bsize⟩ = ⟨s.bsize⟩ | _, h => let ⟨l, m, r, h⟩ := h.validFor; by simp [h.bsize, h.next_stop] theorem prev : ∀ {s}, Valid s → s.toString.1 = m₁ ++ c :: m₂ → - s.prev ⟨utf8Len m₁ + csize c⟩ = ⟨utf8Len m₁⟩ + s.prev ⟨utf8Len m₁ + c.utf8Size⟩ = ⟨utf8Len m₁⟩ | _, h, e => by let ⟨l, m, r, h⟩ := h.validFor simp only [h.toString] at e; subst e; simp [h.prev] diff --git a/Batteries/Data/UInt.lean b/Batteries/Data/UInt.lean index 46fcb26b59..a9ce1b561a 100644 --- a/Batteries/Data/UInt.lean +++ b/Batteries/Data/UInt.lean @@ -33,7 +33,7 @@ theorem UInt8.toNat_zero : (0 : UInt8).toNat = 0 := rfl theorem UInt8.toNat_add (x y : UInt8) : (x + y).toNat = (x.toNat + y.toNat) % UInt8.size := rfl theorem UInt8.toNat_sub (x y : UInt8) : - (x - y).toNat = (x.toNat + (UInt8.size - y.toNat)) % UInt8.size := rfl + (x - y).toNat = (UInt8.size - y.toNat + x.toNat) % UInt8.size := rfl theorem UInt8.toNat_mul (x y : UInt8) : (x * y).toNat = (x.toNat * y.toNat) % UInt8.size := rfl @@ -80,7 +80,7 @@ theorem UInt16.toNat_zero : (0 : UInt16).toNat = 0 := rfl theorem UInt16.toNat_add (x y : UInt16) : (x + y).toNat = (x.toNat + y.toNat) % UInt16.size := rfl theorem UInt16.toNat_sub (x y : UInt16) : - (x - y).toNat = (x.toNat + (UInt16.size - y.toNat)) % UInt16.size := rfl + (x - y).toNat = (UInt16.size - y.toNat + x.toNat) % UInt16.size := rfl theorem UInt16.toNat_mul (x y : UInt16) : (x * y).toNat = (x.toNat * y.toNat) % UInt16.size := rfl @@ -127,7 +127,7 @@ theorem UInt32.toNat_zero : (0 : UInt32).toNat = 0 := rfl theorem UInt32.toNat_add (x y : UInt32) : (x + y).toNat = (x.toNat + y.toNat) % UInt32.size := rfl theorem UInt32.toNat_sub (x y : UInt32) : - (x - y).toNat = (x.toNat + (UInt32.size - y.toNat)) % UInt32.size := rfl + (x - y).toNat = (UInt32.size - y.toNat + x.toNat) % UInt32.size := rfl theorem UInt32.toNat_mul (x y : UInt32) : (x * y).toNat = (x.toNat * y.toNat) % UInt32.size := rfl @@ -174,7 +174,7 @@ theorem UInt64.toNat_zero : (0 : UInt64).toNat = 0 := rfl theorem UInt64.toNat_add (x y : UInt64) : (x + y).toNat = (x.toNat + y.toNat) % UInt64.size := rfl theorem UInt64.toNat_sub (x y : UInt64) : - (x - y).toNat = (x.toNat + (UInt64.size - y.toNat)) % UInt64.size := rfl + (x - y).toNat = (UInt64.size - y.toNat + x.toNat) % UInt64.size := rfl theorem UInt64.toNat_mul (x y : UInt64) : (x * y).toNat = (x.toNat * y.toNat) % UInt64.size := rfl @@ -235,7 +235,7 @@ theorem USize.toNat_zero : (0 : USize).toNat = 0 := rfl theorem USize.toNat_add (x y : USize) : (x + y).toNat = (x.toNat + y.toNat) % USize.size := rfl theorem USize.toNat_sub (x y : USize) : - (x - y).toNat = (x.toNat + (USize.size - y.toNat)) % USize.size := rfl + (x - y).toNat = (USize.size - y.toNat + x.toNat) % USize.size := rfl theorem USize.toNat_mul (x y : USize) : (x * y).toNat = (x.toNat * y.toNat) % USize.size := rfl diff --git a/Batteries/Lean/AttributeExtra.lean b/Batteries/Lean/AttributeExtra.lean index 7b8c92af3f..aeac0267c5 100644 --- a/Batteries/Lean/AttributeExtra.lean +++ b/Batteries/Lean/AttributeExtra.lean @@ -81,7 +81,7 @@ Registers a new parametric attribute. The `extra` field is a list of definitions which will be "pre-tagged" and are not subject to the usual restriction on tagging in the same file as the declaration. -/ -def registerParametricAttributeExtra [Inhabited α] (impl : ParametricAttributeImpl α) +def registerParametricAttributeExtra (impl : ParametricAttributeImpl α) (extra : List (Name × α)) : IO (ParametricAttributeExtra α) := do let attr ← registerParametricAttribute impl pure { attr, base := extra.foldl (fun s (a, b) => s.insert a b) {} } diff --git a/Batteries/Lean/Name.lean b/Batteries/Lean/Name.lean deleted file mode 100644 index b5530e49d1..0000000000 --- a/Batteries/Lean/Name.lean +++ /dev/null @@ -1,28 +0,0 @@ -/- -Copyright (c) 2023 Mario Carneiro. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro --/ -import Lean.Data.Name - -namespace Lean.Name - -/-- -Returns true if this a part of name that is internal or dynamically -generated so that it may easily be changed. - -Generally, user code should not explicitly use internal names. --/ -def isInternalDetail : Name → Bool - | .str p s => - s.startsWith "_" - || matchPrefix s "eq_" - || matchPrefix s "match_" - || matchPrefix s "proof_" - || p.isInternalOrNum - | .num _ _ => true - | p => p.isInternalOrNum -where - /-- Check that a string begins with the given prefix, and then is only digit characters. -/ - matchPrefix (s : String) (pre : String) := - s.startsWith pre && (s |>.drop pre.length |>.all Char.isDigit) diff --git a/Batteries/Lean/PersistentHashMap.lean b/Batteries/Lean/PersistentHashMap.lean index 5054e15758..b297f2fa52 100644 --- a/Batteries/Lean/PersistentHashMap.lean +++ b/Batteries/Lean/PersistentHashMap.lean @@ -10,15 +10,6 @@ namespace Lean.PersistentHashMap variable [BEq α] [Hashable α] -/-- -Similar to `insert`, but also returns a Boolean flag indicating whether an -existing entry has been replaced with `a => b`. --/ -def insert' (m : PersistentHashMap α β) (a : α) (b : β) : PersistentHashMap α β × Bool := - let oldSize := m.size - let m := m.insert a b - (m, m.size == oldSize) - /-- Builds a `PersistentHashMap` from a list of key-value pairs. Values of duplicated keys are replaced by their respective last occurrences. diff --git a/Batteries/Lean/PersistentHashSet.lean b/Batteries/Lean/PersistentHashSet.lean index 51e33e8487..0a8231d723 100644 --- a/Batteries/Lean/PersistentHashSet.lean +++ b/Batteries/Lean/PersistentHashSet.lean @@ -56,16 +56,6 @@ def all (s : PersistentHashSet α) (f : α → Bool) : Bool := instance : BEq (PersistentHashSet α) where beq s t := s.all (t.contains ·) && t.all (s.contains ·) -/-- -Similar to `insert`, but also returns a Boolean flag indicating whether an -existing entry has been replaced with `a => b`. --/ -@[inline] -def insert' (s : PersistentHashSet α) (a : α) : PersistentHashSet α × Bool := - let oldSize := s.size - let s := s.insert a - (s, s.size == oldSize) - /-- Insert all elements from a collection into a `PersistentHashSet`. -/ diff --git a/Batteries/Tactic/Alias.lean b/Batteries/Tactic/Alias.lean index b7eaeaf713..07b91dc76e 100644 --- a/Batteries/Tactic/Alias.lean +++ b/Batteries/Tactic/Alias.lean @@ -6,6 +6,7 @@ Authors: Mario Carneiro, David Renshaw, François G. Dorais import Lean.Elab.Command import Lean.Elab.DeclarationRange import Lean.Compiler.NoncomputableAttr +import Lean.DocString import Batteries.CodeAction.Deprecated /-! diff --git a/Batteries/Tactic/Lint/Simp.lean b/Batteries/Tactic/Lint/Simp.lean index 3dc801fe0a..9196e4ab74 100644 --- a/Batteries/Tactic/Lint/Simp.lean +++ b/Batteries/Tactic/Lint/Simp.lean @@ -96,7 +96,7 @@ def decorateError (msg : MessageData) (k : MetaM α) : MetaM α := do def formatLemmas (usedSimps : Simp.UsedSimps) (simpName : String) : MetaM MessageData := do let mut args := #[] let env ← getEnv - for (thm, _) in usedSimps.toArray.qsort (·.2 < ·.2) do + for (thm, _) in usedSimps.map.toArray.qsort (·.2 < ·.2) do if let .decl declName := thm then if env.contains declName && declName != ``eq_self then args := args.push (← mkConstWithFreshMVarLevels declName) @@ -120,7 +120,7 @@ https://leanprover-community.github.io/mathlib_docs/notes.html#simp-normal%20for else do let (e, s) ← dsimp lhs ctx return (Simp.Result.mk e .none .true, s) - if prf1Stats.usedTheorems.contains (.decl declName) then return none + if prf1Stats.usedTheorems.map.contains (.decl declName) then return none let ({ expr := rhs', .. }, stats) ← decorateError "simplify fails on right-hand side:" <| if !isRfl then diff --git a/Batteries/Tactic/NoMatch.lean b/Batteries/Tactic/NoMatch.lean index 9dc517b33f..7efe68e7d2 100644 --- a/Batteries/Tactic/NoMatch.lean +++ b/Batteries/Tactic/NoMatch.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Lean.Elab.ElabRules +import Lean.DocString /-! Deprecation warnings for `match ⋯ with.`, `fun.`, `λ.`, and `intro.`. diff --git a/Batteries/Tactic/PrintPrefix.lean b/Batteries/Tactic/PrintPrefix.lean index 0046f0a4d7..63d2c3bbf5 100644 --- a/Batteries/Tactic/PrintPrefix.lean +++ b/Batteries/Tactic/PrintPrefix.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Shing Tak Lam. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Shing Tak Lam, Daniel Selsam, Mario Carneiro -/ -import Batteries.Lean.Name import Batteries.Lean.Util.EnvSearch import Batteries.Lean.Delaborator import Lean.Elab.Tactic.Config diff --git a/Batteries/Tactic/SqueezeScope.lean b/Batteries/Tactic/SqueezeScope.lean index 78a1d2a540..df289dc982 100644 --- a/Batteries/Tactic/SqueezeScope.lean +++ b/Batteries/Tactic/SqueezeScope.lean @@ -85,7 +85,8 @@ elab_rules : tactic throw e if let some new := new then for (_, stx, usedSimps) in new do - let usedSimps := usedSimps.foldl (fun s usedSimps => usedSimps.foldl .insert s) {} + let usedSimps := usedSimps.reverse.foldl + (fun s usedSimps => usedSimps.toArray.foldl .insert s) {} let stx' ← mkSimpCallStx stx usedSimps TryThis.addSuggestion stx[0] stx' (origSpan? := stx) diff --git a/lean-toolchain b/lean-toolchain index 4ef27c472e..d69d1ed200 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.9.0 +leanprover/lean4:v4.10.0-rc1 diff --git a/test/absurd.lean b/test/absurd.lean index 4974467e5a..1817c0e715 100644 --- a/test/absurd.lean +++ b/test/absurd.lean @@ -43,5 +43,5 @@ h : p → False ⊢ p -/ #guard_msgs in -example (h : p → False) : 0 = 3 := by +example {p : Prop} (h : p → False) : 0 = 3 := by absurd h diff --git a/test/congr.lean b/test/congr.lean index 5303fbd74b..a71aed0315 100644 --- a/test/congr.lean +++ b/test/congr.lean @@ -38,7 +38,7 @@ section -- Adaptation note: the next two examples have always failed if `List.ext` was in scope, -- but until nightly-2024-04-24 (when `List.ext` was upstreamed), it wasn't in scope. -- In order to preserve the test behaviour we locally remove the `ext` attribute. -attribute [-ext] List.ext +attribute [-ext] List.ext_getElem? private opaque List.sum : List Nat → Nat From bba0af6e930ebcbabfacf021b21dd881d58aaa9d Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 1 Jul 2024 16:20:32 +1000 Subject: [PATCH 019/177] chore: remove 'meta if' from lakefile (#869) --- lakefile.lean | 3 --- 1 file changed, 3 deletions(-) diff --git a/lakefile.lean b/lakefile.lean index a59342f68c..7ee5efafe6 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -18,9 +18,6 @@ lean_exe runLinter where srcDir := "scripts" supportInterpreter := true -meta if get_config? doc |>.isSome then -require «doc-gen4» from git "https://github.com/leanprover/doc-gen4" @ "main" - @[test_driver] lean_exe test where srcDir := "scripts" From 0682b7591187e8b13cd02fdc4e614bca2cd0f3a1 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 4 Jul 2024 09:43:31 +1000 Subject: [PATCH 020/177] chore: adaptations for nightly-2024-07-01 (#870) --- Batteries/Data/Array/Basic.lean | 18 ------------ Batteries/Data/List.lean | 1 - Batteries/Data/List/Init/Attach.lean | 44 ---------------------------- Batteries/Data/List/Perm.lean | 3 -- Batteries/Data/Range/Lemmas.lean | 1 - lean-toolchain | 2 +- 6 files changed, 1 insertion(+), 68 deletions(-) delete mode 100644 Batteries/Data/List/Init/Attach.lean diff --git a/Batteries/Data/Array/Basic.lean b/Batteries/Data/Array/Basic.lean index 1dddea3c14..ad6b8129e6 100644 --- a/Batteries/Data/Array/Basic.lean +++ b/Batteries/Data/Array/Basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Arthur Paulino, Floris van Doorn, Jannis Limperg -/ -import Batteries.Data.List.Init.Attach import Batteries.Data.Array.Init.Lemmas /-! @@ -130,23 +129,6 @@ protected def maxI [ord : Ord α] [Inhabited α] (xs : Array α) (start := 0) (stop := xs.size) : α := xs.minI (ord := ord.opposite) start stop -/-- -Unsafe implementation of `attachWith`, taking advantage of the fact that the representation of -`Array {x // P x}` is the same as the input `Array α`. --/ -@[inline] private unsafe def attachWithImpl - (xs : Array α) (P : α → Prop) (_ : ∀ x ∈ xs, P x) : Array {x // P x} := unsafeCast xs - -/-- `O(1)`. "Attach" a proof `P x` that holds for all the elements of `xs` to produce a new array - with the same elements but in the type `{x // P x}`. -/ -@[implemented_by attachWithImpl] def attachWith - (xs : Array α) (P : α → Prop) (H : ∀ x ∈ xs, P x) : Array {x // P x} := - ⟨xs.data.attachWith P fun x h => H x (Array.Mem.mk h)⟩ - -/-- `O(1)`. "Attach" the proof that the elements of `xs` are in `xs` to produce a new array - with the same elements but in the type `{x // x ∈ xs}`. -/ -@[inline] def attach (xs : Array α) : Array {x // x ∈ xs} := xs.attachWith _ fun _ => id - /-- `O(|join L|)`. `join L` concatenates all the arrays in `L` into one array. * `join #[#[a], #[], #[b, c], #[d, e, f]] = #[a, b, c, d, e, f]` diff --git a/Batteries/Data/List.lean b/Batteries/Data/List.lean index d060ba89c8..998160b904 100644 --- a/Batteries/Data/List.lean +++ b/Batteries/Data/List.lean @@ -1,7 +1,6 @@ import Batteries.Data.List.Basic import Batteries.Data.List.Count import Batteries.Data.List.EraseIdx -import Batteries.Data.List.Init.Attach import Batteries.Data.List.Init.Lemmas import Batteries.Data.List.Lemmas import Batteries.Data.List.Pairwise diff --git a/Batteries/Data/List/Init/Attach.lean b/Batteries/Data/List/Init/Attach.lean deleted file mode 100644 index d2b2bf0990..0000000000 --- a/Batteries/Data/List/Init/Attach.lean +++ /dev/null @@ -1,44 +0,0 @@ -/- -Copyright (c) 2023 Mario Carneiro. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro --/ - -namespace List - -/-- `O(n)`. Partial map. If `f : Π a, P a → β` is a partial function defined on - `a : α` satisfying `P`, then `pmap f l h` is essentially the same as `map f l` - but is defined only when all members of `l` satisfy `P`, using the proof - to apply `f`. -/ -@[simp] def pmap {P : α → Prop} (f : ∀ a, P a → β) : ∀ l : List α, (H : ∀ a ∈ l, P a) → List β - | [], _ => [] - | a :: l, H => f a (forall_mem_cons.1 H).1 :: pmap f l (forall_mem_cons.1 H).2 - -/-- -Unsafe implementation of `attachWith`, taking advantage of the fact that the representation of -`List {x // P x}` is the same as the input `List α`. -(Someday, the compiler might do this optimization automatically, but until then...) --/ -@[inline] private unsafe def attachWithImpl - (l : List α) (P : α → Prop) (_ : ∀ x ∈ l, P x) : List {x // P x} := unsafeCast l - -/-- `O(1)`. "Attach" a proof `P x` that holds for all the elements of `l` to produce a new list - with the same elements but in the type `{x // P x}`. -/ -@[implemented_by attachWithImpl] def attachWith - (l : List α) (P : α → Prop) (H : ∀ x ∈ l, P x) : List {x // P x} := pmap Subtype.mk l H - -/-- `O(1)`. "Attach" the proof that the elements of `l` are in `l` to produce a new list - with the same elements but in the type `{x // x ∈ l}`. -/ -@[inline] def attach (l : List α) : List {x // x ∈ l} := attachWith l _ fun _ => id - -/-- Implementation of `pmap` using the zero-copy version of `attach`. -/ -@[inline] private def pmapImpl {P : α → Prop} (f : ∀ a, P a → β) (l : List α) (H : ∀ a ∈ l, P a) : - List β := (l.attachWith _ H).map fun ⟨x, h'⟩ => f x h' - -@[csimp] private theorem pmap_eq_pmapImpl : @pmap = @pmapImpl := by - funext α β p f L h' - let rec go : ∀ L' (hL' : ∀ ⦃x⦄, x ∈ L' → p x), - pmap f L' hL' = map (fun ⟨x, hx⟩ => f x hx) (pmap Subtype.mk L' hL') - | nil, hL' => rfl - | cons _ L', hL' => congrArg _ <| go L' fun _ hx => hL' (.tail _ hx) - exact go L h' diff --git a/Batteries/Data/List/Perm.lean b/Batteries/Data/List/Perm.lean index ab12db10b1..49e45ecc8b 100644 --- a/Batteries/Data/List/Perm.lean +++ b/Batteries/Data/List/Perm.lean @@ -4,10 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Batteries.Tactic.Alias -import Batteries.Data.List.Init.Attach import Batteries.Data.List.Pairwise --- Adaptation note: nightly-2024-03-18. We should be able to remove this after nightly-2024-03-19. -import Lean.Elab.Tactic.Rfl /-! # List Permutations diff --git a/Batteries/Data/Range/Lemmas.lean b/Batteries/Data/Range/Lemmas.lean index 586f32a7e8..d71082848f 100644 --- a/Batteries/Data/Range/Lemmas.lean +++ b/Batteries/Data/Range/Lemmas.lean @@ -5,7 +5,6 @@ Authors: Mario Carneiro -/ import Batteries.Tactic.SeqFocus import Batteries.Data.List.Lemmas -import Batteries.Data.List.Init.Attach namespace Std.Range diff --git a/lean-toolchain b/lean-toolchain index d69d1ed200..c73e540213 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.10.0-rc1 +leanprover/lean4:nightly-2024-07-01 From dc4a6b1ac3cd502988e283d5c9c5fdf261125a07 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 4 Jul 2024 17:08:55 +1000 Subject: [PATCH 021/177] Revert 'chore: adaptations for nightly-2024-07-01 (#870)' (#872) * Revert 'chore: adaptations for nightly-2024-07-01 (#870)' * Revert 'chore: adaptations for nightly-2024-07-01 (#870)' --- Batteries/Data/Array/Basic.lean | 18 ++++++++++++ Batteries/Data/List.lean | 1 + Batteries/Data/List/Init/Attach.lean | 44 ++++++++++++++++++++++++++++ Batteries/Data/List/Perm.lean | 3 ++ Batteries/Data/Range/Lemmas.lean | 1 + lean-toolchain | 2 +- 6 files changed, 68 insertions(+), 1 deletion(-) create mode 100644 Batteries/Data/List/Init/Attach.lean diff --git a/Batteries/Data/Array/Basic.lean b/Batteries/Data/Array/Basic.lean index ad6b8129e6..1dddea3c14 100644 --- a/Batteries/Data/Array/Basic.lean +++ b/Batteries/Data/Array/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Arthur Paulino, Floris van Doorn, Jannis Limperg -/ +import Batteries.Data.List.Init.Attach import Batteries.Data.Array.Init.Lemmas /-! @@ -129,6 +130,23 @@ protected def maxI [ord : Ord α] [Inhabited α] (xs : Array α) (start := 0) (stop := xs.size) : α := xs.minI (ord := ord.opposite) start stop +/-- +Unsafe implementation of `attachWith`, taking advantage of the fact that the representation of +`Array {x // P x}` is the same as the input `Array α`. +-/ +@[inline] private unsafe def attachWithImpl + (xs : Array α) (P : α → Prop) (_ : ∀ x ∈ xs, P x) : Array {x // P x} := unsafeCast xs + +/-- `O(1)`. "Attach" a proof `P x` that holds for all the elements of `xs` to produce a new array + with the same elements but in the type `{x // P x}`. -/ +@[implemented_by attachWithImpl] def attachWith + (xs : Array α) (P : α → Prop) (H : ∀ x ∈ xs, P x) : Array {x // P x} := + ⟨xs.data.attachWith P fun x h => H x (Array.Mem.mk h)⟩ + +/-- `O(1)`. "Attach" the proof that the elements of `xs` are in `xs` to produce a new array + with the same elements but in the type `{x // x ∈ xs}`. -/ +@[inline] def attach (xs : Array α) : Array {x // x ∈ xs} := xs.attachWith _ fun _ => id + /-- `O(|join L|)`. `join L` concatenates all the arrays in `L` into one array. * `join #[#[a], #[], #[b, c], #[d, e, f]] = #[a, b, c, d, e, f]` diff --git a/Batteries/Data/List.lean b/Batteries/Data/List.lean index 998160b904..d060ba89c8 100644 --- a/Batteries/Data/List.lean +++ b/Batteries/Data/List.lean @@ -1,6 +1,7 @@ import Batteries.Data.List.Basic import Batteries.Data.List.Count import Batteries.Data.List.EraseIdx +import Batteries.Data.List.Init.Attach import Batteries.Data.List.Init.Lemmas import Batteries.Data.List.Lemmas import Batteries.Data.List.Pairwise diff --git a/Batteries/Data/List/Init/Attach.lean b/Batteries/Data/List/Init/Attach.lean new file mode 100644 index 0000000000..d2b2bf0990 --- /dev/null +++ b/Batteries/Data/List/Init/Attach.lean @@ -0,0 +1,44 @@ +/- +Copyright (c) 2023 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ + +namespace List + +/-- `O(n)`. Partial map. If `f : Π a, P a → β` is a partial function defined on + `a : α` satisfying `P`, then `pmap f l h` is essentially the same as `map f l` + but is defined only when all members of `l` satisfy `P`, using the proof + to apply `f`. -/ +@[simp] def pmap {P : α → Prop} (f : ∀ a, P a → β) : ∀ l : List α, (H : ∀ a ∈ l, P a) → List β + | [], _ => [] + | a :: l, H => f a (forall_mem_cons.1 H).1 :: pmap f l (forall_mem_cons.1 H).2 + +/-- +Unsafe implementation of `attachWith`, taking advantage of the fact that the representation of +`List {x // P x}` is the same as the input `List α`. +(Someday, the compiler might do this optimization automatically, but until then...) +-/ +@[inline] private unsafe def attachWithImpl + (l : List α) (P : α → Prop) (_ : ∀ x ∈ l, P x) : List {x // P x} := unsafeCast l + +/-- `O(1)`. "Attach" a proof `P x` that holds for all the elements of `l` to produce a new list + with the same elements but in the type `{x // P x}`. -/ +@[implemented_by attachWithImpl] def attachWith + (l : List α) (P : α → Prop) (H : ∀ x ∈ l, P x) : List {x // P x} := pmap Subtype.mk l H + +/-- `O(1)`. "Attach" the proof that the elements of `l` are in `l` to produce a new list + with the same elements but in the type `{x // x ∈ l}`. -/ +@[inline] def attach (l : List α) : List {x // x ∈ l} := attachWith l _ fun _ => id + +/-- Implementation of `pmap` using the zero-copy version of `attach`. -/ +@[inline] private def pmapImpl {P : α → Prop} (f : ∀ a, P a → β) (l : List α) (H : ∀ a ∈ l, P a) : + List β := (l.attachWith _ H).map fun ⟨x, h'⟩ => f x h' + +@[csimp] private theorem pmap_eq_pmapImpl : @pmap = @pmapImpl := by + funext α β p f L h' + let rec go : ∀ L' (hL' : ∀ ⦃x⦄, x ∈ L' → p x), + pmap f L' hL' = map (fun ⟨x, hx⟩ => f x hx) (pmap Subtype.mk L' hL') + | nil, hL' => rfl + | cons _ L', hL' => congrArg _ <| go L' fun _ hx => hL' (.tail _ hx) + exact go L h' diff --git a/Batteries/Data/List/Perm.lean b/Batteries/Data/List/Perm.lean index 49e45ecc8b..ab12db10b1 100644 --- a/Batteries/Data/List/Perm.lean +++ b/Batteries/Data/List/Perm.lean @@ -4,7 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Batteries.Tactic.Alias +import Batteries.Data.List.Init.Attach import Batteries.Data.List.Pairwise +-- Adaptation note: nightly-2024-03-18. We should be able to remove this after nightly-2024-03-19. +import Lean.Elab.Tactic.Rfl /-! # List Permutations diff --git a/Batteries/Data/Range/Lemmas.lean b/Batteries/Data/Range/Lemmas.lean index d71082848f..586f32a7e8 100644 --- a/Batteries/Data/Range/Lemmas.lean +++ b/Batteries/Data/Range/Lemmas.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro -/ import Batteries.Tactic.SeqFocus import Batteries.Data.List.Lemmas +import Batteries.Data.List.Init.Attach namespace Std.Range diff --git a/lean-toolchain b/lean-toolchain index c73e540213..d69d1ed200 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-07-01 +leanprover/lean4:v4.10.0-rc1 From 9496541a88f4a9e906870434b1c7c12aa3524f00 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Mon, 8 Jul 2024 17:31:05 +0200 Subject: [PATCH 022/177] chore: remove `proof_wanted` in HashMap.Lemmas (#874) --- Batteries/Data/HashMap/Lemmas.lean | 14 +++++++++++--- 1 file changed, 11 insertions(+), 3 deletions(-) diff --git a/Batteries/Data/HashMap/Lemmas.lean b/Batteries/Data/HashMap/Lemmas.lean index 50efb40455..290f4c0cff 100644 --- a/Batteries/Data/HashMap/Lemmas.lean +++ b/Batteries/Data/HashMap/Lemmas.lean @@ -5,7 +5,13 @@ Authors: Scott Morrison -/ import Batteries.Data.HashMap.Basic import Batteries.Data.Array.Lemmas -import Batteries.Util.ProofWanted + +/-! +# Lemmas for `Batteries.HashMap` + +Note that Lean core provides an alternative hash map implementation, `Std.HashMap`, which comes with +more lemmas. See the module `Std.Data.HashMap.Lemmas`. +-/ namespace Batteries.HashMap @@ -21,5 +27,7 @@ end Imp @[simp] theorem empty_find? [BEq α] [Hashable α] {a : α} : (∅ : HashMap α β).find? a = none := by simp [find?, Imp.find?] -proof_wanted insert_find? [BEq α] [Hashable α] (m : HashMap α β) (a a' : α) (b : β) : - (m.insert a b).find? a' = if a' == a then some b else m.find? a' +-- `Std.HashMap` has this lemma (as `getElem?_insert`) and many more, so working on this +-- `proof_wanted` is likely not a good use of your time. +-- proof_wanted insert_find? [BEq α] [Hashable α] (m : HashMap α β) (a a' : α) (b : β) : +-- (m.insert a b).find? a' = if a' == a then some b else m.find? a' From 920a09209f3eff179ff2555e9809748f87606c62 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 9 Jul 2024 01:31:46 +1000 Subject: [PATCH 023/177] chore: remove nonterminal simps in UnionFind (#868) --- Batteries/Data/UnionFind/Basic.lean | 104 ++++++++++++++++++---------- 1 file changed, 68 insertions(+), 36 deletions(-) diff --git a/Batteries/Data/UnionFind/Basic.lean b/Batteries/Data/UnionFind/Basic.lean index 24966452c2..292f4285eb 100644 --- a/Batteries/Data/UnionFind/Basic.lean +++ b/Batteries/Data/UnionFind/Basic.lean @@ -46,13 +46,17 @@ theorem lt_of_parentD : parentD arr i ≠ i → i < arr.size := theorem parentD_set {arr : Array UFNode} {x v i} : parentD (arr.set x v) i = if x.1 = i then v.parent else parentD arr i := by - rw [parentD]; simp [Array.get_eq_getElem, parentD] - split <;> [split <;> simp [Array.get_set, *]; split <;> [(subst i; cases ‹¬_› x.2); rfl]] + rw [parentD]; simp only [Array.size_set, Array.get_eq_getElem, parentD] + split + · split <;> simp_all + · split <;> [(subst i; cases ‹¬_› x.2); rfl] theorem rankD_set {arr : Array UFNode} {x v i} : rankD (arr.set x v) i = if x.1 = i then v.rank else rankD arr i := by - rw [rankD]; simp [Array.get_eq_getElem, rankD] - split <;> [split <;> simp [Array.get_set, *]; split <;> [(subst i; cases ‹¬_› x.2); rfl]] + rw [rankD]; simp only [Array.size_set, Array.get_eq_getElem, rankD] + split + · split <;> simp_all + · split <;> [(subst i; cases ‹¬_› x.2); rfl] end UnionFind @@ -146,7 +150,7 @@ theorem rank'_lt_rankMax (self : UnionFind) (i : Fin self.size) : let rec go : ∀ {l} {x : UFNode}, x ∈ l → x.rank ≤ List.foldr (max ·.rank) 0 l | a::l, _, List.Mem.head _ => by dsimp; apply Nat.le_max_left | a::l, _, .tail _ h => by dsimp; exact Nat.le_trans (go h) (Nat.le_max_right ..) - simp [rankMax, Array.foldr_eq_foldr_data] + simp only [Array.get_eq_getElem, rankMax, Array.foldr_eq_foldr_data] exact Nat.lt_succ.2 <| go (self.arr.data.get_mem i.1 i.2) theorem rankD_lt_rankMax (self : UnionFind) (i : Nat) : @@ -156,11 +160,11 @@ theorem rankD_lt_rankMax (self : UnionFind) (i : Nat) : theorem lt_rankMax (self : UnionFind) (i : Nat) : self.rank i < self.rankMax := rankD_lt_rankMax .. theorem push_rankD (arr : Array UFNode) : rankD (arr.push ⟨arr.size, 0⟩) i = rankD arr i := by - simp [rankD, Array.get_eq_getElem, Array.get_push] + simp only [rankD, Array.size_push, Array.get_eq_getElem, Array.get_push, dite_eq_ite] split <;> split <;> first | simp | cases ‹¬_› (Nat.lt_succ_of_lt ‹_›) theorem push_parentD (arr : Array UFNode) : parentD (arr.push ⟨arr.size, 0⟩) i = parentD arr i := by - simp [parentD, Array.get_eq_getElem, Array.get_push] + simp only [parentD, Array.size_push, Array.get_eq_getElem, Array.get_push, dite_eq_ite] split <;> split <;> try simp · exact Nat.le_antisymm (Nat.ge_of_not_lt ‹_›) (Nat.le_of_lt_succ ‹_›) · cases ‹¬_› (Nat.lt_succ_of_lt ‹_›) @@ -169,9 +173,9 @@ theorem push_parentD (arr : Array UFNode) : parentD (arr.push ⟨arr.size, 0⟩) def push (self : UnionFind) : UnionFind where arr := self.arr.push ⟨self.arr.size, 0⟩ parentD_lt {i} := by - simp [push_parentD]; simp [parentD] + simp only [Array.size_push, push_parentD]; simp only [parentD, Array.get_eq_getElem] split <;> [exact fun _ => Nat.lt_succ_of_lt (self.parent'_lt _); exact id] - rankD_lt := by simp [push_parentD, push_rankD]; exact self.rank_lt + rankD_lt := by simp only [push_parentD, ne_eq, push_rankD]; exact self.rank_lt /-- Root of a union-find node. -/ def root (self : UnionFind) (x : Fin self.size) : Fin self.size := @@ -205,18 +209,23 @@ termination_by self.rankMax - self.rank x theorem parent_rootD (self : UnionFind) (x : Nat) : self.parent (self.rootD x) = self.rootD x := by - rw [rootD]; split <;> - [simp [parentD, parent_root, -Array.get_eq_getElem]; simp [parentD_of_not_lt, *]] + rw [rootD] + split + · simp [parentD, parent_root, -Array.get_eq_getElem] + · simp [parentD_of_not_lt, *] @[nolint unusedHavesSuffices] theorem rootD_parent (self : UnionFind) (x : Nat) : self.rootD (self.parent x) = self.rootD x := by - simp [rootD, parent_lt]; split <;> simp [parentD, parentD_of_not_lt, *, -Array.get_eq_getElem] - (conv => rhs; rw [root]); split - · rw [root, dif_pos] <;> simp [*, -Array.get_eq_getElem] - · simp + simp only [rootD, Array.data_length, parent_lt] + split + · simp only [parentD, ↓reduceDIte, *] + (conv => rhs; rw [root]); split + · rw [root, dif_pos] <;> simp_all + · simp + · simp only [not_false_eq_true, parentD_of_not_lt, *] theorem rootD_lt {self : UnionFind} {x : Nat} : self.rootD x < self.size ↔ x < self.size := by - simp [rootD]; split <;> simp [*] + simp only [rootD, Array.data_length]; split <;> simp [*] @[nolint unusedHavesSuffices] theorem rootD_eq_self {self : UnionFind} {x : Nat} : self.rootD x = x ↔ self.parent x = x := by @@ -273,7 +282,9 @@ termination_by self.rankMax - self.rank x @[nolint unusedHavesSuffices] theorem findAux_root {self : UnionFind} {x : Fin self.size} : (findAux self x).root = self.root x := by - rw [findAux, root]; simp; split <;> simp + rw [findAux, root] + simp only [Array.data_length, Array.get_eq_getElem, dite_eq_ite] + split <;> simp only have := Nat.sub_lt_sub_left (self.lt_rankMax x) (self.rank'_lt _ ‹_›) exact findAux_root termination_by self.rankMax - self.rank x @@ -286,7 +297,7 @@ theorem findAux_s {self : UnionFind} {x : Fin self.size} : rw [show self.rootD _ = (self.findAux ⟨_, self.parent'_lt x⟩).root from _] · rw [findAux]; split <;> rfl · rw [← rootD_parent, parent, parentD_eq] - simp [findAux_root, rootD] + simp only [rootD, Array.get_eq_getElem, Array.data_length, findAux_root] apply dif_pos exact parent'_lt .. @@ -299,7 +310,8 @@ theorem rankD_findAux {self : UnionFind} {x : Fin self.size} : rw [rankD_eq' (by simp [FindAux.size_eq, h]), Array.get_modify (by rwa [FindAux.size_eq])] split <;> simp [← rankD_eq, rankD_findAux (x := ⟨_, self.parent'_lt x⟩), -Array.get_eq_getElem] else - simp [rank, rankD]; rw [dif_neg (by rwa [FindAux.size_eq]), dif_neg h] + simp only [rankD, Array.data_length, Array.get_eq_getElem, rank] + rw [dif_neg (by rwa [FindAux.size_eq]), dif_neg h] termination_by self.rankMax - self.rank x theorem parentD_findAux {self : UnionFind} {x : Fin self.size} : @@ -311,7 +323,7 @@ theorem parentD_findAux {self : UnionFind} {x : Fin self.size} : · next h => rw [parentD]; split <;> rename_i h' · rw [Array.get_modify (by simpa using h')] - simp [@eq_comm _ i, -Array.get_eq_getElem] + simp only [Array.data_length, @eq_comm _ i] split <;> simp [← parentD_eq, -Array.get_eq_getElem] · rw [if_neg (mt (by rintro rfl; simp [FindAux.size_eq]) h')] rw [parentD, dif_neg]; simpa using h' @@ -330,9 +342,11 @@ theorem parentD_findAux_lt {self : UnionFind} {x : Fin self.size} (h : i < self. if h' : (self.arr.get x).parent = x then rw [findAux_s, if_pos h']; apply self.parentD_lt h else - rw [parentD_findAux]; split <;> [simp [rootD_lt]; skip] - have := Nat.sub_lt_sub_left (self.lt_rankMax x) (self.rank'_lt _ ‹_›) - apply parentD_findAux_lt h + rw [parentD_findAux] + split + · simp [rootD_lt] + · have := Nat.sub_lt_sub_left (self.lt_rankMax x) (self.rank'_lt _ ‹_›) + apply parentD_findAux_lt h termination_by self.rankMax - self.rank x theorem parentD_findAux_or (self : UnionFind) (x : Fin self.size) (i) : @@ -341,10 +355,12 @@ theorem parentD_findAux_or (self : UnionFind) (x : Fin self.size) (i) : if h' : (self.arr.get x).parent = x then rw [findAux_s, if_pos h']; exact .inr rfl else - rw [parentD_findAux]; split <;> [simp [*]; skip] - have := Nat.sub_lt_sub_left (self.lt_rankMax x) (self.rank'_lt _ ‹_›) - exact (parentD_findAux_or self ⟨_, self.parent'_lt x⟩ i).imp_left <| .imp_right fun h => by - simp only [h, ← parentD_eq, rootD_parent, Array.data_length] + rw [parentD_findAux] + split + · simp [*] + · have := Nat.sub_lt_sub_left (self.lt_rankMax x) (self.rank'_lt _ ‹_›) + exact (parentD_findAux_or self ⟨_, self.parent'_lt x⟩ i).imp_left <| .imp_right fun h => by + simp only [h, ← parentD_eq, rootD_parent, Array.data_length] termination_by self.rankMax - self.rank x theorem lt_rankD_findAux {self : UnionFind} {x : Fin self.size} : @@ -365,7 +381,9 @@ def find (self : UnionFind) (x : Fin self.size) : let r := self.findAux x { 1.arr := r.s 2.1.val := r.root - 1.parentD_lt := fun h => by simp [FindAux.size_eq] at *; exact parentD_findAux_lt h + 1.parentD_lt := fun h => by + simp only [Array.data_length, FindAux.size_eq] at * + exact parentD_findAux_lt h 1.rankD_lt := fun h => by rw [rankD_findAux, rankD_findAux]; exact lt_rankD_findAux h 2.1.isLt := show _ < r.s.size by rw [r.size_eq]; exact r.root.2 2.2 := by simp [size, r.size_eq] } @@ -398,7 +416,8 @@ def findD (self : UnionFind) (x : Nat) : UnionFind × Nat := @[simp] theorem find_parent_1 (self : UnionFind) (x : Fin self.size) : (self.find x).1.parent x = self.rootD x := by - simp [find, parent]; rw [parentD_findAux, if_pos rfl] + simp only [parent, Array.data_length, find] + rw [parentD_findAux, if_pos rfl] theorem find_parent_or (self : UnionFind) (x : Fin self.size) (i) : (self.find x).1.parent i = self.rootD i ∧ self.rootD i = self.rootD x ∨ @@ -449,7 +468,8 @@ theorem setParentBump_rankD_lt {arr : Array UFNode} {x y : Fin arr.size} simp [hP, hR, -Array.get_eq_getElem] at *; split <;> rename_i h₁ <;> [simp [← h₁]; skip] <;> split <;> rename_i h₂ <;> intro h · simp [h₂] at h - · simp [rankD_eq]; split <;> rename_i h₃ + · simp only [rankD_eq, Array.get_eq_getElem] + split <;> rename_i h₃ · rw [← h₃]; apply Nat.lt_succ_self · exact Nat.lt_of_le_of_ne H h₃ · cases h₂.1 @@ -469,13 +489,16 @@ theorem setParent_rankD_lt {arr : Array UFNode} {x y : Fin arr.size} (by simp [rankD_set, Nat.ne_of_lt h, rankD_eq, -Array.get_eq_getElem]) @[simp] theorem linkAux_size : (linkAux self x y).size = self.size := by - simp [linkAux]; split <;> [rfl; split] <;> [skip; split] <;> simp + simp only [linkAux, Array.get_eq_getElem] + split <;> [rfl; split] <;> [skip; split] <;> simp /-- Link a union-find node to a root node. -/ def link (self : UnionFind) (x y : Fin self.size) (yroot : self.parent y = y) : UnionFind where arr := linkAux self.arr x y parentD_lt h := by - simp at *; simp [linkAux]; split <;> [skip; split <;> [skip; split]] + simp only [Array.data_length, linkAux_size] at * + simp only [linkAux, Array.get_eq_getElem] + split <;> [skip; split <;> [skip; split]] · exact self.parentD_lt h · rw [parentD_set]; split <;> [exact x.2; exact self.parentD_lt h] · rw [parentD_set]; split @@ -484,12 +507,21 @@ def link (self : UnionFind) (x y : Fin self.size) (yroot : self.parent y = y) : · rw [parentD_set]; split <;> [exact y.2; exact self.parentD_lt h] rankD_lt := by rw [parent, parentD_eq] at yroot - simp [linkAux]; split <;> [skip; split <;> [skip; split]] + simp only [linkAux, Array.get_eq_getElem, ne_eq] + split <;> [skip; split <;> [skip; split]] · exact self.rankD_lt · exact setParent_rankD_lt ‹_› self.rankD_lt - · refine setParentBump_rankD_lt (.inr yroot) (Nat.le_of_eq ‹_›) self.rankD_lt - (by simp [parentD_set]; rintro rfl; simp [*, parentD_eq]) fun {i} => ?_ - simp [rankD_set]; split <;> simp [*]; rintro rfl; simp [rankD_eq, *] + · refine setParentBump_rankD_lt (.inr yroot) (Nat.le_of_eq ‹_›) self.rankD_lt (by + simp only [parentD_set, ite_eq_right_iff] + rintro rfl + simp [*, parentD_eq]) fun {i} => ?_ + simp only [rankD_set, Fin.eta, Array.get_eq_getElem] + split + · simp_all + · simp_all only [Array.get_eq_getElem, Array.data_length, Nat.lt_irrefl, not_false_eq_true, + and_true, ite_false, ite_eq_right_iff] + rintro rfl + simp [rankD_eq, *] · exact setParent_rankD_lt (Nat.lt_of_le_of_ne (Nat.not_lt.1 ‹_›) ‹_›) self.rankD_lt @[inherit_doc link] From 7391b9f9b06eca6fcc36d211db143ba7f84dbe29 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 9 Jul 2024 09:46:20 +1000 Subject: [PATCH 024/177] chore: add missing since dates to deprecations (#875) --- Batteries/Data/Array/Merge.lean | 14 ++++---- Batteries/Data/BitVec/Lemmas.lean | 9 ++--- Batteries/Data/Bool.lean | 6 ++-- Batteries/Data/Int/Order.lean | 2 +- Batteries/Data/List/Basic.lean | 10 +++--- Batteries/Data/List/Count.lean | 4 +-- Batteries/Data/List/Lemmas.lean | 6 ++-- Batteries/Data/List/Pairwise.lean | 2 +- Batteries/Data/Nat/Lemmas.lean | 16 +++++---- Batteries/Data/Option/Lemmas.lean | 4 +-- Batteries/Data/RBMap/Basic.lean | 12 +++---- Batteries/Lean/Delaborator.lean | 2 +- Batteries/Logic.lean | 2 +- Batteries/StdDeprecations.lean | 59 ++++++++++++++++--------------- 14 files changed, 74 insertions(+), 74 deletions(-) diff --git a/Batteries/Data/Array/Merge.lean b/Batteries/Data/Array/Merge.lean index 1486520565..aa83e0200b 100644 --- a/Batteries/Data/Array/Merge.lean +++ b/Batteries/Data/Array/Merge.lean @@ -27,7 +27,7 @@ where termination_by xs.size + ys.size - (i + j) set_option linter.unusedVariables false in -@[deprecated merge, inherit_doc merge] +@[deprecated merge (since := "2024-04-24"), inherit_doc merge] def mergeSortedPreservingDuplicates [ord : Ord α] (xs ys : Array α) : Array α := merge (compare · · |>.isLT) xs ys @@ -55,7 +55,7 @@ where | .eq => go (acc.push (merge x y)) (i + 1) (j + 1) termination_by xs.size + ys.size - (i + j) -@[deprecated] alias mergeSortedMergingDuplicates := mergeDedupWith +@[deprecated (since := "2024-04-24")] alias mergeSortedMergingDuplicates := mergeDedupWith /-- `O(|xs| + |ys|)`. Merge arrays `xs` and `ys`, which must be sorted according to `compare` and must @@ -64,7 +64,7 @@ not contain duplicates. If an element appears in both `xs` and `ys`, only one co @[inline] def mergeDedup [ord : Ord α] (xs ys : Array α) : Array α := mergeDedupWith (ord := ord) xs ys fun x _ => x -@[deprecated] alias mergeSortedDeduplicating := mergeDedup +@[deprecated (since := "2024-04-24")] alias mergeSortedDeduplicating := mergeDedup set_option linter.unusedVariables false in /-- @@ -83,7 +83,7 @@ where ys.foldl (init := xs) fun xs y => if xs.any (· == y) (stop := xsSize) then xs else xs.push y -@[deprecated] alias mergeUnsortedDeduplicating := mergeUnsortedDedup +@[deprecated (since := "2024-04-24")] alias mergeUnsortedDeduplicating := mergeUnsortedDedup /-- `O(|xs|)`. Replace each run `[x₁, ⋯, xₙ]` of equal elements in `xs` with @@ -101,7 +101,7 @@ where acc.push hd termination_by xs.size - i -@[deprecated] alias mergeAdjacentDuplicates := mergeAdjacentDups +@[deprecated (since := "2024-04-24")] alias mergeAdjacentDuplicates := mergeAdjacentDups /-- `O(|xs|)`. Deduplicate a sorted array. The array must be sorted with to an order which agrees with @@ -110,13 +110,13 @@ where def dedupSorted [eq : BEq α] (xs : Array α) : Array α := xs.mergeAdjacentDups (eq := eq) fun x _ => x -@[deprecated] alias deduplicateSorted := dedupSorted +@[deprecated (since := "2024-04-24")] alias deduplicateSorted := dedupSorted /-- `O(|xs| log |xs|)`. Sort and deduplicate an array. -/ def sortDedup [ord : Ord α] (xs : Array α) : Array α := have := ord.toBEq dedupSorted <| xs.qsort (compare · · |>.isLT) -@[deprecated] alias sortAndDeduplicate := sortDedup +@[deprecated (since := "2024-04-24")] alias sortAndDeduplicate := sortDedup end Array diff --git a/Batteries/Data/BitVec/Lemmas.lean b/Batteries/Data/BitVec/Lemmas.lean index e0e45b8cb0..2bcd3ce4aa 100644 --- a/Batteries/Data/BitVec/Lemmas.lean +++ b/Batteries/Data/BitVec/Lemmas.lean @@ -7,13 +7,10 @@ import Batteries.Tactic.Alias namespace BitVec -/-- Replaced 2024-02-07. -/ -@[deprecated] alias zero_is_unique := eq_nil +@[deprecated (since := "2024-02-07")] alias zero_is_unique := eq_nil /-! ### sub/neg -/ -/-- Replaced 2024-02-06. -/ -@[deprecated] alias sub_toNat := toNat_sub +@[deprecated (since := "2024-02-07")] alias sub_toNat := toNat_sub -/-- Replaced 2024-02-06. -/ -@[deprecated] alias neg_toNat := toNat_neg +@[deprecated (since := "2024-02-07")] alias neg_toNat := toNat_neg diff --git a/Batteries/Data/Bool.lean b/Batteries/Data/Bool.lean index ab48cbc578..b6c0c9b4e7 100644 --- a/Batteries/Data/Bool.lean +++ b/Batteries/Data/Bool.lean @@ -10,10 +10,10 @@ namespace Bool /-! ### injectivity lemmas -/ -@[deprecated] alias not_inj' := not_inj_iff +@[deprecated (since := "2023-10-27")] alias not_inj' := not_inj_iff -@[deprecated] alias and_or_inj_right' := and_or_inj_right_iff +@[deprecated (since := "2023-10-27")] alias and_or_inj_right' := and_or_inj_right_iff -@[deprecated] alias and_or_inj_left' := and_or_inj_left_iff +@[deprecated (since := "2023-10-27")] alias and_or_inj_left' := and_or_inj_left_iff end Bool diff --git a/Batteries/Data/Int/Order.lean b/Batteries/Data/Int/Order.lean index 3855f19885..daab75c612 100644 --- a/Batteries/Data/Int/Order.lean +++ b/Batteries/Data/Int/Order.lean @@ -7,4 +7,4 @@ import Batteries.Tactic.Alias namespace Int -@[deprecated] alias ofNat_natAbs_eq_of_nonneg := natAbs_of_nonneg +@[deprecated (since := "2024-01-24")] alias ofNat_natAbs_eq_of_nonneg := natAbs_of_nonneg diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index a8c3fef85a..2cff94c670 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -93,9 +93,9 @@ drop_while (· != 1) [0, 1, 2, 3] = [1, 2, 3] /-- Returns the index of the first element equal to `a`, or the length of the list otherwise. -/ def indexOf [BEq α] (a : α) : List α → Nat := findIdx (· == a) -@[deprecated] alias removeNth := eraseIdx -@[deprecated] alias removeNthTR := eraseIdxTR -@[deprecated] alias removeNth_eq_removeNthTR := eraseIdx_eq_eraseIdxTR +@[deprecated (since := "2024-05-06")] alias removeNth := eraseIdx +@[deprecated (since := "2024-05-06")] alias removeNthTR := eraseIdxTR +@[deprecated (since := "2024-05-06")] alias removeNth_eq_removeNthTR := eraseIdx_eq_eraseIdxTR /-- Replaces the first element of the list for which `f` returns `some` with the returned value. -/ @[simp] def replaceF (f : α → Option α) : List α → List α @@ -926,7 +926,7 @@ def range' : (start len : Nat) → (step : Nat := 1) → List Nat `ilast' x xs` returns the last element of `xs` if `xs` is non-empty; it returns `x` otherwise. Use `List.getLastD` instead. -/ -@[simp, deprecated getLastD] def ilast' {α} : α → List α → α +@[simp, deprecated getLastD (since := "2024-01-09")] def ilast' {α} : α → List α → α | a, [] => a | _, b :: l => ilast' b l @@ -934,7 +934,7 @@ Use `List.getLastD` instead. `last' xs` returns the last element of `xs` if `xs` is non-empty; it returns `none` otherwise. Use `List.getLast?` instead -/ -@[simp, deprecated getLast?] def last' {α} : List α → Option α +@[simp, deprecated getLast? (since := "2024-01-09")] def last' {α} : List α → Option α | [] => none | [a] => some a | _ :: l => last' l diff --git a/Batteries/Data/List/Count.lean b/Batteries/Data/List/Count.lean index 81548ff437..dfd1481052 100644 --- a/Batteries/Data/List/Count.lean +++ b/Batteries/Data/List/Count.lean @@ -196,11 +196,11 @@ theorem filter_beq (l : List α) (a : α) : l.filter (· == a) = replicate (coun theorem filter_eq (l : List α) (a : α) : l.filter (· = a) = replicate (count a l) a := filter_beq l a -@[deprecated filter_eq] +@[deprecated filter_eq (since := "2023-12-14")] theorem filter_eq' (l : List α) (a : α) : l.filter (a = ·) = replicate (count a l) a := by simpa only [eq_comm] using filter_eq l a -@[deprecated filter_beq] +@[deprecated filter_beq (since := "2023-12-14")] theorem filter_beq' (l : List α) (a : α) : l.filter (a == ·) = replicate (count a l) a := by simpa only [eq_comm (b := a)] using filter_eq l a diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 44d9a5c18d..1cf017c534 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -308,7 +308,7 @@ theorem eraseIdx_eq_modifyNthTail : ∀ n (l : List α), eraseIdx l n = modifyNt | n+1, [] => rfl | n+1, a :: l => congrArg (cons _) (eraseIdx_eq_modifyNthTail _ _) -@[deprecated] alias removeNth_eq_nth_tail := eraseIdx_eq_modifyNthTail +@[deprecated (since := "2024-05-06")] alias removeNth_eq_nth_tail := eraseIdx_eq_modifyNthTail theorem getElem?_modifyNth (f : α → α) : ∀ n (l : List α) m, (modifyNth f n l)[m]? = (fun a => if n = m then f a else a) <$> l[m]? @@ -472,7 +472,7 @@ theorem length_eraseIdx : ∀ {l i}, i < length l → length (@eraseIdx α l i) simp [eraseIdx, ← Nat.add_one] rw [length_eraseIdx this, Nat.sub_add_cancel (Nat.lt_of_le_of_lt (Nat.zero_le _) this)] -@[deprecated] alias length_removeNth := length_eraseIdx +@[deprecated (since := "2024-05-06")] alias length_removeNth := length_eraseIdx /-! ### tail -/ @@ -619,7 +619,7 @@ theorem erase_subset (a : α) (l : List α) : l.erase a ⊆ l := (erase_sublist theorem Sublist.erase (a : α) {l₁ l₂ : List α} (h : l₁ <+ l₂) : l₁.erase a <+ l₂.erase a := by simp only [erase_eq_eraseP']; exact h.eraseP -@[deprecated] alias sublist.erase := Sublist.erase +@[deprecated (since := "2024-04-22")] alias sublist.erase := Sublist.erase theorem mem_of_mem_erase {a b : α} {l : List α} (h : a ∈ l.erase b) : a ∈ l := erase_subset _ _ h diff --git a/Batteries/Data/List/Pairwise.lean b/Batteries/Data/List/Pairwise.lean index 8583a41f2b..83b935f8ae 100644 --- a/Batteries/Data/List/Pairwise.lean +++ b/Batteries/Data/List/Pairwise.lean @@ -188,7 +188,7 @@ theorem pairwise_iff_forall_sublist : l.Pairwise R ↔ (∀ {a b}, [a,b] <+ l intro a b hab apply h; exact hab.cons _ -@[deprecated pairwise_iff_forall_sublist] +@[deprecated pairwise_iff_forall_sublist (since := "2023-09-18")] theorem pairwise_of_reflexive_on_dupl_of_forall_ne [DecidableEq α] {l : List α} {r : α → α → Prop} (hr : ∀ a, 1 < count a l → r a a) (h : ∀ a ∈ l, ∀ b ∈ l, a ≠ b → r a b) : l.Pairwise r := by apply pairwise_iff_forall_sublist.mpr diff --git a/Batteries/Data/Nat/Lemmas.lean b/Batteries/Data/Nat/Lemmas.lean index 951f330e51..842ca09e95 100644 --- a/Batteries/Data/Nat/Lemmas.lean +++ b/Batteries/Data/Nat/Lemmas.lean @@ -151,19 +151,21 @@ protected def sum_trichotomy (a b : Nat) : a < b ⊕' a = b ⊕' b < a := /-! ## add -/ -@[deprecated] alias succ_add_eq_succ_add := Nat.succ_add_eq_add_succ +@[deprecated (since := "2023-11-25")] alias succ_add_eq_succ_add := Nat.succ_add_eq_add_succ /-! ## sub -/ -@[deprecated] protected alias le_of_le_of_sub_le_sub_right := Nat.le_of_sub_le_sub_right +@[deprecated (since := "2023-11-25")] +protected alias le_of_le_of_sub_le_sub_right := Nat.le_of_sub_le_sub_right -@[deprecated] protected alias le_of_le_of_sub_le_sub_left := Nat.le_of_sub_le_sub_left +@[deprecated (since := "2023-11-25")] +protected alias le_of_le_of_sub_le_sub_left := Nat.le_of_sub_le_sub_left /-! ### mul -/ -@[deprecated] protected alias mul_lt_mul := Nat.mul_lt_mul_of_lt_of_le' +@[deprecated (since := "2024-01-11")] protected alias mul_lt_mul := Nat.mul_lt_mul_of_lt_of_le' -@[deprecated] protected alias mul_lt_mul' := Nat.mul_lt_mul_of_le_of_lt +@[deprecated (since := "2024-01-11")] protected alias mul_lt_mul' := Nat.mul_lt_mul_of_le_of_lt /-! ### div/mod -/ @@ -182,5 +184,5 @@ protected def sum_trichotomy (a b : Nat) : a < b ⊕' a = b ⊕' b < a := @[simp] theorem sum_append : Nat.sum (l₁ ++ l₂) = Nat.sum l₁ + Nat.sum l₂ := by induction l₁ <;> simp [*, Nat.add_assoc] -@[deprecated] protected alias lt_connex := Nat.lt_or_gt_of_ne -@[deprecated] alias pow_two_pos := Nat.two_pow_pos -- deprecated 2024-02-09 +@[deprecated (since := "2024-03-05")] protected alias lt_connex := Nat.lt_or_gt_of_ne +@[deprecated (since := "2024-02-09")] alias pow_two_pos := Nat.two_pow_pos diff --git a/Batteries/Data/Option/Lemmas.lean b/Batteries/Data/Option/Lemmas.lean index a768a776ce..05a38b25a7 100644 --- a/Batteries/Data/Option/Lemmas.lean +++ b/Batteries/Data/Option/Lemmas.lean @@ -7,7 +7,7 @@ import Batteries.Tactic.Alias namespace Option -@[deprecated] alias to_list_some := toList_some -@[deprecated] alias to_list_none := toList_none +@[deprecated (since := "2024-03-05")] alias to_list_some := toList_some +@[deprecated (since := "2024-03-05")] alias to_list_none := toList_none end Option diff --git a/Batteries/Data/RBMap/Basic.lean b/Batteries/Data/RBMap/Basic.lean index a9ba215ad9..d1aff2dc3b 100644 --- a/Batteries/Data/RBMap/Basic.lean +++ b/Batteries/Data/RBMap/Basic.lean @@ -67,8 +67,8 @@ protected def max? : RBNode α → Option α | node _ _ v nil => some v | node _ _ _ r => r.max? -@[deprecated] protected alias min := RBNode.min? -@[deprecated] protected alias max := RBNode.max? +@[deprecated (since := "2024-04-17")] protected alias min := RBNode.min? +@[deprecated (since := "2024-04-17")] protected alias max := RBNode.max? /-- Fold a function in tree order along the nodes. `v₀` is used at `nil` nodes and @@ -669,8 +669,8 @@ instance : ToStream (RBSet α cmp) (RBNode.Stream α) := ⟨fun x => x.1.toStrea /-- `O(log n)`. Returns the entry `a` such that `a ≥ k` for all keys in the RBSet. -/ @[inline] protected def max? (t : RBSet α cmp) : Option α := t.1.max? -@[deprecated] protected alias min := RBSet.min? -@[deprecated] protected alias max := RBSet.max? +@[deprecated (since := "2024-04-17")] protected alias min := RBSet.min? +@[deprecated (since := "2024-04-17")] protected alias max := RBSet.max? instance [Repr α] : Repr (RBSet α cmp) where reprPrec m prec := Repr.addAppParen ("RBSet.ofList " ++ repr m.toList) prec @@ -1053,8 +1053,8 @@ instance : Stream (Values.Stream α β) β := ⟨Values.Stream.next?⟩ /-- `O(log n)`. Returns the key-value pair `(a, b)` such that `a ≥ k` for all keys in the RBMap. -/ @[inline] protected def max? : RBMap α β cmp → Option (α × β) := RBSet.max? -@[deprecated] protected alias min := RBMap.min? -@[deprecated] protected alias max := RBMap.max? +@[deprecated (since := "2024-04-17")] protected alias min := RBMap.min? +@[deprecated (since := "2024-04-17")] protected alias max := RBMap.max? instance [Repr α] [Repr β] : Repr (RBMap α β cmp) where reprPrec m prec := Repr.addAppParen ("RBMap.ofList " ++ repr m.toList) prec diff --git a/Batteries/Lean/Delaborator.lean b/Batteries/Lean/Delaborator.lean index d722eef9d8..e09f2198de 100644 --- a/Batteries/Lean/Delaborator.lean +++ b/Batteries/Lean/Delaborator.lean @@ -8,6 +8,6 @@ import Lean.PrettyPrinter open Lean PrettyPrinter Delaborator SubExpr /-- Abbreviation for `Lean.MessageData.ofConst`. -/ -@[deprecated Lean.MessageData.ofConst] +@[deprecated Lean.MessageData.ofConst (since := "2024-05-18")] def Lean.ppConst (e : Expr) : MessageData := Lean.MessageData.ofConst e diff --git a/Batteries/Logic.lean b/Batteries/Logic.lean index 10b5bd2fcb..d0fd84e688 100644 --- a/Batteries/Logic.lean +++ b/Batteries/Logic.lean @@ -10,7 +10,7 @@ import Batteries.Tactic.Lint.Misc instance {f : α → β} [DecidablePred p] : DecidablePred (p ∘ f) := inferInstanceAs <| DecidablePred fun x => p (f x) -@[deprecated] alias proofIrrel := proof_irrel +@[deprecated (since := "2024-03-15")] alias proofIrrel := proof_irrel /-! ## id -/ diff --git a/Batteries/StdDeprecations.lean b/Batteries/StdDeprecations.lean index 49357a0573..c705ee839d 100644 --- a/Batteries/StdDeprecations.lean +++ b/Batteries/StdDeprecations.lean @@ -21,35 +21,36 @@ but it would be much harder to generate the deprecations. Let's hope that people using the tactic implementations can work this out themselves. -/ -@[deprecated] alias Std.AssocList := Batteries.AssocList -@[deprecated] alias Std.HashMap := Batteries.HashMap -@[deprecated] alias Std.mkHashMap := Batteries.mkHashMap -@[deprecated] alias Std.DList := Batteries.DList -@[deprecated] alias Std.PairingHeapImp.Heap := Batteries.PairingHeapImp.Heap -@[deprecated] alias Std.TotalBLE := Batteries.TotalBLE -@[deprecated] alias Std.OrientedCmp := Batteries.OrientedCmp -@[deprecated] alias Std.TransCmp := Batteries.TransCmp -@[deprecated] alias Std.BEqCmp := Batteries.BEqCmp -@[deprecated] alias Std.LTCmp := Batteries.LTCmp -@[deprecated] alias Std.LECmp := Batteries.LECmp -@[deprecated] alias Std.LawfulCmp := Batteries.LawfulCmp -@[deprecated] alias Std.OrientedOrd := Batteries.OrientedOrd -@[deprecated] alias Std.TransOrd := Batteries.TransOrd -@[deprecated] alias Std.BEqOrd := Batteries.BEqOrd -@[deprecated] alias Std.LTOrd := Batteries.LTOrd -@[deprecated] alias Std.LEOrd := Batteries.LEOrd -@[deprecated] alias Std.LawfulOrd := Batteries.LawfulOrd -@[deprecated] alias Std.compareOfLessAndEq_eq_lt := Batteries.compareOfLessAndEq_eq_lt -@[deprecated] alias Std.RBColor := Batteries.RBColor -@[deprecated] alias Std.RBNode := Batteries.RBNode -@[deprecated] alias Std.RBSet := Batteries.RBSet -@[deprecated] alias Std.mkRBSet := Batteries.mkRBSet -@[deprecated] alias Std.RBMap := Batteries.RBMap -@[deprecated] alias Std.mkRBMap := Batteries.mkRBMap -@[deprecated] alias Std.BinomialHeap := Batteries.BinomialHeap -@[deprecated] alias Std.mkBinomialHeap := Batteries.mkBinomialHeap -@[deprecated] alias Std.UFNode := Batteries.UFNode -@[deprecated] alias Std.UnionFind := Batteries.UnionFind +@[deprecated (since := "2024-05-07")] alias Std.AssocList := Batteries.AssocList +@[deprecated (since := "2024-05-07")] alias Std.HashMap := Batteries.HashMap +@[deprecated (since := "2024-05-07")] alias Std.mkHashMap := Batteries.mkHashMap +@[deprecated (since := "2024-05-07")] alias Std.DList := Batteries.DList +@[deprecated (since := "2024-05-07")] alias Std.PairingHeapImp.Heap := Batteries.PairingHeapImp.Heap +@[deprecated (since := "2024-05-07")] alias Std.TotalBLE := Batteries.TotalBLE +@[deprecated (since := "2024-05-07")] alias Std.OrientedCmp := Batteries.OrientedCmp +@[deprecated (since := "2024-05-07")] alias Std.TransCmp := Batteries.TransCmp +@[deprecated (since := "2024-05-07")] alias Std.BEqCmp := Batteries.BEqCmp +@[deprecated (since := "2024-05-07")] alias Std.LTCmp := Batteries.LTCmp +@[deprecated (since := "2024-05-07")] alias Std.LECmp := Batteries.LECmp +@[deprecated (since := "2024-05-07")] alias Std.LawfulCmp := Batteries.LawfulCmp +@[deprecated (since := "2024-05-07")] alias Std.OrientedOrd := Batteries.OrientedOrd +@[deprecated (since := "2024-05-07")] alias Std.TransOrd := Batteries.TransOrd +@[deprecated (since := "2024-05-07")] alias Std.BEqOrd := Batteries.BEqOrd +@[deprecated (since := "2024-05-07")] alias Std.LTOrd := Batteries.LTOrd +@[deprecated (since := "2024-05-07")] alias Std.LEOrd := Batteries.LEOrd +@[deprecated (since := "2024-05-07")] alias Std.LawfulOrd := Batteries.LawfulOrd +@[deprecated (since := "2024-05-07")] +alias Std.compareOfLessAndEq_eq_lt := Batteries.compareOfLessAndEq_eq_lt +@[deprecated (since := "2024-05-07")] alias Std.RBColor := Batteries.RBColor +@[deprecated (since := "2024-05-07")] alias Std.RBNode := Batteries.RBNode +@[deprecated (since := "2024-05-07")] alias Std.RBSet := Batteries.RBSet +@[deprecated (since := "2024-05-07")] alias Std.mkRBSet := Batteries.mkRBSet +@[deprecated (since := "2024-05-07")] alias Std.RBMap := Batteries.RBMap +@[deprecated (since := "2024-05-07")] alias Std.mkRBMap := Batteries.mkRBMap +@[deprecated (since := "2024-05-07")] alias Std.BinomialHeap := Batteries.BinomialHeap +@[deprecated (since := "2024-05-07")] alias Std.mkBinomialHeap := Batteries.mkBinomialHeap +@[deprecated (since := "2024-05-07")] alias Std.UFNode := Batteries.UFNode +@[deprecated (since := "2024-05-07")] alias Std.UnionFind := Batteries.UnionFind -- Check that these generate usable deprecated hints -- when referring to names inside these namespaces. From c0efc1fd2a0bec51bd55c5b17348af13d7419239 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 10 Jul 2024 02:30:09 +1000 Subject: [PATCH 025/177] chore: bump toolchain to v4.10.0-rc2 (#876) --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index d69d1ed200..6a4259fa56 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.10.0-rc1 +leanprover/lean4:v4.10.0-rc2 From c65cf2ed7aa172d4553db080506ba1242c2ccdd2 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 15 Jul 2024 20:18:51 +0100 Subject: [PATCH 026/177] chore: upstream most LazyList results from Mathlib (#835) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: François G. Dorais --- Batteries/Data/LazyList.lean | 159 ++++++++++++++++++++++++++++++++++- 1 file changed, 158 insertions(+), 1 deletion(-) diff --git a/Batteries/Data/LazyList.lean b/Batteries/Data/LazyList.lean index 9deb41e77b..48edafaf45 100644 --- a/Batteries/Data/LazyList.lean +++ b/Batteries/Data/LazyList.lean @@ -1,8 +1,9 @@ /- Copyright (c) 2017 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Leonardo de Moura +Authors: Leonardo de Moura, Simon Hudon -/ +import Batteries.Data.Thunk /-! # Lazy lists @@ -117,4 +118,160 @@ partial def iterates (f : α → α) : α → LazyList α partial def iota (i : Nat) : LazyList Nat := iterates Nat.succ i + +instance decidableEq {α : Type u} [DecidableEq α] : DecidableEq (LazyList α) + | nil, nil => isTrue rfl + | cons x xs, cons y ys => + if h : x = y then + match decidableEq xs.get ys.get with + | isFalse h2 => by + apply isFalse; simp only [cons.injEq, not_and]; intro _ xs_ys; apply h2; rw [xs_ys] + | isTrue h2 => by apply isTrue; congr; ext; exact h2 + else by apply isFalse; simp only [cons.injEq, not_and]; intro; contradiction + | nil, cons _ _ => by apply isFalse; simp + | cons _ _, nil => by apply isFalse; simp + +/-- Traversal of lazy lists using an applicative effect. -/ +protected def traverse {m : Type u → Type u} [Applicative m] {α β : Type u} (f : α → m β) : + LazyList α → m (LazyList β) + | LazyList.nil => pure LazyList.nil + | LazyList.cons x xs => LazyList.cons <$> f x <*> Thunk.pure <$> xs.get.traverse f + +/-- `init xs`, if `xs` non-empty, drops the last element of the list. +Otherwise, return the empty list. -/ +def init {α} : LazyList α → LazyList α + | LazyList.nil => LazyList.nil + | LazyList.cons x xs => + let xs' := xs.get + match xs' with + | LazyList.nil => LazyList.nil + | LazyList.cons _ _ => LazyList.cons x (init xs') + +/-- Return the first object contained in the list that satisfies +predicate `p` -/ +def find {α} (p : α → Prop) [DecidablePred p] : LazyList α → Option α + | nil => none + | cons h t => if p h then some h else t.get.find p + +/-- `interleave xs ys` creates a list where elements of `xs` and `ys` alternate. -/ +def interleave {α} : LazyList α → LazyList α → LazyList α + | LazyList.nil, xs => xs + | a@(LazyList.cons _ _), LazyList.nil => a + | LazyList.cons x xs, LazyList.cons y ys => + LazyList.cons x (LazyList.cons y (interleave xs.get ys.get)) + +/-- `interleaveAll (xs::ys::zs::xss)` creates a list where elements of `xs`, `ys` +and `zs` and the rest alternate. Every other element of the resulting list is taken from +`xs`, every fourth is taken from `ys`, every eighth is taken from `zs` and so on. -/ +def interleaveAll {α} : List (LazyList α) → LazyList α + | [] => LazyList.nil + | x :: xs => interleave x (interleaveAll xs) + +/-- Monadic bind operation for `LazyList`. -/ +protected def bind {α β} : LazyList α → (α → LazyList β) → LazyList β + | LazyList.nil, _ => LazyList.nil + | LazyList.cons x xs, f => (f x).append (xs.get.bind f) + +/-- Reverse the order of a `LazyList`. +It is done by converting to a `List` first because reversal involves evaluating all +the list and if the list is all evaluated, `List` is a better representation for +it than a series of thunks. -/ +def reverse {α} (xs : LazyList α) : LazyList α := + ofList xs.toList.reverse + +instance : Monad LazyList where + pure := @LazyList.singleton + bind := @LazyList.bind + +-- Porting note: Added `Thunk.pure` to definition. +theorem append_nil {α} (xs : LazyList α) : xs.append (Thunk.pure LazyList.nil) = xs := by + induction xs using LazyList.rec with + | nil => simp only [Thunk.pure, append, Thunk.get] + | cons _ _ => simpa only [append, cons.injEq, true_and] + | mk _ ih => ext; apply ih + +theorem append_assoc {α} (xs ys zs : LazyList α) : + (xs.append ys).append zs = xs.append (ys.append zs) := by + induction xs using LazyList.rec with + | nil => simp only [append, Thunk.get] + | cons _ _ => simpa only [append, cons.injEq, true_and] + | mk _ ih => ext; apply ih + +-- Porting note: Rewrote proof of `append_bind`. +theorem append_bind {α β} (xs : LazyList α) (ys : Thunk (LazyList α)) (f : α → LazyList β) : + (xs.append ys).bind f = (xs.bind f).append (ys.get.bind f) := by + match xs with + | LazyList.nil => + simp only [append, Thunk.get, LazyList.bind] + | LazyList.cons x xs => + simp only [append, Thunk.get, LazyList.bind] + have := append_bind xs.get ys f + simp only [Thunk.get] at this + rw [this, append_assoc] + +-- TODO: upstream from mathlib +-- instance : LawfulMonad LazyList.{u} + +/-- Try applying function `f` to every element of a `LazyList` and +return the result of the first attempt that succeeds. -/ +def mfirst {m} [Alternative m] {α β} (f : α → m β) : LazyList α → m β + | nil => failure + | cons x xs => f x <|> xs.get.mfirst f + +/-- Membership in lazy lists -/ +protected def Mem {α} (x : α) : LazyList α → Prop + | nil => False + | cons y ys => x = y ∨ ys.get.Mem x + +instance {α} : Membership α (LazyList α) := + ⟨LazyList.Mem⟩ + +instance Mem.decidable {α} [DecidableEq α] (x : α) : ∀ xs : LazyList α, Decidable (x ∈ xs) + | LazyList.nil => by + apply Decidable.isFalse + simp [Membership.mem, LazyList.Mem] + | LazyList.cons y ys => + if h : x = y then by + apply Decidable.isTrue + simp only [Membership.mem, LazyList.Mem] + exact Or.inl h + else by + have := Mem.decidable x ys.get + have : (x ∈ ys.get) ↔ (x ∈ cons y ys) := by simp [(· ∈ ·), LazyList.Mem, h] + exact decidable_of_decidable_of_iff this + +@[simp] +theorem mem_nil {α} (x : α) : x ∈ @LazyList.nil α ↔ False := by + simp [Membership.mem, LazyList.Mem] + +@[simp] +theorem mem_cons {α} (x y : α) (ys : Thunk (LazyList α)) : + x ∈ @LazyList.cons α y ys ↔ x = y ∨ x ∈ ys.get := by + simp [Membership.mem, LazyList.Mem] + +theorem forall_mem_cons {α} {p : α → Prop} {a : α} {l : Thunk (LazyList α)} : + (∀ x ∈ @LazyList.cons _ a l, p x) ↔ p a ∧ ∀ x ∈ l.get, p x := by + simp only [Membership.mem, LazyList.Mem, or_imp, forall_and, forall_eq] + +/-! ### map for partial functions -/ + + +/-- Partial map. If `f : ∀ a, p a → β` is a partial function defined on + `a : α` satisfying `p`, then `pmap f l h` is essentially the same as `map f l` + but is defined only when all members of `l` satisfy `p`, using the proof + to apply `f`. -/ +@[simp] +def pmap {α β} {p : α → Prop} (f : ∀ a, p a → β) : ∀ l : LazyList α, (∀ a ∈ l, p a) → LazyList β + | LazyList.nil, _ => LazyList.nil + | LazyList.cons x xs, H => + LazyList.cons (f x (forall_mem_cons.1 H).1) (xs.get.pmap f (forall_mem_cons.1 H).2) + +/-- "Attach" the proof that the elements of `l` are in `l` to produce a new `LazyList` + with the same elements but in the type `{x // x ∈ l}`. -/ +def attach {α} (l : LazyList α) : LazyList { x // x ∈ l } := + pmap Subtype.mk l fun _ => id + +instance {α} [Repr α] : Repr (LazyList α) := + ⟨fun xs _ => repr xs.toList⟩ + end LazyList From 8011496fd710cc4c6fffe7438416a0fe59649bd0 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 15 Jul 2024 14:35:53 -0500 Subject: [PATCH 027/177] Revert "chore: upstream most LazyList results from Mathlib" (#882) --- Batteries/Data/LazyList.lean | 159 +---------------------------------- 1 file changed, 1 insertion(+), 158 deletions(-) diff --git a/Batteries/Data/LazyList.lean b/Batteries/Data/LazyList.lean index 48edafaf45..9deb41e77b 100644 --- a/Batteries/Data/LazyList.lean +++ b/Batteries/Data/LazyList.lean @@ -1,9 +1,8 @@ /- Copyright (c) 2017 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Leonardo de Moura, Simon Hudon +Authors: Leonardo de Moura -/ -import Batteries.Data.Thunk /-! # Lazy lists @@ -118,160 +117,4 @@ partial def iterates (f : α → α) : α → LazyList α partial def iota (i : Nat) : LazyList Nat := iterates Nat.succ i - -instance decidableEq {α : Type u} [DecidableEq α] : DecidableEq (LazyList α) - | nil, nil => isTrue rfl - | cons x xs, cons y ys => - if h : x = y then - match decidableEq xs.get ys.get with - | isFalse h2 => by - apply isFalse; simp only [cons.injEq, not_and]; intro _ xs_ys; apply h2; rw [xs_ys] - | isTrue h2 => by apply isTrue; congr; ext; exact h2 - else by apply isFalse; simp only [cons.injEq, not_and]; intro; contradiction - | nil, cons _ _ => by apply isFalse; simp - | cons _ _, nil => by apply isFalse; simp - -/-- Traversal of lazy lists using an applicative effect. -/ -protected def traverse {m : Type u → Type u} [Applicative m] {α β : Type u} (f : α → m β) : - LazyList α → m (LazyList β) - | LazyList.nil => pure LazyList.nil - | LazyList.cons x xs => LazyList.cons <$> f x <*> Thunk.pure <$> xs.get.traverse f - -/-- `init xs`, if `xs` non-empty, drops the last element of the list. -Otherwise, return the empty list. -/ -def init {α} : LazyList α → LazyList α - | LazyList.nil => LazyList.nil - | LazyList.cons x xs => - let xs' := xs.get - match xs' with - | LazyList.nil => LazyList.nil - | LazyList.cons _ _ => LazyList.cons x (init xs') - -/-- Return the first object contained in the list that satisfies -predicate `p` -/ -def find {α} (p : α → Prop) [DecidablePred p] : LazyList α → Option α - | nil => none - | cons h t => if p h then some h else t.get.find p - -/-- `interleave xs ys` creates a list where elements of `xs` and `ys` alternate. -/ -def interleave {α} : LazyList α → LazyList α → LazyList α - | LazyList.nil, xs => xs - | a@(LazyList.cons _ _), LazyList.nil => a - | LazyList.cons x xs, LazyList.cons y ys => - LazyList.cons x (LazyList.cons y (interleave xs.get ys.get)) - -/-- `interleaveAll (xs::ys::zs::xss)` creates a list where elements of `xs`, `ys` -and `zs` and the rest alternate. Every other element of the resulting list is taken from -`xs`, every fourth is taken from `ys`, every eighth is taken from `zs` and so on. -/ -def interleaveAll {α} : List (LazyList α) → LazyList α - | [] => LazyList.nil - | x :: xs => interleave x (interleaveAll xs) - -/-- Monadic bind operation for `LazyList`. -/ -protected def bind {α β} : LazyList α → (α → LazyList β) → LazyList β - | LazyList.nil, _ => LazyList.nil - | LazyList.cons x xs, f => (f x).append (xs.get.bind f) - -/-- Reverse the order of a `LazyList`. -It is done by converting to a `List` first because reversal involves evaluating all -the list and if the list is all evaluated, `List` is a better representation for -it than a series of thunks. -/ -def reverse {α} (xs : LazyList α) : LazyList α := - ofList xs.toList.reverse - -instance : Monad LazyList where - pure := @LazyList.singleton - bind := @LazyList.bind - --- Porting note: Added `Thunk.pure` to definition. -theorem append_nil {α} (xs : LazyList α) : xs.append (Thunk.pure LazyList.nil) = xs := by - induction xs using LazyList.rec with - | nil => simp only [Thunk.pure, append, Thunk.get] - | cons _ _ => simpa only [append, cons.injEq, true_and] - | mk _ ih => ext; apply ih - -theorem append_assoc {α} (xs ys zs : LazyList α) : - (xs.append ys).append zs = xs.append (ys.append zs) := by - induction xs using LazyList.rec with - | nil => simp only [append, Thunk.get] - | cons _ _ => simpa only [append, cons.injEq, true_and] - | mk _ ih => ext; apply ih - --- Porting note: Rewrote proof of `append_bind`. -theorem append_bind {α β} (xs : LazyList α) (ys : Thunk (LazyList α)) (f : α → LazyList β) : - (xs.append ys).bind f = (xs.bind f).append (ys.get.bind f) := by - match xs with - | LazyList.nil => - simp only [append, Thunk.get, LazyList.bind] - | LazyList.cons x xs => - simp only [append, Thunk.get, LazyList.bind] - have := append_bind xs.get ys f - simp only [Thunk.get] at this - rw [this, append_assoc] - --- TODO: upstream from mathlib --- instance : LawfulMonad LazyList.{u} - -/-- Try applying function `f` to every element of a `LazyList` and -return the result of the first attempt that succeeds. -/ -def mfirst {m} [Alternative m] {α β} (f : α → m β) : LazyList α → m β - | nil => failure - | cons x xs => f x <|> xs.get.mfirst f - -/-- Membership in lazy lists -/ -protected def Mem {α} (x : α) : LazyList α → Prop - | nil => False - | cons y ys => x = y ∨ ys.get.Mem x - -instance {α} : Membership α (LazyList α) := - ⟨LazyList.Mem⟩ - -instance Mem.decidable {α} [DecidableEq α] (x : α) : ∀ xs : LazyList α, Decidable (x ∈ xs) - | LazyList.nil => by - apply Decidable.isFalse - simp [Membership.mem, LazyList.Mem] - | LazyList.cons y ys => - if h : x = y then by - apply Decidable.isTrue - simp only [Membership.mem, LazyList.Mem] - exact Or.inl h - else by - have := Mem.decidable x ys.get - have : (x ∈ ys.get) ↔ (x ∈ cons y ys) := by simp [(· ∈ ·), LazyList.Mem, h] - exact decidable_of_decidable_of_iff this - -@[simp] -theorem mem_nil {α} (x : α) : x ∈ @LazyList.nil α ↔ False := by - simp [Membership.mem, LazyList.Mem] - -@[simp] -theorem mem_cons {α} (x y : α) (ys : Thunk (LazyList α)) : - x ∈ @LazyList.cons α y ys ↔ x = y ∨ x ∈ ys.get := by - simp [Membership.mem, LazyList.Mem] - -theorem forall_mem_cons {α} {p : α → Prop} {a : α} {l : Thunk (LazyList α)} : - (∀ x ∈ @LazyList.cons _ a l, p x) ↔ p a ∧ ∀ x ∈ l.get, p x := by - simp only [Membership.mem, LazyList.Mem, or_imp, forall_and, forall_eq] - -/-! ### map for partial functions -/ - - -/-- Partial map. If `f : ∀ a, p a → β` is a partial function defined on - `a : α` satisfying `p`, then `pmap f l h` is essentially the same as `map f l` - but is defined only when all members of `l` satisfy `p`, using the proof - to apply `f`. -/ -@[simp] -def pmap {α β} {p : α → Prop} (f : ∀ a, p a → β) : ∀ l : LazyList α, (∀ a ∈ l, p a) → LazyList β - | LazyList.nil, _ => LazyList.nil - | LazyList.cons x xs, H => - LazyList.cons (f x (forall_mem_cons.1 H).1) (xs.get.pmap f (forall_mem_cons.1 H).2) - -/-- "Attach" the proof that the elements of `l` are in `l` to produce a new `LazyList` - with the same elements but in the type `{x // x ∈ l}`. -/ -def attach {α} (l : LazyList α) : LazyList { x // x ∈ l } := - pmap Subtype.mk l fun _ => id - -instance {α} [Repr α] : Repr (LazyList α) := - ⟨fun xs _ => repr xs.toList⟩ - end LazyList From 937cd3219c0beffa7b623d2905707d1304da259e Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 15 Jul 2024 15:12:23 -0500 Subject: [PATCH 028/177] chore: upstream most LazyList results from Mathlib (attempt 2) (#883) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Eric Wieser Co-authored-by: François G. Dorais --- Batteries/Data/LazyList.lean | 159 ++++++++++++++++++++++++++++++++++- 1 file changed, 158 insertions(+), 1 deletion(-) diff --git a/Batteries/Data/LazyList.lean b/Batteries/Data/LazyList.lean index 9deb41e77b..48edafaf45 100644 --- a/Batteries/Data/LazyList.lean +++ b/Batteries/Data/LazyList.lean @@ -1,8 +1,9 @@ /- Copyright (c) 2017 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Leonardo de Moura +Authors: Leonardo de Moura, Simon Hudon -/ +import Batteries.Data.Thunk /-! # Lazy lists @@ -117,4 +118,160 @@ partial def iterates (f : α → α) : α → LazyList α partial def iota (i : Nat) : LazyList Nat := iterates Nat.succ i + +instance decidableEq {α : Type u} [DecidableEq α] : DecidableEq (LazyList α) + | nil, nil => isTrue rfl + | cons x xs, cons y ys => + if h : x = y then + match decidableEq xs.get ys.get with + | isFalse h2 => by + apply isFalse; simp only [cons.injEq, not_and]; intro _ xs_ys; apply h2; rw [xs_ys] + | isTrue h2 => by apply isTrue; congr; ext; exact h2 + else by apply isFalse; simp only [cons.injEq, not_and]; intro; contradiction + | nil, cons _ _ => by apply isFalse; simp + | cons _ _, nil => by apply isFalse; simp + +/-- Traversal of lazy lists using an applicative effect. -/ +protected def traverse {m : Type u → Type u} [Applicative m] {α β : Type u} (f : α → m β) : + LazyList α → m (LazyList β) + | LazyList.nil => pure LazyList.nil + | LazyList.cons x xs => LazyList.cons <$> f x <*> Thunk.pure <$> xs.get.traverse f + +/-- `init xs`, if `xs` non-empty, drops the last element of the list. +Otherwise, return the empty list. -/ +def init {α} : LazyList α → LazyList α + | LazyList.nil => LazyList.nil + | LazyList.cons x xs => + let xs' := xs.get + match xs' with + | LazyList.nil => LazyList.nil + | LazyList.cons _ _ => LazyList.cons x (init xs') + +/-- Return the first object contained in the list that satisfies +predicate `p` -/ +def find {α} (p : α → Prop) [DecidablePred p] : LazyList α → Option α + | nil => none + | cons h t => if p h then some h else t.get.find p + +/-- `interleave xs ys` creates a list where elements of `xs` and `ys` alternate. -/ +def interleave {α} : LazyList α → LazyList α → LazyList α + | LazyList.nil, xs => xs + | a@(LazyList.cons _ _), LazyList.nil => a + | LazyList.cons x xs, LazyList.cons y ys => + LazyList.cons x (LazyList.cons y (interleave xs.get ys.get)) + +/-- `interleaveAll (xs::ys::zs::xss)` creates a list where elements of `xs`, `ys` +and `zs` and the rest alternate. Every other element of the resulting list is taken from +`xs`, every fourth is taken from `ys`, every eighth is taken from `zs` and so on. -/ +def interleaveAll {α} : List (LazyList α) → LazyList α + | [] => LazyList.nil + | x :: xs => interleave x (interleaveAll xs) + +/-- Monadic bind operation for `LazyList`. -/ +protected def bind {α β} : LazyList α → (α → LazyList β) → LazyList β + | LazyList.nil, _ => LazyList.nil + | LazyList.cons x xs, f => (f x).append (xs.get.bind f) + +/-- Reverse the order of a `LazyList`. +It is done by converting to a `List` first because reversal involves evaluating all +the list and if the list is all evaluated, `List` is a better representation for +it than a series of thunks. -/ +def reverse {α} (xs : LazyList α) : LazyList α := + ofList xs.toList.reverse + +instance : Monad LazyList where + pure := @LazyList.singleton + bind := @LazyList.bind + +-- Porting note: Added `Thunk.pure` to definition. +theorem append_nil {α} (xs : LazyList α) : xs.append (Thunk.pure LazyList.nil) = xs := by + induction xs using LazyList.rec with + | nil => simp only [Thunk.pure, append, Thunk.get] + | cons _ _ => simpa only [append, cons.injEq, true_and] + | mk _ ih => ext; apply ih + +theorem append_assoc {α} (xs ys zs : LazyList α) : + (xs.append ys).append zs = xs.append (ys.append zs) := by + induction xs using LazyList.rec with + | nil => simp only [append, Thunk.get] + | cons _ _ => simpa only [append, cons.injEq, true_and] + | mk _ ih => ext; apply ih + +-- Porting note: Rewrote proof of `append_bind`. +theorem append_bind {α β} (xs : LazyList α) (ys : Thunk (LazyList α)) (f : α → LazyList β) : + (xs.append ys).bind f = (xs.bind f).append (ys.get.bind f) := by + match xs with + | LazyList.nil => + simp only [append, Thunk.get, LazyList.bind] + | LazyList.cons x xs => + simp only [append, Thunk.get, LazyList.bind] + have := append_bind xs.get ys f + simp only [Thunk.get] at this + rw [this, append_assoc] + +-- TODO: upstream from mathlib +-- instance : LawfulMonad LazyList.{u} + +/-- Try applying function `f` to every element of a `LazyList` and +return the result of the first attempt that succeeds. -/ +def mfirst {m} [Alternative m] {α β} (f : α → m β) : LazyList α → m β + | nil => failure + | cons x xs => f x <|> xs.get.mfirst f + +/-- Membership in lazy lists -/ +protected def Mem {α} (x : α) : LazyList α → Prop + | nil => False + | cons y ys => x = y ∨ ys.get.Mem x + +instance {α} : Membership α (LazyList α) := + ⟨LazyList.Mem⟩ + +instance Mem.decidable {α} [DecidableEq α] (x : α) : ∀ xs : LazyList α, Decidable (x ∈ xs) + | LazyList.nil => by + apply Decidable.isFalse + simp [Membership.mem, LazyList.Mem] + | LazyList.cons y ys => + if h : x = y then by + apply Decidable.isTrue + simp only [Membership.mem, LazyList.Mem] + exact Or.inl h + else by + have := Mem.decidable x ys.get + have : (x ∈ ys.get) ↔ (x ∈ cons y ys) := by simp [(· ∈ ·), LazyList.Mem, h] + exact decidable_of_decidable_of_iff this + +@[simp] +theorem mem_nil {α} (x : α) : x ∈ @LazyList.nil α ↔ False := by + simp [Membership.mem, LazyList.Mem] + +@[simp] +theorem mem_cons {α} (x y : α) (ys : Thunk (LazyList α)) : + x ∈ @LazyList.cons α y ys ↔ x = y ∨ x ∈ ys.get := by + simp [Membership.mem, LazyList.Mem] + +theorem forall_mem_cons {α} {p : α → Prop} {a : α} {l : Thunk (LazyList α)} : + (∀ x ∈ @LazyList.cons _ a l, p x) ↔ p a ∧ ∀ x ∈ l.get, p x := by + simp only [Membership.mem, LazyList.Mem, or_imp, forall_and, forall_eq] + +/-! ### map for partial functions -/ + + +/-- Partial map. If `f : ∀ a, p a → β` is a partial function defined on + `a : α` satisfying `p`, then `pmap f l h` is essentially the same as `map f l` + but is defined only when all members of `l` satisfy `p`, using the proof + to apply `f`. -/ +@[simp] +def pmap {α β} {p : α → Prop} (f : ∀ a, p a → β) : ∀ l : LazyList α, (∀ a ∈ l, p a) → LazyList β + | LazyList.nil, _ => LazyList.nil + | LazyList.cons x xs, H => + LazyList.cons (f x (forall_mem_cons.1 H).1) (xs.get.pmap f (forall_mem_cons.1 H).2) + +/-- "Attach" the proof that the elements of `l` are in `l` to produce a new `LazyList` + with the same elements but in the type `{x // x ∈ l}`. -/ +def attach {α} (l : LazyList α) : LazyList { x // x ∈ l } := + pmap Subtype.mk l fun _ => id + +instance {α} [Repr α] : Repr (LazyList α) := + ⟨fun xs _ => repr xs.toList⟩ + end LazyList From 8236a3610b3ed9d6161ce4d3239230499bb5dba0 Mon Sep 17 00:00:00 2001 From: grunweg Date: Wed, 17 Jul 2024 20:29:10 +0200 Subject: [PATCH 029/177] doc: fix outdated docstring (#880) --- Batteries/Tactic/Lint/Misc.lean | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/Batteries/Tactic/Lint/Misc.lean b/Batteries/Tactic/Lint/Misc.lean index 4bb3f2125b..3f0e69f648 100644 --- a/Batteries/Tactic/Lint/Misc.lean +++ b/Batteries/Tactic/Lint/Misc.lean @@ -239,9 +239,7 @@ def findUnusedHaves (e : Expr) : MetaM (Array MessageData) := do | _ => return res.get -/-- A linter for checking that declarations don't have unused term mode have statements. We do not -tag this as `@[env_linter]` so that it is not in the default linter set as it is slow and an -uncommon problem. -/ +/-- A linter for checking that declarations don't have unused term mode have statements. -/ @[env_linter] def unusedHavesSuffices : Linter where noErrorsFound := "No declarations have unused term mode have statements." errorsFound := "THE FOLLOWING DECLARATIONS HAVE INEFFECTUAL TERM MODE HAVE/SUFFICES BLOCKS. \ From 27b7b8814b9bc4634fcf3ace2388b171784f76e3 Mon Sep 17 00:00:00 2001 From: grunweg Date: Wed, 17 Jul 2024 20:29:43 +0200 Subject: [PATCH 030/177] feat: allow configuring additional linters to run (#881) --- Batteries/Tactic/Lint/Frontend.lean | 38 ++++++++++++++++++----------- scripts/runLinter.lean | 2 +- 2 files changed, 25 insertions(+), 15 deletions(-) diff --git a/Batteries/Tactic/Lint/Frontend.lean b/Batteries/Tactic/Lint/Frontend.lean index 072e23743f..841c696af8 100644 --- a/Batteries/Tactic/Lint/Frontend.lean +++ b/Batteries/Tactic/Lint/Frontend.lean @@ -65,20 +65,28 @@ inductive LintVerbosity | high deriving Inhabited, DecidableEq, Repr -/-- `getChecks slow extra use_only` produces a list of linters. -`extras` is a list of names that should resolve to declarations with type `linter`. -If `useOnly` is true, it only uses the linters in `extra`. -Otherwise, it uses all linters in the environment tagged with `@[env_linter]`. +/-- `getChecks slow runOnly runAlways` produces a list of linters. +`runOnly` is an optional list of names that should resolve to declarations with type `NamedLinter`. +If populated, only these linters are run (regardless of the default configuration). +`runAlways` is an optional list of names that should resolve to declarations with type +`NamedLinter`. If populated, these linters are always run (regardless of their configuration). +Specifying a linter in `runAlways` but not `runOnly` is an error. +Otherwise, it uses all enabled linters in the environment tagged with `@[env_linter]`. If `slow` is false, it only uses the fast default tests. -/ -def getChecks (slow : Bool) (useOnly : Bool) : CoreM (Array NamedLinter) := do +def getChecks (slow : Bool) (runOnly : Option (List Name)) (runAlways : Option (List Name)) : + CoreM (Array NamedLinter) := do let mut result := #[] - unless useOnly do - for (name, declName, dflt) in batteriesLinterExt.getState (← getEnv) do - if dflt then - let linter ← getLinter name declName - if slow || linter.isFast then - let _ := Inhabited.mk linter - result := result.binInsert (·.name.lt ·.name) linter + for (name, declName, default) in batteriesLinterExt.getState (← getEnv) do + let shouldRun := match (runOnly, runAlways) with + | (some only, some always) => only.contains name && (always.contains name || default) + | (some only, none) => only.contains name + | (none, some always) => default || always.contains name + | _ => default + if shouldRun then + let linter ← getLinter name declName + if slow || linter.isFast then + let _ := Inhabited.mk linter + result := result.binInsert (·.name.lt ·.name) linter pure result -- Note: we have to use the same context as `runTermElabM` here so that the `simpNF` @@ -238,9 +246,11 @@ elab tk:"#lint" verbosity:("+" <|> "-")? fast:"*"? only:(&" only")? | some ⟨.node _ `token.«-» _⟩ => pure .low | _ => throwUnsupportedSyntax let fast := fast.isSome - let only := only.isSome + let onlyNames : Option (List Name) := match only.isSome with + | true => some (linters.map fun l => l.getId).toList + | false => none let linters ← liftCoreM do - let mut result ← getChecks (slow := !fast) only + let mut result ← getChecks (slow := !fast) (runOnly := onlyNames) none let linterState := batteriesLinterExt.getState (← getEnv) for id in linters do let name := id.getId.eraseMacroScopes diff --git a/scripts/runLinter.lean b/scripts/runLinter.lean index 87d8774ce9..3d98a1c902 100644 --- a/scripts/runLinter.lean +++ b/scripts/runLinter.lean @@ -57,7 +57,7 @@ unsafe def main (args : List String) : IO Unit := do let state := { env } Prod.fst <$> (CoreM.toIO · ctx state) do let decls ← getDeclsInPackage module.getRoot - let linters ← getChecks (slow := true) (useOnly := false) + let linters ← getChecks (slow := true) (runAlways := none) (runOnly := none) let results ← lintCore decls linters if update then writeJsonFile (α := NoLints) nolintsFile <| From f27beb10b53350d6c1257ba3a8899df369135cc3 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 17 Jul 2024 13:30:10 -0500 Subject: [PATCH 031/177] chore: test file for linting lean4 (#871) --- test/lint_lean.lean | 47 +++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 47 insertions(+) create mode 100644 test/lint_lean.lean diff --git a/test/lint_lean.lean b/test/lint_lean.lean new file mode 100644 index 0000000000..2b2f8b89e1 --- /dev/null +++ b/test/lint_lean.lean @@ -0,0 +1,47 @@ +import Batteries.Tactic.Lint + +/-! +This test file runs the environment linters set up in Batteries on the core Lean4 repository. + +Everything is commented out as it is too slow to run in regular Batteries CI +(and in any case there are many failures still), +but it is useful to run locally to see what the linters would catch. +-/ + +-- We can't apply `nolint` attributes to imported declarations, +-- but if we move the environment linters up to Lean, +-- these nolints should be installed there. +-- (And in the meantime you can "manually" ignore them!) +-- attribute [nolint dupNamespace] Lean.Elab.Tactic.Tactic +-- attribute [nolint dupNamespace] Lean.Parser.Parser Lean.Parser.Parser.rec Lean.Parser.Parser.mk +-- Lean.Parser.Parser.info Lean.Parser.Parser.fn + +/-! Failing lints that need work. -/ + +-- #lint only explicitVarsOfIff in all -- Found 109 errors + +-- Many fixed in https://github.com/leanprover/lean4/pull/4620 +-- and should be checked again. +-- #lint only simpNF in all -- Found 34 errors + +/-! Lints that fail, but that we're not intending to do anything about. -/ + +-- Mostly fixed in https://github.com/leanprover/lean4/pull/4621 +-- There are many false positives here. +-- To run this properly we would need to ignore all declarations with `@[extern]`. +-- #lint only unusedArguments in all -- Found 89 errors + +-- After https://github.com/leanprover/lean4/pull/4616, these are all intentional and have +-- `nolint` attributes above. +-- #lint only dupNamespace in all -- Found 6 errors + +-- After https://github.com/leanprover/lean4/pull/4619 all of these should be caused by `abbrev`. +-- Unless we decide to upstream something like `alias`, we're not planning on fixing these. +-- #lint only defLemma in all -- Found 31 errors + +/-! Lints that have succeeded in the past, and hopefully still do! -/ + +-- #lint only impossibleInstance in all -- Found 0 errors +-- #lint only simpVarHead in all -- Found 0 error +-- #lint only unusedHavesSuffices in all -- Found 0 errors +-- #lint only checkUnivs in all -- Found 0 errors From ac82ca1064a77eb76d37ae2ab2d1cdeb942d57fe Mon Sep 17 00:00:00 2001 From: Shrys Date: Fri, 19 Jul 2024 23:10:01 +0200 Subject: [PATCH 032/177] feat: Array lemma for empty map added (#877) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: François G. Dorais --- Batteries/Data/Array/Lemmas.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/Batteries/Data/Array/Lemmas.lean b/Batteries/Data/Array/Lemmas.lean index cf069eb703..f59644a0fa 100644 --- a/Batteries/Data/Array/Lemmas.lean +++ b/Batteries/Data/Array/Lemmas.lean @@ -126,3 +126,10 @@ theorem size_shrink_loop (a : Array α) (n) : (shrink.loop n a).size = a.size - theorem size_shrink (a : Array α) (n) : (a.shrink n).size = min a.size n := by simp [shrink, size_shrink_loop] omega + +/-! ### map -/ + +theorem mapM_empty [Monad m] (f : α → m β) : mapM f #[] = pure #[] := by + rw [mapM, mapM.map]; rfl + +@[simp] theorem map_empty (f : α → β) : map f #[] = #[] := mapM_empty .. From d2b1546c5fc05a06426e3f6ee1cb020e71be5592 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 22 Jul 2024 14:15:43 -0500 Subject: [PATCH 033/177] chore: deprecate LazyList (#884) --- Batteries/Data/LazyList.lean | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/Batteries/Data/LazyList.lean b/Batteries/Data/LazyList.lean index 48edafaf45..426cdd2248 100644 --- a/Batteries/Data/LazyList.lean +++ b/Batteries/Data/LazyList.lean @@ -4,10 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Simon Hudon -/ import Batteries.Data.Thunk - /-! # Lazy lists +Deprecated. This module is deprecated and will be removed in the future. +Most use cases can use `MLList`. Without custom support from the kernel +(previously provided in Lean 3) this type is not very useful, +but was ported from Lean 3 anyway. + The type `LazyList α` is a lazy list with elements of type `α`. In the VM, these are potentially infinite lists where all elements after the first are computed on-demand. @@ -18,6 +22,7 @@ logically we can prove that `LazyList α` is isomorphic to `List α`.) /-- Lazy list. All elements (except the first) are computed lazily. -/ +@[deprecated "Consider using `MLList`." (since := "2024-07-15")] inductive LazyList (α : Type u) : Type u /-- The empty lazy list. -/ | nil : LazyList α @@ -25,6 +30,7 @@ inductive LazyList (α : Type u) : Type u | cons (hd : α) (tl : Thunk <| LazyList α) : LazyList α +set_option linter.deprecated false namespace LazyList From e7897807913fafdab31b01b9f627550bcc96cff2 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 26 Jul 2024 10:48:35 +1000 Subject: [PATCH 034/177] chore: robustify for byAsSorry (#892) --- Batteries/CodeAction/Deprecated.lean | 2 +- Batteries/Data/List/Lemmas.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/Batteries/CodeAction/Deprecated.lean b/Batteries/CodeAction/Deprecated.lean index f308992885..ec6a2f11c2 100644 --- a/Batteries/CodeAction/Deprecated.lean +++ b/Batteries/CodeAction/Deprecated.lean @@ -32,7 +32,7 @@ def deprecatedCodeActionProvider : CodeActionProvider := fun params snap => do for m in snap.msgLog.toList do if m.data.isDeprecationWarning then if h : _ then - msgs := msgs.push (snap.cmdState.messages.toList[i]) + msgs := msgs.push (snap.cmdState.messages.toList[i]'h) i := i + 1 if msgs.isEmpty then return #[] let start := doc.meta.text.lspPosToUtf8Pos params.range.start diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 1cf017c534..b932b5ff71 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -379,7 +379,7 @@ theorem modifyNth_eq_take_drop (f : α → α) : ∀ n l, modifyNth f n l = take n l ++ modifyHead f (drop n l) := modifyNthTail_eq_take_drop _ rfl -theorem modifyNth_eq_take_cons_drop (f : α → α) {n l} (h) : +theorem modifyNth_eq_take_cons_drop (f : α → α) {n l} (h : n < length l) : modifyNth f n l = take n l ++ f l[n] :: drop (n + 1) l := by rw [modifyNth_eq_take_drop, drop_eq_getElem_cons h]; rfl From 3dfb59cffd1fdeaa41a92588c0d57e9a70cba8b6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 29 Jul 2024 20:15:16 -0400 Subject: [PATCH 035/177] feat: some array lemma aliases (#891) --- Batteries/Data/Array/Lemmas.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Batteries/Data/Array/Lemmas.lean b/Batteries/Data/Array/Lemmas.lean index f59644a0fa..753cbb4955 100644 --- a/Batteries/Data/Array/Lemmas.lean +++ b/Batteries/Data/Array/Lemmas.lean @@ -133,3 +133,15 @@ theorem mapM_empty [Monad m] (f : α → m β) : mapM f #[] = pure #[] := by rw [mapM, mapM.map]; rfl @[simp] theorem map_empty (f : α → β) : map f #[] = #[] := mapM_empty .. + +/-! ### mem -/ + +alias not_mem_empty := not_mem_nil + +theorem mem_singleton : a ∈ #[b] ↔ a = b := by simp + +/-! ### append -/ + +alias append_empty := append_nil + +alias empty_append := nil_append From f96a8aaf795d50bb08d0f77d452980ef4d918f3e Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 31 Jul 2024 15:37:43 +1000 Subject: [PATCH 036/177] chore: backports for leanprover/lean4#4814 (#894) --- Batteries/Classes/Order.lean | 2 +- Batteries/Data/List/Lemmas.lean | 5 +++-- 2 files changed, 4 insertions(+), 3 deletions(-) diff --git a/Batteries/Classes/Order.lean b/Batteries/Classes/Order.lean index f54f82eb24..8cbfae25b0 100644 --- a/Batteries/Classes/Order.lean +++ b/Batteries/Classes/Order.lean @@ -102,7 +102,7 @@ theorem cmp_congr_left (xy : cmp x y = .eq) : cmp x z = cmp y z := theorem cmp_congr_left' (xy : cmp x y = .eq) : cmp x = cmp y := funext fun _ => cmp_congr_left xy -theorem cmp_congr_right [TransCmp cmp] (yz : cmp y z = .eq) : cmp x y = cmp x z := by +theorem cmp_congr_right (yz : cmp y z = .eq) : cmp x y = cmp x z := by rw [← Ordering.swap_inj, symm, symm, cmp_congr_left yz] end TransCmp diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index b932b5ff71..533a140f45 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -912,7 +912,7 @@ section union variable [BEq α] -theorem union_def [BEq α] (l₁ l₂ : List α) : l₁ ∪ l₂ = foldr .insert l₂ l₁ := rfl +theorem union_def (l₁ l₂ : List α) : l₁ ∪ l₂ = foldr .insert l₂ l₁ := rfl @[simp] theorem nil_union (l : List α) : nil ∪ l = l := by simp [List.union_def, foldr] @@ -977,10 +977,11 @@ theorem forIn_eq_bindList [Monad m] [LawfulMonad m] section Diff variable [BEq α] -variable [LawfulBEq α] @[simp] theorem diff_nil (l : List α) : l.diff [] = l := rfl +variable [LawfulBEq α] + @[simp] theorem diff_cons (l₁ l₂ : List α) (a : α) : l₁.diff (a :: l₂) = (l₁.erase a).diff l₂ := by simp_all [List.diff, erase_of_not_mem] From 0f3e143dffdc3a591662f3401ce1d7a3405227c0 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 31 Jul 2024 19:46:56 +1000 Subject: [PATCH 037/177] chore: bump toolchain to v4.10.0 (#900) --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 6a4259fa56..7f0ea50aaa 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.10.0-rc2 +leanprover/lean4:v4.10.0 From 41bc768e2224d6c75128a877f1d6e198859b3178 Mon Sep 17 00:00:00 2001 From: Bulhwi Cha Date: Thu, 1 Aug 2024 05:52:18 +0000 Subject: [PATCH 038/177] refactor: rename theorems about `Char.utf8Size` (#901) --- Batteries/Data/String/Lemmas.lean | 54 +++++++++++++++---------------- 1 file changed, 27 insertions(+), 27 deletions(-) diff --git a/Batteries/Data/String/Lemmas.lean b/Batteries/Data/String/Lemmas.lean index 1ea4e4f788..e046783fee 100644 --- a/Batteries/Data/String/Lemmas.lean +++ b/Batteries/Data/String/Lemmas.lean @@ -34,14 +34,14 @@ instance : Batteries.BEqOrd String := .compareOfLessAndEq String.lt_irrefl attribute [simp] toList -- prefer `String.data` over `String.toList` in lemmas -private theorem add_csize_pos : 0 < i + Char.utf8Size c := +private theorem add_utf8Size_pos : 0 < i + Char.utf8Size c := Nat.add_pos_right _ (Char.utf8Size_pos c) -private theorem ne_add_csize_add_self : i ≠ n + Char.utf8Size c + i := - Nat.ne_of_lt (Nat.lt_add_of_pos_left add_csize_pos) +private theorem ne_add_utf8Size_add_self : i ≠ n + Char.utf8Size c + i := + Nat.ne_of_lt (Nat.lt_add_of_pos_left add_utf8Size_pos) -private theorem ne_self_add_add_csize : i ≠ i + (n + Char.utf8Size c) := - Nat.ne_of_lt (Nat.lt_add_of_pos_right add_csize_pos) +private theorem ne_self_add_add_utf8Size : i ≠ i + (n + Char.utf8Size c) := + Nat.ne_of_lt (Nat.lt_add_of_pos_right add_utf8Size_pos) /-- The UTF-8 byte length of a list of characters. (This is intended for specification purposes.) -/ @[inline] def utf8Len : List Char → Nat := utf8ByteSize.go @@ -64,7 +64,7 @@ private theorem ne_self_add_add_csize : i ≠ i + (n + Char.utf8Size c) := @[simp] theorem utf8Len_reverse (cs) : utf8Len cs.reverse = utf8Len cs := utf8Len_reverseAux .. @[simp] theorem utf8Len_eq_zero : utf8Len l = 0 ↔ l = [] := by - cases l <;> simp [Nat.ne_of_gt add_csize_pos] + cases l <;> simp [Nat.ne_of_gt add_utf8Size_pos] section open List @@ -92,7 +92,7 @@ attribute [ext] ext theorem lt_addChar (p : Pos) (c : Char) : p < p + c := Nat.lt_add_of_pos_right (Char.utf8Size_pos _) private theorem zero_ne_addChar {i : Pos} {c : Char} : 0 ≠ i + c := - ne_of_lt add_csize_pos + ne_of_lt add_utf8Size_pos /-- A string position is valid if it is equal to the UTF-8 length of an initial substring of `s`. -/ def Valid (s : String) (p : Pos) : Prop := @@ -154,7 +154,7 @@ theorem utf8GetAux_of_valid (cs cs' : List Char) {i p : Nat} (hp : i + utf8Len c | c::cs, cs' => simp only [utf8GetAux, List.append_eq, Char.reduceDefault, ↓Char.isValue] rw [if_neg] - case hnc => simp only [← hp, utf8Len_cons, Pos.ext_iff]; exact ne_self_add_add_csize + case hnc => simp only [← hp, utf8Len_cons, Pos.ext_iff]; exact ne_self_add_add_utf8Size refine utf8GetAux_of_valid cs cs' ?_ simpa [Nat.add_assoc, Nat.add_comm] using hp @@ -173,7 +173,7 @@ theorem utf8GetAux?_of_valid (cs cs' : List Char) {i p : Nat} (hp : i + utf8Len | c::cs, cs' => simp only [utf8GetAux?, List.append_eq] rw [if_neg] - case hnc => simp [← hp, Pos.ext_iff]; exact ne_self_add_add_csize + case hnc => simp [← hp, Pos.ext_iff]; exact ne_self_add_add_utf8Size refine utf8GetAux?_of_valid cs cs' ?_ simpa [Nat.add_assoc, Nat.add_comm] using hp @@ -188,7 +188,7 @@ theorem utf8SetAux_of_valid (c' : Char) (cs cs' : List Char) {i p : Nat} (hp : i | c::cs, cs' => simp only [utf8SetAux, List.append_eq, List.cons_append] rw [if_neg] - case hnc => simp [← hp, Pos.ext_iff]; exact ne_self_add_add_csize + case hnc => simp [← hp, Pos.ext_iff]; exact ne_self_add_add_utf8Size refine congrArg (c::·) (utf8SetAux_of_valid c' cs cs' ?_) simpa [Nat.add_assoc, Nat.add_comm] using hp @@ -228,14 +228,14 @@ theorem utf8PrevAux_of_valid {cs cs' : List Char} {c : Char} {i p : Nat} case hnc => simp only [Pos.ext_iff] rw [Nat.add_right_comm, Nat.add_left_comm] - apply ne_add_csize_add_self + apply ne_add_utf8Size_add_self refine (utf8PrevAux_of_valid (by simp [Nat.add_assoc, Nat.add_left_comm])).trans ?_ simp [Nat.add_assoc, Nat.add_comm] theorem prev_of_valid (cs : List Char) (c : Char) (cs' : List Char) : prev ⟨cs ++ c :: cs'⟩ ⟨utf8Len cs + c.utf8Size⟩ = ⟨utf8Len cs⟩ := by simp only [prev] - refine (if_neg (Pos.ne_of_gt add_csize_pos)).trans ?_ + refine (if_neg (Pos.ne_of_gt add_utf8Size_pos)).trans ?_ rw [utf8PrevAux_of_valid] <;> simp theorem prev_of_valid' (cs cs' : List Char) : @@ -255,7 +255,7 @@ theorem back_eq (s : String) : back s = s.1.getLastD default := by theorem atEnd_of_valid (cs : List Char) (cs' : List Char) : atEnd ⟨cs ++ cs'⟩ ⟨utf8Len cs⟩ ↔ cs' = [] := by rw [atEnd_iff] - cases cs' <;> simp [Nat.lt_add_of_pos_right add_csize_pos] + cases cs' <;> simp [Nat.lt_add_of_pos_right add_utf8Size_pos] unseal posOfAux findAux in theorem posOfAux_eq (s c) : posOfAux s c = findAux s (· == c) := rfl @@ -276,7 +276,7 @@ theorem findAux_of_valid (p) : ∀ l m r, | l, [], r => by unfold findAux List.takeWhile; simp | l, c::m, r => by unfold findAux List.takeWhile - rw [dif_pos (by exact Nat.lt_add_of_pos_right add_csize_pos)] + rw [dif_pos (by exact Nat.lt_add_of_pos_right add_utf8Size_pos)] have h1 := get_of_valid l (c::m++r); have h2 := next_of_valid l c (m++r) simp only [List.cons_append, Char.reduceDefault, List.headD_cons] at h1 h2 simp only [List.append_assoc, List.cons_append, h1, utf8Len_cons, h2] @@ -298,7 +298,7 @@ theorem revFindAux_of_valid (p) : ∀ l r, | [], r => by unfold revFindAux List.dropWhile; simp | c::l, r => by unfold revFindAux List.dropWhile - rw [dif_neg (by exact Pos.ne_of_gt add_csize_pos)] + rw [dif_neg (by exact Pos.ne_of_gt add_utf8Size_pos)] have h1 := get_of_valid l.reverse (c::r); have h2 := prev_of_valid l.reverse c r simp only [utf8Len_reverse, Char.reduceDefault, List.headD_cons] at h1 h2 simp only [List.reverse_cons, List.append_assoc, List.singleton_append, utf8Len_cons, h2, h1] @@ -320,7 +320,7 @@ theorem firstDiffPos_loop_eq (l₁ l₂ r₁ r₂ stop p) rw [ dif_pos <| by rw [hstop, ← hl₁, ← hl₂] - refine Nat.lt_min.2 ⟨?_, ?_⟩ <;> exact Nat.lt_add_of_pos_right add_csize_pos, + refine Nat.lt_min.2 ⟨?_, ?_⟩ <;> exact Nat.lt_add_of_pos_right add_utf8Size_pos, show get ⟨l₁ ++ a :: r₁⟩ ⟨p⟩ = a by simp [hl₁, get_of_valid], show get ⟨l₂ ++ b :: r₂⟩ ⟨p⟩ = b by simp [hl₂, get_of_valid]] simp only [bne_iff_ne, ne_eq, ite_not, decide_eq_true_eq] @@ -360,7 +360,7 @@ theorem extract.go₂_append_left : ∀ (s t : List Char) (i e : Nat), e = utf8Len s + i → go₂ (s ++ t) ⟨i⟩ ⟨e⟩ = s | [], t, i, _, rfl => by cases t <;> simp [go₂] | c :: cs, t, i, _, rfl => by - simp only [go₂, utf8Len_cons, Pos.ext_iff, ne_add_csize_add_self, ↓reduceIte, List.append_eq, + simp only [go₂, utf8Len_cons, Pos.ext_iff, ne_add_utf8Size_add_self, ↓reduceIte, List.append_eq, Pos.addChar_eq, List.cons.injEq, true_and] apply go₂_append_left; rw [Nat.add_right_comm, Nat.add_assoc] @@ -379,7 +379,7 @@ theorem extract.go₁_add_right_cancel (s : List Char) (i b e n : Nat) : theorem extract.go₁_cons_addChar (c : Char) (cs : List Char) (b e : Pos) : go₁ (c :: cs) 0 (b + c) (e + c) = go₁ cs 0 b e := by - simp only [go₁, Pos.ext_iff, Pos.byteIdx_zero, pos_add_char, Nat.ne_of_lt add_csize_pos, + simp only [go₁, Pos.ext_iff, Pos.byteIdx_zero, pos_add_char, Nat.ne_of_lt add_utf8Size_pos, ↓reduceIte] apply go₁_add_right_cancel @@ -387,7 +387,7 @@ theorem extract.go₁_append_right : ∀ (s t : List Char) (i b : Nat) (e : Pos) b = utf8Len s + i → go₁ (s ++ t) ⟨i⟩ ⟨b⟩ e = go₂ t ⟨b⟩ e | [], t, i, _, e, rfl => by cases t <;> simp [go₁, go₂] | c :: cs, t, i, _, e, rfl => by - simp only [go₁, utf8Len_cons, Pos.ext_iff, ne_add_csize_add_self, ↓reduceIte, List.append_eq, + simp only [go₁, utf8Len_cons, Pos.ext_iff, ne_add_utf8Size_add_self, ↓reduceIte, List.append_eq, Pos.addChar_eq] apply go₁_append_right; rw [Nat.add_right_comm, Nat.add_assoc] @@ -404,7 +404,7 @@ theorem extract_zero_endPos : ∀ (s : String), s.extract 0 (endPos s) = s | ⟨[]⟩ => rfl | ⟨c :: cs⟩ => by simp only [extract, Pos.byteIdx_zero, endPos_eq, utf8Len_cons, ge_iff_le, Nat.le_zero_eq, - Nat.ne_of_gt add_csize_pos, ↓reduceIte] + Nat.ne_of_gt add_utf8Size_pos, ↓reduceIte] congr apply extract.go₁_zero_utf8Len @@ -645,7 +645,7 @@ theorem offsetOfPosAux_of_valid : ∀ l m r n, | l, [], r, n => by unfold offsetOfPosAux; simp | l, c::m, r, n => by unfold offsetOfPosAux - rw [if_neg (by exact Nat.not_le.2 (Nat.lt_add_of_pos_right add_csize_pos))] + rw [if_neg (by exact Nat.not_le.2 (Nat.lt_add_of_pos_right add_utf8Size_pos))] simp only [List.append_assoc, atEnd_of_valid l (c::m++r)] simp only [List.cons_append, utf8Len_cons, next_of_valid l c (m ++ r)] simpa [← Nat.add_assoc, Nat.add_right_comm] using @@ -660,7 +660,7 @@ theorem foldlAux_of_valid (f : α → Char → α) : ∀ l m r a, | l, [], r, a => by unfold foldlAux; simp | l, c::m, r, a => by unfold foldlAux - rw [dif_pos (by exact Nat.lt_add_of_pos_right add_csize_pos)] + rw [dif_pos (by exact Nat.lt_add_of_pos_right add_utf8Size_pos)] simp only [List.append_assoc, List.cons_append, utf8Len_cons, next_of_valid l c (m ++ r), get_of_valid l (c :: (m ++ r)), Char.reduceDefault, List.headD_cons, List.foldl_cons] simpa [← Nat.add_assoc, Nat.add_right_comm] using foldlAux_of_valid f (l++[c]) m r (f a c) @@ -674,7 +674,7 @@ theorem foldrAux_of_valid (f : Char → α → α) (l m r a) : rw [← m.reverse_reverse] induction m.reverse generalizing r a with (unfold foldrAux; simp) | cons c m IH => - rw [if_pos (by exact Nat.lt_add_of_pos_right add_csize_pos)] + rw [if_pos (by exact Nat.lt_add_of_pos_right add_utf8Size_pos)] simp only [← Nat.add_assoc, by simpa using prev_of_valid (l ++ m.reverse) c r] simp only [by simpa using get_of_valid (l ++ m.reverse) (c :: r)] simpa using IH (c::r) (f c a) @@ -688,7 +688,7 @@ theorem anyAux_of_valid (p : Char → Bool) : ∀ l m r, | l, [], r => by unfold anyAux; simp | l, c::m, r => by unfold anyAux - rw [dif_pos (by exact Nat.lt_add_of_pos_right add_csize_pos)] + rw [dif_pos (by exact Nat.lt_add_of_pos_right add_utf8Size_pos)] simp only [List.append_assoc, List.cons_append, get_of_valid l (c :: (m ++ r)), Char.reduceDefault, List.headD_cons, utf8Len_cons, next_of_valid l c (m ++ r), Bool.if_true_left, Bool.decide_eq_true, List.any_cons] @@ -732,7 +732,7 @@ theorem takeWhileAux_of_valid (p : Char → Bool) : ∀ l m r, | l, [], r => by unfold Substring.takeWhileAux List.takeWhile; simp | l, c::m, r => by unfold Substring.takeWhileAux List.takeWhile - rw [dif_pos (by exact Nat.lt_add_of_pos_right add_csize_pos)] + rw [dif_pos (by exact Nat.lt_add_of_pos_right add_utf8Size_pos)] simp only [List.append_assoc, List.cons_append, get_of_valid l (c :: (m ++ r)), Char.reduceDefault, List.headD_cons, utf8Len_cons, next_of_valid l c (m ++ r)] cases p c <;> simp @@ -809,7 +809,7 @@ theorem next : ∀ {s}, ValidFor l (m₁ ++ c :: m₂) r s → rw [if_neg (mt Pos.ext_iff.1 ?a)] case a => simpa [Nat.add_assoc, Nat.add_comm, Nat.add_left_comm] using - @ne_add_csize_add_self (utf8Len l + utf8Len m₁) (utf8Len m₂) c + @ne_add_utf8Size_add_self (utf8Len l + utf8Len m₁) (utf8Len m₂) c have := next_of_valid (l ++ m₁) c (m₂ ++ r) simp only [List.append_assoc, utf8Len_append, Pos.add_eq] at this ⊢; rw [this] simp [Nat.add_assoc, Nat.add_sub_cancel_left] @@ -822,7 +822,7 @@ theorem prev : ∀ {s}, ValidFor l (m₁ ++ c :: m₂) r s → | _, ⟨⟩ => by simp only [Substring.prev, List.append_assoc, List.cons_append] rw [if_neg (mt Pos.ext_iff.1 <| Ne.symm ?a)] - case a => simpa [Nat.add_comm] using @ne_add_csize_add_self (utf8Len l) (utf8Len m₁) c + case a => simpa [Nat.add_comm] using @ne_add_utf8Size_add_self (utf8Len l) (utf8Len m₁) c have := prev_of_valid (l ++ m₁) c (m₂ ++ r) simp only [List.append_assoc, utf8Len_append, Nat.add_assoc, Pos.add_eq] at this ⊢; rw [this] simp [Nat.add_sub_cancel_left] From 021e272cb5cdcc82b7e1e760fe915ff2f64169ad Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 5 Aug 2024 12:16:22 +1000 Subject: [PATCH 039/177] chore: move to v4.11.0-rc1 (#907) --- Batteries.lean | 4 - Batteries/Classes/BEq.lean | 18 - Batteries/Data/Array/Basic.lean | 18 - Batteries/Data/Char.lean | 5 - Batteries/Data/HashMap/Basic.lean | 1 - Batteries/Data/HashMap/WF.lean | 10 +- Batteries/Data/List.lean | 1 - Batteries/Data/List/Basic.lean | 204 +----- Batteries/Data/List/Count.lean | 211 +----- Batteries/Data/List/EraseIdx.lean | 70 +- Batteries/Data/List/Init/Attach.lean | 44 -- Batteries/Data/List/Lemmas.lean | 985 +-------------------------- Batteries/Data/List/Pairwise.lean | 210 ------ Batteries/Data/List/Perm.lean | 17 +- Batteries/Data/Nat/Basic.lean | 3 - Batteries/Data/Nat/Lemmas.lean | 4 - Batteries/Data/RBMap/Lemmas.lean | 1 - Batteries/Data/Range/Lemmas.lean | 1 - Batteries/Data/String/Lemmas.lean | 15 +- Batteries/Data/UInt.lean | 10 - Batteries/Lean/SMap.lean | 17 - Batteries/Lean/Util/EnvSearch.lean | 3 +- Batteries/Lean/Util/Path.lean | 34 - Batteries/Logic.lean | 11 +- Batteries/StdDeprecations.lean | 1 - Batteries/Tactic/Lint/Basic.lean | 2 + Batteries/Util/CheckTactic.lean | 124 ---- lean-toolchain | 2 +- scripts/runLinter.lean | 2 +- test/array.lean | 1 - test/congr.lean | 2 - 31 files changed, 41 insertions(+), 1990 deletions(-) delete mode 100644 Batteries/Classes/BEq.lean delete mode 100644 Batteries/Data/List/Init/Attach.lean delete mode 100644 Batteries/Lean/SMap.lean delete mode 100644 Batteries/Lean/Util/Path.lean delete mode 100644 Batteries/Util/CheckTactic.lean diff --git a/Batteries.lean b/Batteries.lean index 2fac3bf121..8e1e6f1828 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -1,4 +1,3 @@ -import Batteries.Classes.BEq import Batteries.Classes.Cast import Batteries.Classes.Order import Batteries.Classes.RatCast @@ -64,12 +63,10 @@ import Batteries.Lean.NameMapAttribute import Batteries.Lean.PersistentHashMap import Batteries.Lean.PersistentHashSet import Batteries.Lean.Position -import Batteries.Lean.SMap import Batteries.Lean.Syntax import Batteries.Lean.System.IO import Batteries.Lean.TagAttribute import Batteries.Lean.Util.EnvSearch -import Batteries.Lean.Util.Path import Batteries.Linter import Batteries.Linter.UnnecessarySeqFocus import Batteries.Linter.UnreachableTactic @@ -101,7 +98,6 @@ import Batteries.Tactic.Unreachable import Batteries.Tactic.Where import Batteries.Test.Internal.DummyLabelAttr import Batteries.Util.Cache -import Batteries.Util.CheckTactic import Batteries.Util.ExtendedBinder import Batteries.Util.LibraryNote import Batteries.Util.Pickle diff --git a/Batteries/Classes/BEq.lean b/Batteries/Classes/BEq.lean deleted file mode 100644 index 98318f97c9..0000000000 --- a/Batteries/Classes/BEq.lean +++ /dev/null @@ -1,18 +0,0 @@ -/- -Copyright (c) 2022 Mario Carneiro. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro --/ - -/-! ## Boolean equality -/ - -/-- `PartialEquivBEq α` says that the `BEq` implementation is a -partial equivalence relation, that is: -* it is symmetric: `a == b → b == a` -* it is transitive: `a == b → b == c → a == c`. --/ -class PartialEquivBEq (α) [BEq α] : Prop where - /-- Symmetry for `BEq`. If `a == b` then `b == a`. -/ - symm : (a : α) == b → b == a - /-- Transitivity for `BEq`. If `a == b` and `b == c` then `a == c`. -/ - trans : (a : α) == b → b == c → a == c diff --git a/Batteries/Data/Array/Basic.lean b/Batteries/Data/Array/Basic.lean index 1dddea3c14..ad6b8129e6 100644 --- a/Batteries/Data/Array/Basic.lean +++ b/Batteries/Data/Array/Basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Arthur Paulino, Floris van Doorn, Jannis Limperg -/ -import Batteries.Data.List.Init.Attach import Batteries.Data.Array.Init.Lemmas /-! @@ -130,23 +129,6 @@ protected def maxI [ord : Ord α] [Inhabited α] (xs : Array α) (start := 0) (stop := xs.size) : α := xs.minI (ord := ord.opposite) start stop -/-- -Unsafe implementation of `attachWith`, taking advantage of the fact that the representation of -`Array {x // P x}` is the same as the input `Array α`. --/ -@[inline] private unsafe def attachWithImpl - (xs : Array α) (P : α → Prop) (_ : ∀ x ∈ xs, P x) : Array {x // P x} := unsafeCast xs - -/-- `O(1)`. "Attach" a proof `P x` that holds for all the elements of `xs` to produce a new array - with the same elements but in the type `{x // P x}`. -/ -@[implemented_by attachWithImpl] def attachWith - (xs : Array α) (P : α → Prop) (H : ∀ x ∈ xs, P x) : Array {x // P x} := - ⟨xs.data.attachWith P fun x h => H x (Array.Mem.mk h)⟩ - -/-- `O(1)`. "Attach" the proof that the elements of `xs` are in `xs` to produce a new array - with the same elements but in the type `{x // x ∈ xs}`. -/ -@[inline] def attach (xs : Array α) : Array {x // x ∈ xs} := xs.attachWith _ fun _ => id - /-- `O(|join L|)`. `join L` concatenates all the arrays in `L` into one array. * `join #[#[a], #[], #[b, c], #[d, e, f]] = #[a, b, c, d, e, f]` diff --git a/Batteries/Data/Char.lean b/Batteries/Data/Char.lean index 5a9110ae12..4f3c59ef0b 100644 --- a/Batteries/Data/Char.lean +++ b/Batteries/Data/Char.lean @@ -6,11 +6,6 @@ Authors: Jannis Limperg import Batteries.Data.UInt import Batteries.Tactic.Alias -@[ext] theorem Char.ext : {a b : Char} → a.val = b.val → a = b - | ⟨_,_⟩, ⟨_,_⟩, rfl => rfl - -theorem Char.ext_iff {x y : Char} : x = y ↔ x.val = y.val := ⟨congrArg _, Char.ext⟩ - theorem Char.le_antisymm_iff {x y : Char} : x = y ↔ x ≤ y ∧ y ≤ x := Char.ext_iff.trans UInt32.le_antisymm_iff diff --git a/Batteries/Data/HashMap/Basic.lean b/Batteries/Data/HashMap/Basic.lean index 106fea1e4d..3419ba35ab 100644 --- a/Batteries/Data/HashMap/Basic.lean +++ b/Batteries/Data/HashMap/Basic.lean @@ -6,7 +6,6 @@ Authors: Leonardo de Moura, Mario Carneiro import Batteries.Data.AssocList import Batteries.Data.Nat.Basic import Batteries.Data.Array.Monadic -import Batteries.Classes.BEq namespace Batteries.HashMap diff --git a/Batteries/Data/HashMap/WF.lean b/Batteries/Data/HashMap/WF.lean index f6298d3322..e83e32e567 100644 --- a/Batteries/Data/HashMap/WF.lean +++ b/Batteries/Data/HashMap/WF.lean @@ -24,7 +24,7 @@ theorem exists_of_update (self : Buckets α β) (i d h) : ∃ l₁ l₂, self.1.data = l₁ ++ self.1[i] :: l₂ ∧ List.length l₁ = i.toNat ∧ (self.update i d h).1.data = l₁ ++ d :: l₂ := by simp only [Array.data_length, Array.ugetElem_eq_getElem, Array.getElem_eq_data_getElem] - exact List.exists_of_set' h + exact List.exists_of_set h theorem update_update (self : Buckets α β) (i d d' h h') : (self.update i d h).update i d' h' = self.update i d' h := by @@ -95,11 +95,11 @@ where · next H => refine (go (i+1) _ _ fun j hj => ?a).trans ?b <;> simp · case a => - simp [List.getD_eq_getElem?, List.getElem?_set, Option.map_eq_map]; split + simp [List.getD_eq_getElem?_getD, List.getElem?_set, Option.map_eq_map]; split · cases source.data[j]? <;> rfl · next H => exact hs _ (Nat.lt_of_le_of_ne (Nat.le_of_lt_succ hj) (Ne.symm H)) · case b => - refine have ⟨l₁, l₂, h₁, _, eq⟩ := List.exists_of_set' H; eq ▸ ?_ + refine have ⟨l₁, l₂, h₁, _, eq⟩ := List.exists_of_set H; eq ▸ ?_ rw [h₁] simp only [Buckets.size_eq, List.map_append, List.map_cons, AssocList.toList, List.length_nil, Nat.sum_append, Nat.sum_cons, Nat.zero_add, Array.data_length] @@ -315,7 +315,7 @@ theorem WF.mapVal {α β γ} {f : α → β → γ} [BEq α] [Hashable α] have ⟨h₁, h₂⟩ := H.out simp only [Imp.mapVal, h₁, Buckets.mapVal, WF_iff]; refine ⟨?_, ?_, fun i h => ?_⟩ · simp only [Buckets.size, Array.map_data, List.map_map]; congr; funext l; simp - · simp only [Array.map_data, List.forall_mem_map_iff] + · simp only [Array.map_data, List.forall_mem_map] simp only [AssocList.toList_mapVal, List.pairwise_map] exact fun _ => h₂.1 _ · simp only [Array.size_map, AssocList.All, Array.getElem_map, AssocList.toList_mapVal, @@ -361,7 +361,7 @@ theorem WF.filterMap {α β γ} {f : α → β → Option γ} [BEq α] [Hashable simp only [Array.mapM_eq_mapM_data, bind, StateT.bind, H2, List.map_map, Nat.zero_add, g] intro bk sz h e'; cases e' refine .mk (by simp [Buckets.size]) ⟨?_, fun i h => ?_⟩ - · simp only [List.forall_mem_map_iff, List.toList_toAssocList] + · simp only [List.forall_mem_map, List.toList_toAssocList] refine fun l h => (List.pairwise_reverse.2 ?_).imp (mt PartialEquivBEq.symm) have := H.out.2.1 _ h rw [← List.pairwise_map (R := (¬ · == ·))] at this ⊢ diff --git a/Batteries/Data/List.lean b/Batteries/Data/List.lean index d060ba89c8..998160b904 100644 --- a/Batteries/Data/List.lean +++ b/Batteries/Data/List.lean @@ -1,7 +1,6 @@ import Batteries.Data.List.Basic import Batteries.Data.List.Count import Batteries.Data.List.EraseIdx -import Batteries.Data.List.Init.Attach import Batteries.Data.List.Init.Lemmas import Batteries.Data.List.Lemmas import Batteries.Data.List.Pairwise diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index 2cff94c670..43915cd733 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -10,16 +10,6 @@ namespace List /-! ## New definitions -/ -/-- -`l₁ ⊆ l₂` means that every element of `l₁` is also an element of `l₂`, ignoring multiplicity. --/ -protected def Subset (l₁ l₂ : List α) := ∀ ⦃a : α⦄, a ∈ l₁ → a ∈ l₂ - -instance : HasSubset (List α) := ⟨List.Subset⟩ - -instance [DecidableEq α] : DecidableRel (Subset : List α → List α → Prop) := - fun _ _ => decidableBAll _ _ - /-- Computes the "bag intersection" of `l₁` and `l₂`, that is, the collection of elements of `l₁` which are also in `l₂`. As each element @@ -37,15 +27,6 @@ protected def diff {α} [BEq α] : List α → List α → List α open Option Nat -/-- Get the tail of a nonempty list, or return `[]` for `[]`. -/ -def tail : List α → List α - | [] => [] - | _::as => as - --- FIXME: `@[simp]` on the definition simplifies even `tail l` -@[simp] theorem tail_nil : @tail α [] = [] := rfl -@[simp] theorem tail_cons : @tail α (a::as) = as := rfl - /-- Get the head and tail of a list, if it is nonempty. -/ @[inline] def next? : List α → Option (α × List α) | [] => none @@ -83,16 +64,6 @@ drop_while (· != 1) [0, 1, 2, 3] = [1, 2, 3] | [] => [] | x :: xs => bif p x then xs else after p xs -/-- Returns the index of the first element satisfying `p`, or the length of the list otherwise. -/ -@[inline] def findIdx (p : α → Bool) (l : List α) : Nat := go l 0 where - /-- Auxiliary for `findIdx`: `findIdx.go p l n = findIdx p l + n` -/ - @[specialize] go : List α → Nat → Nat - | [], n => n - | a :: l, n => bif p a then n else go l (n + 1) - -/-- Returns the index of the first element equal to `a`, or the length of the list otherwise. -/ -def indexOf [BEq α] (a : α) : List α → Nat := findIdx (· == a) - @[deprecated (since := "2024-05-06")] alias removeNth := eraseIdx @[deprecated (since := "2024-05-06")] alias removeNthTR := eraseIdxTR @[deprecated (since := "2024-05-06")] alias removeNth_eq_removeNthTR := eraseIdx_eq_eraseIdxTR @@ -139,26 +110,6 @@ Unlike `bagInter` this does not preserve multiplicity: `[1, 1].inter [1]` is `[1 instance [BEq α] : Inter (List α) := ⟨List.inter⟩ -/-- `l₁ <+ l₂`, or `Sublist l₁ l₂`, says that `l₁` is a (non-contiguous) subsequence of `l₂`. -/ -inductive Sublist {α} : List α → List α → Prop - /-- the base case: `[]` is a sublist of `[]` -/ - | slnil : Sublist [] [] - /-- If `l₁` is a subsequence of `l₂`, then it is also a subsequence of `a :: l₂`. -/ - | cons a : Sublist l₁ l₂ → Sublist l₁ (a :: l₂) - /-- If `l₁` is a subsequence of `l₂`, then `a :: l₁` is a subsequence of `a :: l₂`. -/ - | cons₂ a : Sublist l₁ l₂ → Sublist (a :: l₁) (a :: l₂) - -@[inherit_doc] scoped infixl:50 " <+ " => Sublist - -/-- True if the first list is a potentially non-contiguous sub-sequence of the second list. -/ -def isSublist [BEq α] : List α → List α → Bool - | [], _ => true - | _, [] => false - | l₁@(hd₁::tl₁), hd₂::tl₂ => - if hd₁ == hd₂ - then tl₁.isSublist tl₂ - else l₁.isSublist tl₂ - /-- Split a list at an index. ``` @@ -299,7 +250,7 @@ theorem insertNthTR_go_eq : ∀ n l, insertNthTR.go a n l acc = acc.data ++ inse @[csimp] theorem insertNth_eq_insertNthTR : @insertNth = @insertNthTR := by funext α f n l; simp [insertNthTR, insertNthTR_go_eq] -@[simp] theorem headD_eq_head? (l) (a : α) : headD l a = (head? l).getD a := by cases l <;> rfl +theorem headD_eq_head? (l) (a : α) : headD l a = (head? l).getD a := by cases l <;> rfl /-- Take `n` elements from a list `l`. If `l` has less than `n` elements, append `n - length l` @@ -332,19 +283,6 @@ theorem takeDTR_go_eq : ∀ n l, takeDTR.go dflt n l acc = acc.data ++ takeD n l @[csimp] theorem takeD_eq_takeDTR : @takeD = @takeDTR := by funext α f n l; simp [takeDTR, takeDTR_go_eq] -/-- -Pads `l : List α` with repeated occurrences of `a : α` until it is of length `n`. -If `l` is initially larger than `n`, just return `l`. --/ -def leftpad (n : Nat) (a : α) (l : List α) : List α := replicate (n - length l) a ++ l - -/-- Optimized version of `leftpad`. -/ -@[inline] def leftpadTR (n : Nat) (a : α) (l : List α) : List α := - replicateTR.loop a (n - length l) l - -@[csimp] theorem leftpad_eq_leftpadTR : @leftpad = @leftpadTR := by - funext α n a l; simp [leftpad, leftpadTR, replicateTR_loop_eq] - /-- Fold a function `f` over the list from the left, returning the list of partial results. ``` @@ -416,14 +354,6 @@ indexesOf a [a, b, a, a] = [0, 2, 3] -/ @[inline] def indexesOf [BEq α] (a : α) : List α → List Nat := findIdxs (· == a) -/-- Return the index of the first occurrence of an element satisfying `p`. -/ -def findIdx? (p : α → Bool) : List α → (start : Nat := 0) → Option Nat -| [], _ => none -| a :: l, i => if p a then some i else findIdx? p l (i + 1) - -/-- Return the index of the first occurrence of `a` in the list. -/ -@[inline] def indexOf? [BEq α] (a : α) : List α → Option Nat := findIdx? (· == a) - /-- `lookmap` is a combination of `lookup` and `filterMap`. `lookmap f l` will apply `f : α → Option α` to each element of the list, @@ -437,40 +367,6 @@ replacing `a → b` at the first value `a` in the list such that `f a = some b`. | some b => acc.toListAppend (b :: l) | none => go l (acc.push a) -/-- `countP p l` is the number of elements of `l` that satisfy `p`. -/ -@[inline] def countP (p : α → Bool) (l : List α) : Nat := go l 0 where - /-- Auxiliary for `countP`: `countP.go p l acc = countP p l + acc`. -/ - @[specialize] go : List α → Nat → Nat - | [], acc => acc - | x :: xs, acc => bif p x then go xs (acc + 1) else go xs acc - -/-- `count a l` is the number of occurrences of `a` in `l`. -/ -@[inline] def count [BEq α] (a : α) : List α → Nat := countP (· == a) - -/-- -`IsPrefix l₁ l₂`, or `l₁ <+: l₂`, means that `l₁` is a prefix of `l₂`, -that is, `l₂` has the form `l₁ ++ t` for some `t`. --/ -def IsPrefix (l₁ : List α) (l₂ : List α) : Prop := ∃ t, l₁ ++ t = l₂ - -/-- -`IsSuffix l₁ l₂`, or `l₁ <:+ l₂`, means that `l₁` is a suffix of `l₂`, -that is, `l₂` has the form `t ++ l₁` for some `t`. --/ -def IsSuffix (l₁ : List α) (l₂ : List α) : Prop := ∃ t, t ++ l₁ = l₂ - -/-- -`IsInfix l₁ l₂`, or `l₁ <:+: l₂`, means that `l₁` is a contiguous -substring of `l₂`, that is, `l₂` has the form `s ++ l₁ ++ t` for some `s, t`. --/ -def IsInfix (l₁ : List α) (l₂ : List α) : Prop := ∃ s t, s ++ l₁ ++ t = l₂ - -@[inherit_doc] infixl:50 " <+: " => IsPrefix - -@[inherit_doc] infixl:50 " <:+ " => IsSuffix - -@[inherit_doc] infixl:50 " <:+: " => IsInfix - /-- `inits l` is the list of initial segments of `l`. ``` @@ -659,29 +555,6 @@ theorem sections_eq_nil_of_isEmpty : ∀ {L}, L.any isEmpty → @sections α L = rw [Array.foldl_eq_foldl_data, Array.foldl_data_eq_bind]; rfl intros; apply Array.foldl_data_eq_map -/-- `eraseP p l` removes the first element of `l` satisfying the predicate `p`. -/ -def eraseP (p : α → Bool) : List α → List α - | [] => [] - | a :: l => bif p a then l else a :: eraseP p l - -/-- Tail-recursive version of `eraseP`. -/ -@[inline] def erasePTR (p : α → Bool) (l : List α) : List α := go l #[] where - /-- Auxiliary for `erasePTR`: `erasePTR.go p l xs acc = acc.toList ++ eraseP p xs`, - unless `xs` does not contain any elements satisfying `p`, where it returns `l`. -/ - @[specialize] go : List α → Array α → List α - | [], _ => l - | a :: l, acc => bif p a then acc.toListAppend l else go l (acc.push a) - -@[csimp] theorem eraseP_eq_erasePTR : @eraseP = @erasePTR := by - funext α p l; simp [erasePTR] - let rec go (acc) : ∀ xs, l = acc.data ++ xs → - erasePTR.go p l xs acc = acc.data ++ xs.eraseP p - | [] => fun h => by simp [erasePTR.go, eraseP, h] - | x::xs => by - simp [erasePTR.go, eraseP]; cases p x <;> simp - · intro h; rw [go _ xs]; {simp}; simp [h] - exact (go #[] _ rfl).symm - /-- `extractP p l` returns a pair of an element `a` of `l` satisfying the predicate `p`, and `l`, with `a` removed. If there is no such element `a` it returns `(none, l)`. @@ -806,46 +679,6 @@ where rename_i a as b bs; unfold cond; cases R a b <;> simp [go as bs] exact (go as bs [] []).symm -section Pairwise - -variable (R : α → α → Prop) - -/-- -`Pairwise R l` means that all the elements with earlier indexes are -`R`-related to all the elements with later indexes. -``` -Pairwise R [1, 2, 3] ↔ R 1 2 ∧ R 1 3 ∧ R 2 3 -``` -For example if `R = (·≠·)` then it asserts `l` has no duplicates, -and if `R = (·<·)` then it asserts that `l` is (strictly) sorted. --/ -inductive Pairwise : List α → Prop - /-- All elements of the empty list are vacuously pairwise related. -/ - | nil : Pairwise [] - /-- `a :: l` is `Pairwise R` if `a` `R`-relates to every element of `l`, - and `l` is `Pairwise R`. -/ - | cons : ∀ {a : α} {l : List α}, (∀ a' ∈ l, R a a') → Pairwise l → Pairwise (a :: l) - -attribute [simp] Pairwise.nil - -variable {R} - -@[simp] theorem pairwise_cons : Pairwise R (a::l) ↔ (∀ a' ∈ l, R a a') ∧ Pairwise R l := - ⟨fun | .cons h₁ h₂ => ⟨h₁, h₂⟩, fun ⟨h₁, h₂⟩ => h₂.cons h₁⟩ - -instance instDecidablePairwise [DecidableRel R] : - (l : List α) → Decidable (Pairwise R l) - | [] => isTrue .nil - | hd :: tl => - match instDecidablePairwise tl with - | isTrue ht => - match decidableBAll (R hd) tl with - | isFalse hf => isFalse fun hf' => hf (pairwise_cons.1 hf').1 - | isTrue ht' => isTrue <| pairwise_cons.mpr (And.intro ht' ht) - | isFalse hf => isFalse fun | .cons _ ih => hf ih - -end Pairwise - /-- `pwFilter R l` is a maximal sublist of `l` which is `Pairwise R`. `pwFilter (·≠·)` is the erase duplicates function (cf. `eraseDup`), and `pwFilter (·<·)` finds @@ -881,47 +714,12 @@ def Chain' : List α → Prop end Chain -/-- `Nodup l` means that `l` has no duplicates, that is, any element appears at most - once in the List. It is defined as `Pairwise (≠)`. -/ -def Nodup : List α → Prop := Pairwise (· ≠ ·) - -instance nodupDecidable [DecidableEq α] : ∀ l : List α, Decidable (Nodup l) := - instDecidablePairwise - /-- `eraseDup l` removes duplicates from `l` (taking only the first occurrence). Defined as `pwFilter (≠)`. eraseDup [1, 0, 2, 2, 1] = [0, 2, 1] -/ @[inline] def eraseDup [BEq α] : List α → List α := pwFilter (· != ·) -/-- `range' start len step` is the list of numbers `[start, start+step, ..., start+(len-1)*step]`. - It is intended mainly for proving properties of `range` and `iota`. -/ -def range' : (start len : Nat) → (step : Nat := 1) → List Nat - | _, 0, _ => [] - | s, n+1, step => s :: range' (s+step) n step - -/-- Optimized version of `range'`. -/ -@[inline] def range'TR (s n : Nat) (step : Nat := 1) : List Nat := go n (s + step * n) [] where - /-- Auxiliary for `range'TR`: `range'TR.go n e = [e-n, ..., e-1] ++ acc`. -/ - go : Nat → Nat → List Nat → List Nat - | 0, _, acc => acc - | n+1, e, acc => go n (e-step) ((e-step) :: acc) - -@[csimp] theorem range'_eq_range'TR : @range' = @range'TR := by - funext s n step - let rec go (s) : ∀ n m, - range'TR.go step n (s + step * n) (range' (s + step * n) m step) = range' s (n + m) step - | 0, m => by simp [range'TR.go] - | n+1, m => by - simp [range'TR.go] - rw [Nat.mul_succ, ← Nat.add_assoc, Nat.add_sub_cancel, Nat.add_right_comm n] - exact go s n (m + 1) - exact (go s n 0).symm - -/-- Drop `none`s from a list, and replace each remaining `some a` with `a`. -/ -@[inline] def reduceOption {α} : List (Option α) → List α := - List.filterMap id - /-- `ilast' x xs` returns the last element of `xs` if `xs` is non-empty; it returns `x` otherwise. Use `List.getLastD` instead. diff --git a/Batteries/Data/List/Count.lean b/Batteries/Data/List/Count.lean index dfd1481052..036a76b4b8 100644 --- a/Batteries/Data/List/Count.lean +++ b/Batteries/Data/List/Count.lean @@ -19,111 +19,6 @@ open Nat namespace List -section countP - -variable (p q : α → Bool) - -@[simp] theorem countP_nil : countP p [] = 0 := rfl - -protected theorem countP_go_eq_add (l) : countP.go p l n = n + countP.go p l 0 := by - induction l generalizing n with - | nil => rfl - | cons head tail ih => - unfold countP.go - rw [ih (n := n + 1), ih (n := n), ih (n := 1)] - if h : p head then simp [h, Nat.add_assoc] else simp [h] - -@[simp] theorem countP_cons_of_pos (l) (pa : p a) : countP p (a :: l) = countP p l + 1 := by - have : countP.go p (a :: l) 0 = countP.go p l 1 := show cond .. = _ by rw [pa]; rfl - unfold countP - rw [this, Nat.add_comm, List.countP_go_eq_add] - -@[simp] theorem countP_cons_of_neg (l) (pa : ¬p a) : countP p (a :: l) = countP p l := by - simp [countP, countP.go, pa] - -theorem countP_cons (a : α) (l) : countP p (a :: l) = countP p l + if p a then 1 else 0 := by - by_cases h : p a <;> simp [h] - -theorem length_eq_countP_add_countP (l) : length l = countP p l + countP (fun a => ¬p a) l := by - induction l with - | nil => rfl - | cons x h ih => - if h : p x then - rw [countP_cons_of_pos _ _ h, countP_cons_of_neg _ _ _, length, ih] - · rw [Nat.add_assoc, Nat.add_comm _ 1, Nat.add_assoc] - · simp only [h, not_true_eq_false, decide_False, not_false_eq_true] - else - rw [countP_cons_of_pos (fun a => ¬p a) _ _, countP_cons_of_neg _ _ h, length, ih] - · rfl - · simp only [h, not_false_eq_true, decide_True] - -theorem countP_eq_length_filter (l) : countP p l = length (filter p l) := by - induction l with - | nil => rfl - | cons x l ih => - if h : p x - then rw [countP_cons_of_pos p l h, ih, filter_cons_of_pos l h, length] - else rw [countP_cons_of_neg p l h, ih, filter_cons_of_neg l h] - -theorem countP_le_length : countP p l ≤ l.length := by - simp only [countP_eq_length_filter] - apply length_filter_le - -@[simp] theorem countP_append (l₁ l₂) : countP p (l₁ ++ l₂) = countP p l₁ + countP p l₂ := by - simp only [countP_eq_length_filter, filter_append, length_append] - -theorem countP_pos : 0 < countP p l ↔ ∃ a ∈ l, p a := by - simp only [countP_eq_length_filter, length_pos_iff_exists_mem, mem_filter, exists_prop] - -theorem countP_eq_zero : countP p l = 0 ↔ ∀ a ∈ l, ¬p a := by - simp only [countP_eq_length_filter, length_eq_zero, filter_eq_nil] - -theorem countP_eq_length : countP p l = l.length ↔ ∀ a ∈ l, p a := by - rw [countP_eq_length_filter, filter_length_eq_length] - -theorem Sublist.countP_le (s : l₁ <+ l₂) : countP p l₁ ≤ countP p l₂ := by - simp only [countP_eq_length_filter] - apply s.filter _ |>.length_le - -theorem countP_filter (l : List α) : - countP p (filter q l) = countP (fun a => p a ∧ q a) l := by - simp only [countP_eq_length_filter, filter_filter] - -@[simp] theorem countP_true {l : List α} : (l.countP fun _ => true) = l.length := by - rw [countP_eq_length] - simp - -@[simp] theorem countP_false {l : List α} : (l.countP fun _ => false) = 0 := by - rw [countP_eq_zero] - simp - -@[simp] theorem countP_map (p : β → Bool) (f : α → β) : - ∀ l, countP p (map f l) = countP (p ∘ f) l - | [] => rfl - | a :: l => by rw [map_cons, countP_cons, countP_cons, countP_map p f l]; rfl - -variable {p q} - -theorem countP_mono_left (h : ∀ x ∈ l, p x → q x) : countP p l ≤ countP q l := by - induction l with - | nil => apply Nat.le_refl - | cons a l ihl => - rw [forall_mem_cons] at h - have ⟨ha, hl⟩ := h - simp [countP_cons] - cases h : p a - . simp - apply Nat.le_trans ?_ (Nat.le_add_right _ _) - apply ihl hl - . simp [ha h] - apply ihl hl - -theorem countP_congr (h : ∀ x ∈ l, p x ↔ q x) : countP p l = countP q l := - Nat.le_antisymm - (countP_mono_left fun x hx => (h x hx).1) - (countP_mono_left fun x hx => (h x hx).2) - -end countP /-! ### count -/ @@ -131,71 +26,10 @@ section count variable [DecidableEq α] -@[simp] theorem count_nil (a : α) : count a [] = 0 := rfl - -theorem count_cons (a b : α) (l : List α) : - count a (b :: l) = count a l + if a = b then 1 else 0 := by - simp [count, countP_cons, eq_comm (a := a)] - -@[simp] theorem count_cons_self (a : α) (l : List α) : count a (a :: l) = count a l + 1 := by - simp [count_cons] - -@[simp] theorem count_cons_of_ne (h : a ≠ b) (l : List α) : count a (b :: l) = count a l := by - simp [count_cons, h] - -theorem count_tail : ∀ (l : List α) (a : α) (h : l ≠ []), - l.tail.count a = l.count a - if a = l.head h then 1 else 0 - | head :: tail, a, h => by simp [count_cons] - -theorem count_le_length (a : α) (l : List α) : count a l ≤ l.length := countP_le_length _ - -theorem Sublist.count_le (h : l₁ <+ l₂) (a : α) : count a l₁ ≤ count a l₂ := h.countP_le _ - -theorem count_le_count_cons (a b : α) (l : List α) : count a l ≤ count a (b :: l) := - (sublist_cons _ _).count_le _ - -theorem count_singleton (a : α) : count a [a] = 1 := by simp - -theorem count_singleton' (a b : α) : count a [b] = if a = b then 1 else 0 := by simp [count_cons] - -@[simp] theorem count_append (a : α) : ∀ l₁ l₂, count a (l₁ ++ l₂) = count a l₁ + count a l₂ := - countP_append _ +theorem count_singleton' (a b : α) : count a [b] = if b = a then 1 else 0 := by simp [count_cons] theorem count_concat (a : α) (l : List α) : count a (concat l a) = succ (count a l) := by simp -@[simp] -theorem count_pos_iff_mem {a : α} {l : List α} : 0 < count a l ↔ a ∈ l := by - simp only [count, countP_pos, beq_iff_eq, exists_eq_right] - -@[simp 900] theorem count_eq_zero_of_not_mem {a : α} {l : List α} (h : a ∉ l) : count a l = 0 := - Decidable.byContradiction fun h' => h <| count_pos_iff_mem.1 (Nat.pos_of_ne_zero h') - -theorem not_mem_of_count_eq_zero {a : α} {l : List α} (h : count a l = 0) : a ∉ l := - fun h' => Nat.ne_of_lt (count_pos_iff_mem.2 h') h.symm - -theorem count_eq_zero {l : List α} : count a l = 0 ↔ a ∉ l := - ⟨not_mem_of_count_eq_zero, count_eq_zero_of_not_mem⟩ - -theorem count_eq_length {l : List α} : count a l = l.length ↔ ∀ b ∈ l, a = b := by - rw [count, countP_eq_length] - refine ⟨fun h b hb => Eq.symm ?_, fun h b hb => ?_⟩ - · simpa using h b hb - · rw [h b hb, beq_self_eq_true] - -@[simp] theorem count_replicate_self (a : α) (n : Nat) : count a (replicate n a) = n := - (count_eq_length.2 <| fun _ h => (eq_of_mem_replicate h).symm).trans (length_replicate ..) - -theorem count_replicate (a b : α) (n : Nat) : count a (replicate n b) = if a = b then n else 0 := by - split - exacts [‹a = b› ▸ count_replicate_self .., count_eq_zero.2 <| mt eq_of_mem_replicate ‹a ≠ b›] - -theorem filter_beq (l : List α) (a : α) : l.filter (· == a) = replicate (count a l) a := by - simp only [count, countP_eq_length_filter, eq_replicate, mem_filter, beq_iff_eq] - exact ⟨trivial, fun _ h => h.2⟩ - -theorem filter_eq (l : List α) (a : α) : l.filter (· = a) = replicate (count a l) a := - filter_beq l a - @[deprecated filter_eq (since := "2023-12-14")] theorem filter_eq' (l : List α) (a : α) : l.filter (a = ·) = replicate (count a l) a := by simpa only [eq_comm] using filter_eq l a @@ -203,46 +37,3 @@ theorem filter_eq' (l : List α) (a : α) : l.filter (a = ·) = replicate (count @[deprecated filter_beq (since := "2023-12-14")] theorem filter_beq' (l : List α) (a : α) : l.filter (a == ·) = replicate (count a l) a := by simpa only [eq_comm (b := a)] using filter_eq l a - -theorem le_count_iff_replicate_sublist {l : List α} : n ≤ count a l ↔ replicate n a <+ l := by - refine ⟨fun h => ?_, fun h => ?_⟩ - · exact ((replicate_sublist_replicate a).2 h).trans <| filter_eq l a ▸ filter_sublist _ - · simpa only [count_replicate_self] using h.count_le a - -theorem replicate_count_eq_of_count_eq_length {l : List α} (h : count a l = length l) : - replicate (count a l) a = l := - (le_count_iff_replicate_sublist.mp (Nat.le_refl _)).eq_of_length <| - (length_replicate (count a l) a).trans h - -@[simp] theorem count_filter {l : List α} (h : p a) : count a (filter p l) = count a l := by - rw [count, countP_filter]; congr; funext b - rw [(by rfl : (b == a) = decide (b = a)), decide_eq_decide] - simp; rintro rfl; exact h - -theorem count_le_count_map [DecidableEq β] (l : List α) (f : α → β) (x : α) : - count x l ≤ count (f x) (map f l) := by - rw [count, count, countP_map] - apply countP_mono_left; simp (config := { contextual := true }) - -theorem count_erase (a b : α) : - ∀ l : List α, count a (l.erase b) = count a l - if a = b then 1 else 0 - | [] => by simp - | c :: l => by - rw [erase_cons] - if hc : c = b then - have hc_beq := (beq_iff_eq _ _).mpr hc - rw [if_pos hc_beq, hc, count_cons, Nat.add_sub_cancel] - else - have hc_beq := beq_false_of_ne hc - simp only [hc_beq, if_false, count_cons, count_cons, count_erase a b l] - if ha : a = b then - rw [← ha, eq_comm] at hc - rw [if_pos ha, if_neg hc, Nat.add_zero, Nat.add_zero] - else - rw [if_neg ha, Nat.sub_zero, Nat.sub_zero] - -@[simp] theorem count_erase_self (a : α) (l : List α) : - count a (List.erase l a) = count a l - 1 := by rw [count_erase, if_pos rfl] - -@[simp] theorem count_erase_of_ne (ab : a ≠ b) (l : List α) : count a (l.erase b) = count a l := by - rw [count_erase, if_neg ab, Nat.sub_zero] diff --git a/Batteries/Data/List/EraseIdx.lean b/Batteries/Data/List/EraseIdx.lean index ca3cf2e716..42b4bfacd1 100644 --- a/Batteries/Data/List/EraseIdx.lean +++ b/Batteries/Data/List/EraseIdx.lean @@ -17,75 +17,15 @@ namespace List universe u v variable {α : Type u} {β : Type v} -@[simp] theorem eraseIdx_zero (l : List α) : eraseIdx l 0 = tail l := by cases l <;> rfl - -theorem eraseIdx_eq_take_drop_succ : - ∀ (l : List α) (i : Nat), l.eraseIdx i = l.take i ++ l.drop (i + 1) - | nil, _ => by simp - | a::l, 0 => by simp - | a::l, i + 1 => by simp [eraseIdx_eq_take_drop_succ l i] - -theorem eraseIdx_sublist : ∀ (l : List α) (k : Nat), eraseIdx l k <+ l - | [], _ => by simp - | a::l, 0 => by simp - | a::l, k + 1 => by simp [eraseIdx_sublist l k] - -theorem eraseIdx_subset (l : List α) (k : Nat) : eraseIdx l k ⊆ l := (eraseIdx_sublist l k).subset - -@[simp] -theorem eraseIdx_eq_self : ∀ {l : List α} {k : Nat}, eraseIdx l k = l ↔ length l ≤ k - | [], _ => by simp - | a::l, 0 => by simp [(cons_ne_self _ _).symm] - | a::l, k + 1 => by simp [eraseIdx_eq_self] - -alias ⟨_, eraseIdx_of_length_le⟩ := eraseIdx_eq_self - -theorem eraseIdx_append_of_lt_length {l : List α} {k : Nat} (hk : k < length l) (l' : List α) : - eraseIdx (l ++ l') k = eraseIdx l k ++ l' := by - rw [eraseIdx_eq_take_drop_succ, take_append_of_le_length, drop_append_of_le_length, - eraseIdx_eq_take_drop_succ, append_assoc] - all_goals omega - -theorem eraseIdx_append_of_length_le {l : List α} {k : Nat} (hk : length l ≤ k) (l' : List α) : - eraseIdx (l ++ l') k = l ++ eraseIdx l' (k - length l) := by - rw [eraseIdx_eq_take_drop_succ, eraseIdx_eq_take_drop_succ, - take_append_eq_append_take, drop_append_eq_append_drop, - take_all_of_le hk, drop_eq_nil_of_le (by omega), nil_append, append_assoc] - congr - omega - -protected theorem IsPrefix.eraseIdx {l l' : List α} (h : l <+: l') (k : Nat) : - eraseIdx l k <+: eraseIdx l' k := by - rcases h with ⟨t, rfl⟩ - if hkl : k < length l then - simp [eraseIdx_append_of_lt_length hkl] - else - rw [Nat.not_lt] at hkl - simp [eraseIdx_append_of_length_le hkl, eraseIdx_of_length_le hkl] - -theorem mem_eraseIdx_iff_getElem {x : α} : - ∀ {l} {k}, x ∈ eraseIdx l k ↔ ∃ i : Fin l.length, ↑i ≠ k ∧ l[i.1] = x - | [], _ => by - simp only [eraseIdx, Fin.exists_iff, not_mem_nil, false_iff] - rintro ⟨i, h, -⟩ - exact Nat.not_lt_zero _ h - | a::l, 0 => by simp [Fin.exists_fin_succ, mem_iff_get] - | a::l, k+1 => by - simp [Fin.exists_fin_succ, mem_eraseIdx_iff_getElem, @eq_comm _ a, k.succ_ne_zero.symm] - @[deprecated mem_eraseIdx_iff_getElem (since := "2024-06-12")] theorem mem_eraseIdx_iff_get {x : α} {l} {k} : x ∈ eraseIdx l k ↔ ∃ i : Fin l.length, ↑i ≠ k ∧ l.get i = x := by - simp [mem_eraseIdx_iff_getElem] - -theorem mem_eraseIdx_iff_getElem? {x : α} {l} {k} : x ∈ eraseIdx l k ↔ ∃ i ≠ k, l[i]? = some x := by - simp only [mem_eraseIdx_iff_getElem, getElem_eq_iff, Fin.exists_iff, exists_and_left] - refine exists_congr fun i => and_congr_right' ?_ + simp only [mem_eraseIdx_iff_getElem, ne_eq, exists_and_left, get_eq_getElem] constructor - · rintro ⟨_, h⟩; exact h - · rintro h; - obtain ⟨h', -⟩ := getElem?_eq_some.1 h - exact ⟨h', h⟩ + · rintro ⟨i, h, w, rfl⟩ + exact ⟨⟨i, w⟩, h, rfl⟩ + · rintro ⟨i, h, rfl⟩ + exact ⟨i.1, h, i.2, rfl⟩ @[deprecated mem_eraseIdx_iff_getElem? (since := "2024-06-12")] theorem mem_eraseIdx_iff_get? {x : α} {l} {k} : x ∈ eraseIdx l k ↔ ∃ i ≠ k, l.get? i = x := by diff --git a/Batteries/Data/List/Init/Attach.lean b/Batteries/Data/List/Init/Attach.lean deleted file mode 100644 index d2b2bf0990..0000000000 --- a/Batteries/Data/List/Init/Attach.lean +++ /dev/null @@ -1,44 +0,0 @@ -/- -Copyright (c) 2023 Mario Carneiro. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro --/ - -namespace List - -/-- `O(n)`. Partial map. If `f : Π a, P a → β` is a partial function defined on - `a : α` satisfying `P`, then `pmap f l h` is essentially the same as `map f l` - but is defined only when all members of `l` satisfy `P`, using the proof - to apply `f`. -/ -@[simp] def pmap {P : α → Prop} (f : ∀ a, P a → β) : ∀ l : List α, (H : ∀ a ∈ l, P a) → List β - | [], _ => [] - | a :: l, H => f a (forall_mem_cons.1 H).1 :: pmap f l (forall_mem_cons.1 H).2 - -/-- -Unsafe implementation of `attachWith`, taking advantage of the fact that the representation of -`List {x // P x}` is the same as the input `List α`. -(Someday, the compiler might do this optimization automatically, but until then...) --/ -@[inline] private unsafe def attachWithImpl - (l : List α) (P : α → Prop) (_ : ∀ x ∈ l, P x) : List {x // P x} := unsafeCast l - -/-- `O(1)`. "Attach" a proof `P x` that holds for all the elements of `l` to produce a new list - with the same elements but in the type `{x // P x}`. -/ -@[implemented_by attachWithImpl] def attachWith - (l : List α) (P : α → Prop) (H : ∀ x ∈ l, P x) : List {x // P x} := pmap Subtype.mk l H - -/-- `O(1)`. "Attach" the proof that the elements of `l` are in `l` to produce a new list - with the same elements but in the type `{x // x ∈ l}`. -/ -@[inline] def attach (l : List α) : List {x // x ∈ l} := attachWith l _ fun _ => id - -/-- Implementation of `pmap` using the zero-copy version of `attach`. -/ -@[inline] private def pmapImpl {P : α → Prop} (f : ∀ a, P a → β) (l : List α) (H : ∀ a ∈ l, P a) : - List β := (l.attachWith _ H).map fun ⟨x, h'⟩ => f x h' - -@[csimp] private theorem pmap_eq_pmapImpl : @pmap = @pmapImpl := by - funext α β p f L h' - let rec go : ∀ L' (hL' : ∀ ⦃x⦄, x ∈ L' → p x), - pmap f L' hL' = map (fun ⟨x, hx⟩ => f x hx) (pmap Subtype.mk L' hL') - | nil, hL' => rfl - | cons _ L', hL' => congrArg _ <| go L' fun _ hx => hL' (.tail _ hx) - exact go L h' diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 533a140f45..2e97cb5264 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -17,225 +17,6 @@ open Nat @[simp] theorem mem_toArray {a : α} {l : List α} : a ∈ l.toArray ↔ a ∈ l := by simp [Array.mem_def] -/-! ### drop -/ - -@[simp] -theorem drop_one : ∀ l : List α, drop 1 l = tail l - | [] | _ :: _ => rfl - -/-! ### zipWith -/ - -theorem zipWith_distrib_tail : (zipWith f l l').tail = zipWith f l.tail l'.tail := by - rw [← drop_one]; simp [zipWith_distrib_drop] - -/-! ### List subset -/ - -theorem subset_def {l₁ l₂ : List α} : l₁ ⊆ l₂ ↔ ∀ {a : α}, a ∈ l₁ → a ∈ l₂ := .rfl - -@[simp] theorem nil_subset (l : List α) : [] ⊆ l := nofun - -@[simp] theorem Subset.refl (l : List α) : l ⊆ l := fun _ i => i - -theorem Subset.trans {l₁ l₂ l₃ : List α} (h₁ : l₁ ⊆ l₂) (h₂ : l₂ ⊆ l₃) : l₁ ⊆ l₃ := - fun _ i => h₂ (h₁ i) - -instance : Trans (Membership.mem : α → List α → Prop) Subset Membership.mem := - ⟨fun h₁ h₂ => h₂ h₁⟩ - -instance : Trans (Subset : List α → List α → Prop) Subset Subset := - ⟨Subset.trans⟩ - -@[simp] theorem subset_cons (a : α) (l : List α) : l ⊆ a :: l := fun _ => Mem.tail _ - -theorem subset_of_cons_subset {a : α} {l₁ l₂ : List α} : a :: l₁ ⊆ l₂ → l₁ ⊆ l₂ := - fun s _ i => s (mem_cons_of_mem _ i) - -theorem subset_cons_of_subset (a : α) {l₁ l₂ : List α} : l₁ ⊆ l₂ → l₁ ⊆ a :: l₂ := - fun s _ i => .tail _ (s i) - -theorem cons_subset_cons {l₁ l₂ : List α} (a : α) (s : l₁ ⊆ l₂) : a :: l₁ ⊆ a :: l₂ := - fun _ => by simp only [mem_cons]; exact Or.imp_right (@s _) - -@[simp] theorem subset_append_left (l₁ l₂ : List α) : l₁ ⊆ l₁ ++ l₂ := fun _ => mem_append_left _ - -@[simp] theorem subset_append_right (l₁ l₂ : List α) : l₂ ⊆ l₁ ++ l₂ := fun _ => mem_append_right _ - -theorem subset_append_of_subset_left (l₂ : List α) : l ⊆ l₁ → l ⊆ l₁ ++ l₂ := -fun s => Subset.trans s <| subset_append_left _ _ - -theorem subset_append_of_subset_right (l₁ : List α) : l ⊆ l₂ → l ⊆ l₁ ++ l₂ := -fun s => Subset.trans s <| subset_append_right _ _ - -@[simp] theorem cons_subset : a :: l ⊆ m ↔ a ∈ m ∧ l ⊆ m := by - simp only [subset_def, mem_cons, or_imp, forall_and, forall_eq] - -@[simp] theorem append_subset {l₁ l₂ l : List α} : - l₁ ++ l₂ ⊆ l ↔ l₁ ⊆ l ∧ l₂ ⊆ l := by simp [subset_def, or_imp, forall_and] - -theorem subset_nil {l : List α} : l ⊆ [] ↔ l = [] := - ⟨fun h => match l with | [] => rfl | _::_ => (nomatch h (.head ..)), fun | rfl => Subset.refl _⟩ - -theorem map_subset {l₁ l₂ : List α} (f : α → β) (H : l₁ ⊆ l₂) : map f l₁ ⊆ map f l₂ := - fun x => by simp only [mem_map]; exact .imp fun a => .imp_left (@H _) - -/-! ### sublists -/ - -@[simp] theorem nil_sublist : ∀ l : List α, [] <+ l - | [] => .slnil - | a :: l => (nil_sublist l).cons a - -@[simp] theorem Sublist.refl : ∀ l : List α, l <+ l - | [] => .slnil - | a :: l => (Sublist.refl l).cons₂ a - -theorem Sublist.trans {l₁ l₂ l₃ : List α} (h₁ : l₁ <+ l₂) (h₂ : l₂ <+ l₃) : l₁ <+ l₃ := by - induction h₂ generalizing l₁ with - | slnil => exact h₁ - | cons _ _ IH => exact (IH h₁).cons _ - | @cons₂ l₂ _ a _ IH => - generalize e : a :: l₂ = l₂' - match e ▸ h₁ with - | .slnil => apply nil_sublist - | .cons a' h₁' => cases e; apply (IH h₁').cons - | .cons₂ a' h₁' => cases e; apply (IH h₁').cons₂ - -instance : Trans (@Sublist α) Sublist Sublist := ⟨Sublist.trans⟩ - -@[simp] theorem sublist_cons (a : α) (l : List α) : l <+ a :: l := (Sublist.refl l).cons _ - -theorem sublist_of_cons_sublist : a :: l₁ <+ l₂ → l₁ <+ l₂ := - (sublist_cons a l₁).trans - -@[simp] theorem sublist_append_left : ∀ l₁ l₂ : List α, l₁ <+ l₁ ++ l₂ - | [], _ => nil_sublist _ - | _ :: l₁, l₂ => (sublist_append_left l₁ l₂).cons₂ _ - -@[simp] theorem sublist_append_right : ∀ l₁ l₂ : List α, l₂ <+ l₁ ++ l₂ - | [], _ => Sublist.refl _ - | _ :: l₁, l₂ => (sublist_append_right l₁ l₂).cons _ - -theorem sublist_append_of_sublist_left (s : l <+ l₁) : l <+ l₁ ++ l₂ := - s.trans <| sublist_append_left .. - -theorem sublist_append_of_sublist_right (s : l <+ l₂) : l <+ l₁ ++ l₂ := - s.trans <| sublist_append_right .. - -@[simp] -theorem cons_sublist_cons : a :: l₁ <+ a :: l₂ ↔ l₁ <+ l₂ := - ⟨fun | .cons _ s => sublist_of_cons_sublist s | .cons₂ _ s => s, .cons₂ _⟩ - -@[simp] theorem append_sublist_append_left : ∀ l, l ++ l₁ <+ l ++ l₂ ↔ l₁ <+ l₂ - | [] => Iff.rfl - | _ :: l => cons_sublist_cons.trans (append_sublist_append_left l) - -theorem Sublist.append_left : l₁ <+ l₂ → ∀ l, l ++ l₁ <+ l ++ l₂ := - fun h l => (append_sublist_append_left l).mpr h - -theorem Sublist.append_right : l₁ <+ l₂ → ∀ l, l₁ ++ l <+ l₂ ++ l - | .slnil, _ => Sublist.refl _ - | .cons _ h, _ => (h.append_right _).cons _ - | .cons₂ _ h, _ => (h.append_right _).cons₂ _ - -theorem sublist_or_mem_of_sublist (h : l <+ l₁ ++ a :: l₂) : l <+ l₁ ++ l₂ ∨ a ∈ l := by - induction l₁ generalizing l with - | nil => match h with - | .cons _ h => exact .inl h - | .cons₂ _ h => exact .inr (.head ..) - | cons b l₁ IH => - match h with - | .cons _ h => exact (IH h).imp_left (Sublist.cons _) - | .cons₂ _ h => exact (IH h).imp (Sublist.cons₂ _) (.tail _) - -theorem Sublist.reverse : l₁ <+ l₂ → l₁.reverse <+ l₂.reverse - | .slnil => Sublist.refl _ - | .cons _ h => by rw [reverse_cons]; exact sublist_append_of_sublist_left h.reverse - | .cons₂ _ h => by rw [reverse_cons, reverse_cons]; exact h.reverse.append_right _ - -@[simp] theorem reverse_sublist : l₁.reverse <+ l₂.reverse ↔ l₁ <+ l₂ := - ⟨fun h => l₁.reverse_reverse ▸ l₂.reverse_reverse ▸ h.reverse, Sublist.reverse⟩ - -@[simp] theorem append_sublist_append_right (l) : l₁ ++ l <+ l₂ ++ l ↔ l₁ <+ l₂ := - ⟨fun h => by - have := h.reverse - simp only [reverse_append, append_sublist_append_left, reverse_sublist] at this - exact this, - fun h => h.append_right l⟩ - -theorem Sublist.append (hl : l₁ <+ l₂) (hr : r₁ <+ r₂) : l₁ ++ r₁ <+ l₂ ++ r₂ := - (hl.append_right _).trans ((append_sublist_append_left _).2 hr) - -theorem Sublist.subset : l₁ <+ l₂ → l₁ ⊆ l₂ - | .slnil, _, h => h - | .cons _ s, _, h => .tail _ (s.subset h) - | .cons₂ .., _, .head .. => .head .. - | .cons₂ _ s, _, .tail _ h => .tail _ (s.subset h) - -instance : Trans (@Sublist α) Subset Subset := - ⟨fun h₁ h₂ => trans h₁.subset h₂⟩ - -instance : Trans Subset (@Sublist α) Subset := - ⟨fun h₁ h₂ => trans h₁ h₂.subset⟩ - -instance : Trans (Membership.mem : α → List α → Prop) Sublist Membership.mem := - ⟨fun h₁ h₂ => h₂.subset h₁⟩ - -theorem Sublist.length_le : l₁ <+ l₂ → length l₁ ≤ length l₂ - | .slnil => Nat.le_refl 0 - | .cons _l s => le_succ_of_le (length_le s) - | .cons₂ _ s => succ_le_succ (length_le s) - -@[simp] theorem sublist_nil {l : List α} : l <+ [] ↔ l = [] := - ⟨fun s => subset_nil.1 s.subset, fun H => H ▸ Sublist.refl _⟩ - -theorem Sublist.eq_of_length : l₁ <+ l₂ → length l₁ = length l₂ → l₁ = l₂ - | .slnil, _ => rfl - | .cons a s, h => nomatch Nat.not_lt.2 s.length_le (h ▸ lt_succ_self _) - | .cons₂ a s, h => by rw [s.eq_of_length (succ.inj h)] - -theorem Sublist.eq_of_length_le (s : l₁ <+ l₂) (h : length l₂ ≤ length l₁) : l₁ = l₂ := - s.eq_of_length <| Nat.le_antisymm s.length_le h - -@[simp] theorem singleton_sublist {a : α} {l} : [a] <+ l ↔ a ∈ l := by - refine ⟨fun h => h.subset (mem_singleton_self _), fun h => ?_⟩ - obtain ⟨_, _, rfl⟩ := append_of_mem h - exact ((nil_sublist _).cons₂ _).trans (sublist_append_right ..) - -@[simp] theorem replicate_sublist_replicate {m n} (a : α) : - replicate m a <+ replicate n a ↔ m ≤ n := by - refine ⟨fun h => ?_, fun h => ?_⟩ - · have := h.length_le; simp only [length_replicate] at this ⊢; exact this - · induction h with - | refl => apply Sublist.refl - | step => simp [*, replicate, Sublist.cons] - -theorem isSublist_iff_sublist [BEq α] [LawfulBEq α] {l₁ l₂ : List α} : - l₁.isSublist l₂ ↔ l₁ <+ l₂ := by - cases l₁ <;> cases l₂ <;> simp [isSublist] - case cons.cons hd₁ tl₁ hd₂ tl₂ => - if h_eq : hd₁ = hd₂ then - simp [h_eq, cons_sublist_cons, isSublist_iff_sublist] - else - simp only [beq_iff_eq, h_eq] - constructor - · intro h_sub - apply Sublist.cons - exact isSublist_iff_sublist.mp h_sub - · intro h_sub - cases h_sub - case cons h_sub => - exact isSublist_iff_sublist.mpr h_sub - case cons₂ => - contradiction - -instance [DecidableEq α] (l₁ l₂ : List α) : Decidable (l₁ <+ l₂) := - decidable_of_iff (l₁.isSublist l₂) isSublist_iff_sublist - -/-! ### tail -/ - -theorem tail_eq_tailD (l) : @tail α l = tailD l [] := by cases l <;> rfl - -theorem tail_eq_tail? (l) : @tail α l = (tail? l).getD [] := by simp [tail_eq_tailD] - /-! ### next? -/ @[simp] theorem next?_nil : @next? α [] = none := rfl @@ -243,51 +24,16 @@ theorem tail_eq_tail? (l) : @tail α l = (tail? l).getD [] := by simp [tail_eq_t /-! ### get? -/ -theorem getElem_eq_iff {l : List α} {n : Nat} {h : n < l.length} : l[n] = x ↔ l[n]? = some x := by - simp only [get_eq_getElem, get?_eq_getElem?, getElem?_eq_some] - exact ⟨fun w => ⟨h, w⟩, fun h => h.2⟩ - @[deprecated getElem_eq_iff (since := "2024-06-12")] theorem get_eq_iff : List.get l n = x ↔ l.get? n.1 = some x := by simp -theorem getElem?_inj - (h₀ : i < xs.length) (h₁ : Nodup xs) (h₂ : xs[i]? = xs[j]?) : i = j := by - induction xs generalizing i j with - | nil => cases h₀ - | cons x xs ih => - match i, j with - | 0, 0 => rfl - | i+1, j+1 => simp; cases h₁ with - | cons ha h₁ => - simp only [getElem?_cons_succ] at h₂ - exact ih (Nat.lt_of_succ_lt_succ h₀) h₁ h₂ - | i+1, 0 => ?_ - | 0, j+1 => ?_ - all_goals - simp only [get?_eq_getElem?, getElem?_cons_zero, getElem?_cons_succ] at h₂ - cases h₁; rename_i h' h - have := h x ?_ rfl; cases this - rw [mem_iff_get?] - simp only [get?_eq_getElem?] - exact ⟨_, h₂⟩; exact ⟨_ , h₂.symm⟩ - @[deprecated getElem?_inj (since := "2024-06-12")] theorem get?_inj (h₀ : i < xs.length) (h₁ : Nodup xs) (h₂ : xs.get? i = xs.get? j) : i = j := by apply getElem?_inj h₀ h₁ simp_all -/-! ### drop -/ - -theorem tail_drop (l : List α) (n : Nat) : (l.drop n).tail = l.drop (n + 1) := by - induction l generalizing n with - | nil => simp - | cons hd tl hl => - cases n - · simp - · simp [hl] - /-! ### modifyNth -/ @[simp] theorem modifyNth_nil (f : α → α) (n) : [].modifyNth f n = [] := by cases n <;> rfl @@ -405,18 +151,13 @@ theorem modifyNth_eq_set_get (f : α → α) {n} {l : List α} (h) : l.modifyNth f n = l.set n (f (l.get ⟨n, h⟩)) := by rw [modifyNth_eq_set_get?, get?_eq_get h]; rfl -theorem exists_of_set {l : List α} (h : n < l.length) : +-- The naming of `exists_of_set'` and `exists_of_set` have been swapped. +-- If no one complains, we will remove this version later. +@[deprecated exists_of_set (since := "2024-07-04")] +theorem exists_of_set' {l : List α} (h : n < l.length) : ∃ l₁ a l₂, l = l₁ ++ a :: l₂ ∧ l₁.length = n ∧ l.set n a' = l₁ ++ a' :: l₂ := by rw [set_eq_modifyNth]; exact exists_of_modifyNth _ h -theorem exists_of_set' {l : List α} (h : n < l.length) : - ∃ l₁ l₂, l = l₁ ++ l[n] :: l₂ ∧ l₁.length = n ∧ l.set n a' = l₁ ++ a' :: l₂ := - have ⟨_, _, _, h₁, h₂, h₃⟩ := exists_of_set h; ⟨_, _, getElem_of_append h₁ h₂ ▸ h₁, h₂, h₃⟩ - -@[simp] -theorem getElem?_set_eq' (a : α) (n) (l : List α) : (set l n a)[n]? = (fun _ => a) <$> l[n]? := by - simp only [set_eq_modifyNth, getElem?_modifyNth_eq] - @[deprecated getElem?_set_eq' (since := "2024-06-12")] theorem get?_set_eq (a : α) (n) (l : List α) : (set l n a).get? n = (fun _ => a) <$> l.get? n := by simp @@ -433,10 +174,6 @@ theorem get?_set_eq_of_lt (a : α) {n} {l : List α} (h : n < length l) : theorem get?_set_ne (a : α) {m n} (l : List α) (h : m ≠ n) : (set l m a).get? n = l.get? n := by simp [h] -theorem getElem?_set' (a : α) {m n} (l : List α) : - (set l m a)[n]? = if m = n then (fun _ => a) <$> l[n]? else l[n]? := by - by_cases m = n <;> simp [*] - @[deprecated getElem?_set (since := "2024-06-12")] theorem get?_set (a : α) {m n} (l : List α) : (set l m a).get? n = if m = n then (fun _ => a) <$> l.get? n else l.get? n := by @@ -450,123 +187,10 @@ theorem get?_set_of_lt' (a : α) {m n} (l : List α) (h : m < length l) : (set l m a).get? n = if m = n then some a else l.get? n := by simp [getElem?_set]; split <;> subst_vars <;> simp [*, getElem?_eq_getElem h] -theorem drop_set_of_lt (a : α) {n m : Nat} (l : List α) (h : n < m) : - (l.set n a).drop m = l.drop m := - List.ext_getElem? fun i => by rw [getElem?_drop, getElem?_drop, getElem?_set_ne (by omega)] - -theorem take_set_of_lt (a : α) {n m : Nat} (l : List α) (h : m < n) : - (l.set n a).take m = l.take m := - List.ext_getElem? fun i => by - rw [getElem?_take_eq_if, getElem?_take_eq_if] - split - · next h' => rw [getElem?_set_ne (by omega)] - · rfl - -/-! ### removeNth -/ - -theorem length_eraseIdx : ∀ {l i}, i < length l → length (@eraseIdx α l i) = length l - 1 - | [], _, _ => rfl - | _::_, 0, _ => by simp [eraseIdx] - | x::xs, i+1, h => by - have : i < length xs := Nat.lt_of_succ_lt_succ h - simp [eraseIdx, ← Nat.add_one] - rw [length_eraseIdx this, Nat.sub_add_cancel (Nat.lt_of_le_of_lt (Nat.zero_le _) this)] - @[deprecated (since := "2024-05-06")] alias length_removeNth := length_eraseIdx -/-! ### tail -/ - -@[simp] theorem length_tail (l : List α) : length (tail l) = length l - 1 := by cases l <;> rfl - /-! ### eraseP -/ -@[simp] theorem eraseP_nil : [].eraseP p = [] := rfl - -theorem eraseP_cons (a : α) (l : List α) : - (a :: l).eraseP p = bif p a then l else a :: l.eraseP p := rfl - -@[simp] theorem eraseP_cons_of_pos {l : List α} (p) (h : p a) : (a :: l).eraseP p = l := by - simp [eraseP_cons, h] - -@[simp] theorem eraseP_cons_of_neg {l : List α} (p) (h : ¬p a) : - (a :: l).eraseP p = a :: l.eraseP p := by simp [eraseP_cons, h] - -theorem eraseP_of_forall_not {l : List α} (h : ∀ a, a ∈ l → ¬p a) : l.eraseP p = l := by - induction l with - | nil => rfl - | cons _ _ ih => simp [h _ (.head ..), ih (forall_mem_cons.1 h).2] - -theorem exists_of_eraseP : ∀ {l : List α} {a} (al : a ∈ l) (pa : p a), - ∃ a l₁ l₂, (∀ b ∈ l₁, ¬p b) ∧ p a ∧ l = l₁ ++ a :: l₂ ∧ l.eraseP p = l₁ ++ l₂ - | b :: l, a, al, pa => - if pb : p b then - ⟨b, [], l, forall_mem_nil _, pb, by simp [pb]⟩ - else - match al with - | .head .. => nomatch pb pa - | .tail _ al => - let ⟨c, l₁, l₂, h₁, h₂, h₃, h₄⟩ := exists_of_eraseP al pa - ⟨c, b::l₁, l₂, (forall_mem_cons ..).2 ⟨pb, h₁⟩, - h₂, by rw [h₃, cons_append], by simp [pb, h₄]⟩ - -theorem exists_or_eq_self_of_eraseP (p) (l : List α) : - l.eraseP p = l ∨ - ∃ a l₁ l₂, (∀ b ∈ l₁, ¬p b) ∧ p a ∧ l = l₁ ++ a :: l₂ ∧ l.eraseP p = l₁ ++ l₂ := - if h : ∃ a ∈ l, p a then - let ⟨_, ha, pa⟩ := h - .inr (exists_of_eraseP ha pa) - else - .inl (eraseP_of_forall_not (h ⟨·, ·, ·⟩)) - -@[simp] theorem length_eraseP_of_mem (al : a ∈ l) (pa : p a) : - length (l.eraseP p) = Nat.pred (length l) := by - let ⟨_, l₁, l₂, _, _, e₁, e₂⟩ := exists_of_eraseP al pa - rw [e₂]; simp [length_append, e₁]; rfl - -theorem eraseP_append_left {a : α} (pa : p a) : - ∀ {l₁ : List α} l₂, a ∈ l₁ → (l₁++l₂).eraseP p = l₁.eraseP p ++ l₂ - | x :: xs, l₂, h => by - by_cases h' : p x <;> simp [h'] - rw [eraseP_append_left pa l₂ ((mem_cons.1 h).resolve_left (mt _ h'))] - intro | rfl => exact pa - -theorem eraseP_append_right : - ∀ {l₁ : List α} l₂, (∀ b ∈ l₁, ¬p b) → eraseP p (l₁++l₂) = l₁ ++ l₂.eraseP p - | [], l₂, _ => rfl - | x :: xs, l₂, h => by - simp [(forall_mem_cons.1 h).1, eraseP_append_right _ (forall_mem_cons.1 h).2] - -theorem eraseP_sublist (l : List α) : l.eraseP p <+ l := by - match exists_or_eq_self_of_eraseP p l with - | .inl h => rw [h]; apply Sublist.refl - | .inr ⟨c, l₁, l₂, _, _, h₃, h₄⟩ => rw [h₄, h₃]; simp - -theorem eraseP_subset (l : List α) : l.eraseP p ⊆ l := (eraseP_sublist l).subset - -protected theorem Sublist.eraseP : l₁ <+ l₂ → l₁.eraseP p <+ l₂.eraseP p - | .slnil => Sublist.refl _ - | .cons a s => by - by_cases h : p a <;> simp [h] - exacts [s.eraseP.trans (eraseP_sublist _), s.eraseP.cons _] - | .cons₂ a s => by - by_cases h : p a <;> simp [h] - exacts [s, s.eraseP] - -theorem mem_of_mem_eraseP {l : List α} : a ∈ l.eraseP p → a ∈ l := (eraseP_subset _ ·) - -@[simp] theorem mem_eraseP_of_neg {l : List α} (pa : ¬p a) : a ∈ l.eraseP p ↔ a ∈ l := by - refine ⟨mem_of_mem_eraseP, fun al => ?_⟩ - match exists_or_eq_self_of_eraseP p l with - | .inl h => rw [h]; assumption - | .inr ⟨c, l₁, l₂, h₁, h₂, h₃, h₄⟩ => - rw [h₄]; rw [h₃] at al - have : a ≠ c := fun h => (h ▸ pa).elim h₂ - simp [this] at al; simp [al] - -theorem eraseP_map (f : β → α) : ∀ (l : List β), (map f l).eraseP p = map f (l.eraseP (p ∘ f)) - | [] => rfl - | b::l => by by_cases h : p (f b) <;> simp [h, eraseP_map f l, eraseP_cons_of_pos] - @[simp] theorem extractP_eq_find?_eraseP (l : List α) : extractP p l = (find? p l, eraseP p l) := by let rec go (acc) : ∀ xs, l = acc.data ++ xs → @@ -579,216 +203,8 @@ theorem eraseP_map (f : β → α) : ∀ (l : List β), (map f l).eraseP p = map /-! ### erase -/ -section erase -variable [BEq α] - -theorem erase_eq_eraseP' (a : α) (l : List α) : l.erase a = l.eraseP (· == a) := by - induction l - · simp - · next b t ih => - rw [erase_cons, eraseP_cons, ih] - if h : b == a then simp [h] else simp [h] - -theorem erase_eq_eraseP [LawfulBEq α] (a : α) : ∀ l : List α, l.erase a = l.eraseP (a == ·) - | [] => rfl - | b :: l => by - if h : a = b then simp [h] else simp [h, Ne.symm h, erase_eq_eraseP a l] - -theorem exists_erase_eq [LawfulBEq α] {a : α} {l : List α} (h : a ∈ l) : - ∃ l₁ l₂, a ∉ l₁ ∧ l = l₁ ++ a :: l₂ ∧ l.erase a = l₁ ++ l₂ := by - let ⟨_, l₁, l₂, h₁, e, h₂, h₃⟩ := exists_of_eraseP h (beq_self_eq_true _) - rw [erase_eq_eraseP]; exact ⟨l₁, l₂, fun h => h₁ _ h (beq_self_eq_true _), eq_of_beq e ▸ h₂, h₃⟩ - -@[simp] theorem length_erase_of_mem [LawfulBEq α] {a : α} {l : List α} (h : a ∈ l) : - length (l.erase a) = Nat.pred (length l) := by - rw [erase_eq_eraseP]; exact length_eraseP_of_mem h (beq_self_eq_true a) - -theorem erase_append_left [LawfulBEq α] {l₁ : List α} (l₂) (h : a ∈ l₁) : - (l₁ ++ l₂).erase a = l₁.erase a ++ l₂ := by - simp [erase_eq_eraseP]; exact eraseP_append_left (beq_self_eq_true a) l₂ h - -theorem erase_append_right [LawfulBEq α] {a : α} {l₁ : List α} (l₂ : List α) (h : a ∉ l₁) : - (l₁ ++ l₂).erase a = (l₁ ++ l₂.erase a) := by - rw [erase_eq_eraseP, erase_eq_eraseP, eraseP_append_right] - intros b h' h''; rw [eq_of_beq h''] at h; exact h h' - -theorem erase_sublist (a : α) (l : List α) : l.erase a <+ l := - erase_eq_eraseP' a l ▸ eraseP_sublist l - -theorem erase_subset (a : α) (l : List α) : l.erase a ⊆ l := (erase_sublist a l).subset - -theorem Sublist.erase (a : α) {l₁ l₂ : List α} (h : l₁ <+ l₂) : l₁.erase a <+ l₂.erase a := by - simp only [erase_eq_eraseP']; exact h.eraseP @[deprecated (since := "2024-04-22")] alias sublist.erase := Sublist.erase -theorem mem_of_mem_erase {a b : α} {l : List α} (h : a ∈ l.erase b) : a ∈ l := erase_subset _ _ h - -@[simp] theorem mem_erase_of_ne [LawfulBEq α] {a b : α} {l : List α} (ab : a ≠ b) : - a ∈ l.erase b ↔ a ∈ l := - erase_eq_eraseP b l ▸ mem_eraseP_of_neg (mt eq_of_beq ab.symm) - -theorem erase_comm [LawfulBEq α] (a b : α) (l : List α) : - (l.erase a).erase b = (l.erase b).erase a := by - if ab : a == b then rw [eq_of_beq ab] else ?_ - if ha : a ∈ l then ?_ else - simp only [erase_of_not_mem ha, erase_of_not_mem (mt mem_of_mem_erase ha)] - if hb : b ∈ l then ?_ else - simp only [erase_of_not_mem hb, erase_of_not_mem (mt mem_of_mem_erase hb)] - match l, l.erase a, exists_erase_eq ha with - | _, _, ⟨l₁, l₂, ha', rfl, rfl⟩ => - if h₁ : b ∈ l₁ then - rw [erase_append_left _ h₁, erase_append_left _ h₁, - erase_append_right _ (mt mem_of_mem_erase ha'), erase_cons_head] - else - rw [erase_append_right _ h₁, erase_append_right _ h₁, erase_append_right _ ha', - erase_cons_tail _ ab, erase_cons_head] - -end erase - -/-! ### filter and partition -/ - -@[simp] theorem filter_sublist {p : α → Bool} : ∀ (l : List α), filter p l <+ l - | [] => .slnil - | a :: l => by rw [filter]; split <;> simp [Sublist.cons, Sublist.cons₂, filter_sublist l] - -/-! ### filterMap -/ - -protected theorem Sublist.filterMap (f : α → Option β) (s : l₁ <+ l₂) : - filterMap f l₁ <+ filterMap f l₂ := by - induction s <;> simp [filterMap_cons] <;> split <;> simp [*, cons, cons₂] - -theorem Sublist.filter (p : α → Bool) {l₁ l₂} (s : l₁ <+ l₂) : filter p l₁ <+ filter p l₂ := by - rw [← filterMap_eq_filter]; apply s.filterMap - -/-! ### findIdx -/ - -@[simp] theorem findIdx_nil {α : Type _} (p : α → Bool) : [].findIdx p = 0 := rfl - -theorem findIdx_cons (p : α → Bool) (b : α) (l : List α) : - (b :: l).findIdx p = bif p b then 0 else (l.findIdx p) + 1 := by - cases H : p b with - | true => simp [H, findIdx, findIdx.go] - | false => simp [H, findIdx, findIdx.go, findIdx_go_succ] -where - findIdx_go_succ (p : α → Bool) (l : List α) (n : Nat) : - List.findIdx.go p l (n + 1) = (findIdx.go p l n) + 1 := by - cases l with - | nil => unfold findIdx.go; exact Nat.succ_eq_add_one n - | cons head tail => - unfold findIdx.go - cases p head <;> simp only [cond_false, cond_true] - exact findIdx_go_succ p tail (n + 1) - -theorem findIdx_of_get?_eq_some {xs : List α} (w : xs.get? (xs.findIdx p) = some y) : p y := by - induction xs with - | nil => simp_all - | cons x xs ih => by_cases h : p x <;> simp_all [findIdx_cons] - -theorem findIdx_get {xs : List α} {w : xs.findIdx p < xs.length} : - p (xs.get ⟨xs.findIdx p, w⟩) := - xs.findIdx_of_get?_eq_some (get?_eq_get w) - -theorem findIdx_lt_length_of_exists {xs : List α} (h : ∃ x ∈ xs, p x) : - xs.findIdx p < xs.length := by - induction xs with - | nil => simp_all - | cons x xs ih => - by_cases p x - · simp_all only [forall_exists_index, and_imp, mem_cons, exists_eq_or_imp, true_or, - findIdx_cons, cond_true, length_cons] - apply Nat.succ_pos - · simp_all [findIdx_cons] - refine Nat.succ_lt_succ ?_ - obtain ⟨x', m', h'⟩ := h - exact ih x' m' h' - -theorem findIdx_get?_eq_get_of_exists {xs : List α} (h : ∃ x ∈ xs, p x) : - xs.get? (xs.findIdx p) = some (xs.get ⟨xs.findIdx p, xs.findIdx_lt_length_of_exists h⟩) := - get?_eq_get (findIdx_lt_length_of_exists h) - - /-! ### findIdx? -/ - -@[simp] theorem findIdx?_nil : ([] : List α).findIdx? p i = none := rfl - -@[simp] theorem findIdx?_cons : - (x :: xs).findIdx? p i = if p x then some i else findIdx? p xs (i + 1) := rfl - -@[simp] theorem findIdx?_succ : - (xs : List α).findIdx? p (i+1) = (xs.findIdx? p i).map fun i => i + 1 := by - induction xs generalizing i with simp - | cons _ _ _ => split <;> simp_all - -theorem findIdx?_eq_some_iff (xs : List α) (p : α → Bool) : - xs.findIdx? p = some i ↔ (xs.take (i + 1)).map p = replicate i false ++ [true] := by - induction xs generalizing i with - | nil => simp - | cons x xs ih => - simp only [findIdx?_cons, Nat.zero_add, findIdx?_succ, take_succ_cons, map_cons] - split <;> cases i <;> simp_all [replicate_succ] - -theorem findIdx?_of_eq_some {xs : List α} {p : α → Bool} (w : xs.findIdx? p = some i) : - match xs.get? i with | some a => p a | none => false := by - induction xs generalizing i with - | nil => simp_all - | cons x xs ih => - simp_all only [findIdx?_cons, Nat.zero_add, findIdx?_succ] - split at w <;> cases i <;> simp_all - -theorem findIdx?_of_eq_none {xs : List α} {p : α → Bool} (w : xs.findIdx? p = none) : - ∀ i, match xs.get? i with | some a => ¬ p a | none => true := by - intro i - induction xs generalizing i with - | nil => simp_all - | cons x xs ih => - simp_all only [Bool.not_eq_true, findIdx?_cons, Nat.zero_add, findIdx?_succ] - cases i with - | zero => - split at w <;> simp_all - | succ i => - simp only [get?_cons_succ] - apply ih - split at w <;> simp_all - -@[simp] theorem findIdx?_append : - (xs ++ ys : List α).findIdx? p = - (xs.findIdx? p <|> (ys.findIdx? p).map fun i => i + xs.length) := by - induction xs with simp - | cons _ _ _ => split <;> simp_all [Option.map_orElse, Option.map_map]; rfl - -@[simp] theorem findIdx?_replicate : - (replicate n a).findIdx? p = if 0 < n ∧ p a then some 0 else none := by - induction n with - | zero => simp - | succ n ih => - simp only [replicate, findIdx?_cons, Nat.zero_add, findIdx?_succ, Nat.zero_lt_succ, true_and] - split <;> simp_all - -/-! ### pairwise -/ - -theorem Pairwise.sublist : l₁ <+ l₂ → l₂.Pairwise R → l₁.Pairwise R - | .slnil, h => h - | .cons _ s, .cons _ h₂ => h₂.sublist s - | .cons₂ _ s, .cons h₁ h₂ => (h₂.sublist s).cons fun _ h => h₁ _ (s.subset h) - -theorem pairwise_map {l : List α} : - (l.map f).Pairwise R ↔ l.Pairwise fun a b => R (f a) (f b) := by - induction l - · simp - · simp only [map, pairwise_cons, forall_mem_map_iff, *] - -theorem pairwise_append {l₁ l₂ : List α} : - (l₁ ++ l₂).Pairwise R ↔ l₁.Pairwise R ∧ l₂.Pairwise R ∧ ∀ a ∈ l₁, ∀ b ∈ l₂, R a b := by - induction l₁ <;> simp [*, or_imp, forall_and, and_assoc, and_left_comm] - -theorem pairwise_reverse {l : List α} : - l.reverse.Pairwise R ↔ l.Pairwise (fun a b => R b a) := by - induction l <;> simp [*, pairwise_append, and_comm] - -theorem Pairwise.imp {α R S} (H : ∀ {a b}, R a b → S a b) : - ∀ {l : List α}, l.Pairwise R → l.Pairwise S - | _, .nil => .nil - | _, .cons h₁ h₂ => .cons (H ∘ h₁ ·) (h₂.imp H) - /-! ### replaceF -/ theorem replaceF_nil : [].replaceF p = [] := rfl @@ -857,10 +273,10 @@ theorem disjoint_of_subset_right (ss : l₂ ⊆ l) (d : Disjoint l₁ l) : Disjo fun _ m m₁ => d m (ss m₁) theorem disjoint_of_disjoint_cons_left {l₁ l₂} : Disjoint (a :: l₁) l₂ → Disjoint l₁ l₂ := -disjoint_of_subset_left (subset_cons _ _) + disjoint_of_subset_left (subset_cons_self _ _) theorem disjoint_of_disjoint_cons_right {l₁ l₂} : Disjoint l₁ (a :: l₂) → Disjoint l₁ l₂ := -disjoint_of_subset_right (subset_cons _ _) + disjoint_of_subset_right (subset_cons_self _ _) @[simp] theorem disjoint_nil_left (l : List α) : Disjoint [] l := fun a => (not_mem_nil a).elim @@ -896,16 +312,6 @@ theorem disjoint_of_disjoint_append_right_left (d : Disjoint l (l₁ ++ l₂)) : theorem disjoint_of_disjoint_append_right_right (d : Disjoint l (l₁ ++ l₂)) : Disjoint l l₂ := (disjoint_append_right.1 d).2 -/-! ### foldl / foldr -/ - -theorem foldl_hom (f : α₁ → α₂) (g₁ : α₁ → β → α₁) (g₂ : α₂ → β → α₂) (l : List β) (init : α₁) - (H : ∀ x y, g₂ (f x) y = f (g₁ x y)) : l.foldl g₂ (f init) = f (l.foldl g₁ init) := by - induction l generalizing init <;> simp [*, H] - -theorem foldr_hom (f : β₁ → β₂) (g₁ : α → β₁ → β₁) (g₂ : α → β₂ → β₂) (l : List α) (init : β₁) - (H : ∀ x y, g₂ x (f y) = f (g₁ x y)) : l.foldr g₂ (f init) = f (l.foldr g₁ init) := by - induction l <;> simp [*, H] - /-! ### union -/ section union @@ -941,38 +347,14 @@ theorem pair_mem_product {xs : List α} {ys : List β} {x : α} {y : β} : simp only [product, and_imp, mem_map, Prod.mk.injEq, exists_eq_right_right, mem_bind, iff_self] -/-! ### leftpad -/ - -/-- The length of the List returned by `List.leftpad n a l` is equal - to the larger of `n` and `l.length` -/ -@[simp] -theorem leftpad_length (n : Nat) (a : α) (l : List α) : - (leftpad n a l).length = max n l.length := by - simp only [leftpad, length_append, length_replicate, Nat.sub_add_eq_max] - -theorem leftpad_prefix (n : Nat) (a : α) (l : List α) : - replicate (n - length l) a <+: leftpad n a l := by - simp only [IsPrefix, leftpad] - exact Exists.intro l rfl - -theorem leftpad_suffix (n : Nat) (a : α) (l : List α) : l <:+ (leftpad n a l) := by - simp only [IsSuffix, leftpad] - exact Exists.intro (replicate (n - length l) a) rfl - /-! ### monadic operations -/ --- we use ForIn.forIn as the simp normal form -@[simp] theorem forIn_eq_forIn [Monad m] : @List.forIn α β m _ = forIn := rfl - theorem forIn_eq_bindList [Monad m] [LawfulMonad m] (f : α → β → m (ForInStep β)) (l : List α) (init : β) : forIn l init f = ForInStep.run <$> (ForInStep.yield init).bindList f l := by induction l generalizing init <;> simp [*, map_eq_pure_bind] congr; ext (b | b) <;> simp -@[simp] theorem forM_append [Monad m] [LawfulMonad m] (l₁ l₂ : List α) (f : α → m PUnit) : - (l₁ ++ l₂).forM f = (do l₁.forM f; l₂.forM f) := by induction l₁ <;> simp [*] - /-! ### diff -/ section Diff @@ -1050,205 +432,8 @@ theorem Sublist.erase_diff_erase_sublist {a : α} : end Diff -/-! ### prefix, suffix, infix -/ - -@[simp] theorem prefix_append (l₁ l₂ : List α) : l₁ <+: l₁ ++ l₂ := ⟨l₂, rfl⟩ - -@[simp] theorem suffix_append (l₁ l₂ : List α) : l₂ <:+ l₁ ++ l₂ := ⟨l₁, rfl⟩ - -theorem infix_append (l₁ l₂ l₃ : List α) : l₂ <:+: l₁ ++ l₂ ++ l₃ := ⟨l₁, l₃, rfl⟩ - -@[simp] theorem infix_append' (l₁ l₂ l₃ : List α) : l₂ <:+: l₁ ++ (l₂ ++ l₃) := by - rw [← List.append_assoc]; apply infix_append - -theorem IsPrefix.isInfix : l₁ <+: l₂ → l₁ <:+: l₂ := fun ⟨t, h⟩ => ⟨[], t, h⟩ - -theorem IsSuffix.isInfix : l₁ <:+ l₂ → l₁ <:+: l₂ := fun ⟨t, h⟩ => ⟨t, [], by rw [h, append_nil]⟩ - -theorem nil_prefix (l : List α) : [] <+: l := ⟨l, rfl⟩ - -theorem nil_suffix (l : List α) : [] <:+ l := ⟨l, append_nil _⟩ - -theorem nil_infix (l : List α) : [] <:+: l := (nil_prefix _).isInfix - -theorem prefix_refl (l : List α) : l <+: l := ⟨[], append_nil _⟩ - -theorem suffix_refl (l : List α) : l <:+ l := ⟨[], rfl⟩ - -theorem infix_refl (l : List α) : l <:+: l := (prefix_refl l).isInfix - -@[simp] theorem suffix_cons (a : α) : ∀ l, l <:+ a :: l := suffix_append [a] - -theorem infix_cons : l₁ <:+: l₂ → l₁ <:+: a :: l₂ := fun ⟨L₁, L₂, h⟩ => ⟨a :: L₁, L₂, h ▸ rfl⟩ - -theorem infix_concat : l₁ <:+: l₂ → l₁ <:+: concat l₂ a := fun ⟨L₁, L₂, h⟩ => - ⟨L₁, concat L₂ a, by simp [← h, concat_eq_append, append_assoc]⟩ - -theorem IsPrefix.trans : ∀ {l₁ l₂ l₃ : List α}, l₁ <+: l₂ → l₂ <+: l₃ → l₁ <+: l₃ - | _, _, _, ⟨r₁, rfl⟩, ⟨r₂, rfl⟩ => ⟨r₁ ++ r₂, (append_assoc _ _ _).symm⟩ - -theorem IsSuffix.trans : ∀ {l₁ l₂ l₃ : List α}, l₁ <:+ l₂ → l₂ <:+ l₃ → l₁ <:+ l₃ - | _, _, _, ⟨l₁, rfl⟩, ⟨l₂, rfl⟩ => ⟨l₂ ++ l₁, append_assoc _ _ _⟩ - -theorem IsInfix.trans : ∀ {l₁ l₂ l₃ : List α}, l₁ <:+: l₂ → l₂ <:+: l₃ → l₁ <:+: l₃ - | l, _, _, ⟨l₁, r₁, rfl⟩, ⟨l₂, r₂, rfl⟩ => ⟨l₂ ++ l₁, r₁ ++ r₂, by simp only [append_assoc]⟩ - -protected theorem IsInfix.sublist : l₁ <:+: l₂ → l₁ <+ l₂ - | ⟨_, _, h⟩ => h ▸ (sublist_append_right ..).trans (sublist_append_left ..) - -protected theorem IsInfix.subset (hl : l₁ <:+: l₂) : l₁ ⊆ l₂ := - hl.sublist.subset - -protected theorem IsPrefix.sublist (h : l₁ <+: l₂) : l₁ <+ l₂ := - h.isInfix.sublist - -protected theorem IsPrefix.subset (hl : l₁ <+: l₂) : l₁ ⊆ l₂ := - hl.sublist.subset - -protected theorem IsSuffix.sublist (h : l₁ <:+ l₂) : l₁ <+ l₂ := - h.isInfix.sublist - -protected theorem IsSuffix.subset (hl : l₁ <:+ l₂) : l₁ ⊆ l₂ := - hl.sublist.subset - -@[simp] theorem reverse_suffix : reverse l₁ <:+ reverse l₂ ↔ l₁ <+: l₂ := - ⟨fun ⟨r, e⟩ => ⟨reverse r, by rw [← reverse_reverse l₁, ← reverse_append, e, reverse_reverse]⟩, - fun ⟨r, e⟩ => ⟨reverse r, by rw [← reverse_append, e]⟩⟩ - -@[simp] theorem reverse_prefix : reverse l₁ <+: reverse l₂ ↔ l₁ <:+ l₂ := by - rw [← reverse_suffix]; simp only [reverse_reverse] - -@[simp] theorem reverse_infix : reverse l₁ <:+: reverse l₂ ↔ l₁ <:+: l₂ := by - refine ⟨fun ⟨s, t, e⟩ => ⟨reverse t, reverse s, ?_⟩, fun ⟨s, t, e⟩ => ⟨reverse t, reverse s, ?_⟩⟩ - · rw [← reverse_reverse l₁, append_assoc, ← reverse_append, ← reverse_append, e, - reverse_reverse] - · rw [append_assoc, ← reverse_append, ← reverse_append, e] - -theorem IsInfix.length_le (h : l₁ <:+: l₂) : l₁.length ≤ l₂.length := - h.sublist.length_le - -theorem IsPrefix.length_le (h : l₁ <+: l₂) : l₁.length ≤ l₂.length := - h.sublist.length_le - -theorem IsSuffix.length_le (h : l₁ <:+ l₂) : l₁.length ≤ l₂.length := - h.sublist.length_le - -@[simp] theorem infix_nil : l <:+: [] ↔ l = [] := ⟨(sublist_nil.1 ·.sublist), (· ▸ infix_refl _)⟩ - -@[simp] theorem prefix_nil : l <+: [] ↔ l = [] := ⟨(sublist_nil.1 ·.sublist), (· ▸ prefix_refl _)⟩ - -@[simp] theorem suffix_nil : l <:+ [] ↔ l = [] := ⟨(sublist_nil.1 ·.sublist), (· ▸ suffix_refl _)⟩ - -theorem infix_iff_prefix_suffix (l₁ l₂ : List α) : l₁ <:+: l₂ ↔ ∃ t, l₁ <+: t ∧ t <:+ l₂ := - ⟨fun ⟨_, t, e⟩ => ⟨l₁ ++ t, ⟨_, rfl⟩, e ▸ append_assoc .. ▸ ⟨_, rfl⟩⟩, - fun ⟨_, ⟨t, rfl⟩, s, e⟩ => ⟨s, t, append_assoc .. ▸ e⟩⟩ - -theorem IsInfix.eq_of_length (h : l₁ <:+: l₂) : l₁.length = l₂.length → l₁ = l₂ := - h.sublist.eq_of_length - -theorem IsPrefix.eq_of_length (h : l₁ <+: l₂) : l₁.length = l₂.length → l₁ = l₂ := - h.sublist.eq_of_length - -theorem IsSuffix.eq_of_length (h : l₁ <:+ l₂) : l₁.length = l₂.length → l₁ = l₂ := - h.sublist.eq_of_length - -theorem prefix_of_prefix_length_le : - ∀ {l₁ l₂ l₃ : List α}, l₁ <+: l₃ → l₂ <+: l₃ → length l₁ ≤ length l₂ → l₁ <+: l₂ - | [], l₂, _, _, _, _ => nil_prefix _ - | a :: l₁, b :: l₂, _, ⟨r₁, rfl⟩, ⟨r₂, e⟩, ll => by - injection e with _ e'; subst b - rcases prefix_of_prefix_length_le ⟨_, rfl⟩ ⟨_, e'⟩ (le_of_succ_le_succ ll) with ⟨r₃, rfl⟩ - exact ⟨r₃, rfl⟩ - -theorem prefix_or_prefix_of_prefix (h₁ : l₁ <+: l₃) (h₂ : l₂ <+: l₃) : l₁ <+: l₂ ∨ l₂ <+: l₁ := - (Nat.le_total (length l₁) (length l₂)).imp (prefix_of_prefix_length_le h₁ h₂) - (prefix_of_prefix_length_le h₂ h₁) - -theorem suffix_of_suffix_length_le - (h₁ : l₁ <:+ l₃) (h₂ : l₂ <:+ l₃) (ll : length l₁ ≤ length l₂) : l₁ <:+ l₂ := - reverse_prefix.1 <| - prefix_of_prefix_length_le (reverse_prefix.2 h₁) (reverse_prefix.2 h₂) (by simp [ll]) - -theorem suffix_or_suffix_of_suffix (h₁ : l₁ <:+ l₃) (h₂ : l₂ <:+ l₃) : l₁ <:+ l₂ ∨ l₂ <:+ l₁ := - (prefix_or_prefix_of_prefix (reverse_prefix.2 h₁) (reverse_prefix.2 h₂)).imp reverse_prefix.1 - reverse_prefix.1 - -theorem suffix_cons_iff : l₁ <:+ a :: l₂ ↔ l₁ = a :: l₂ ∨ l₁ <:+ l₂ := by - constructor - · rintro ⟨⟨hd, tl⟩, hl₃⟩ - · exact Or.inl hl₃ - · simp only [cons_append] at hl₃ - injection hl₃ with _ hl₄ - exact Or.inr ⟨_, hl₄⟩ - · rintro (rfl | hl₁) - · exact (a :: l₂).suffix_refl - · exact hl₁.trans (l₂.suffix_cons _) - -theorem infix_cons_iff : l₁ <:+: a :: l₂ ↔ l₁ <+: a :: l₂ ∨ l₁ <:+: l₂ := by - constructor - · rintro ⟨⟨hd, tl⟩, t, hl₃⟩ - · exact Or.inl ⟨t, hl₃⟩ - · simp only [cons_append] at hl₃ - injection hl₃ with _ hl₄ - exact Or.inr ⟨_, t, hl₄⟩ - · rintro (h | hl₁) - · exact h.isInfix - · exact infix_cons hl₁ - -theorem infix_of_mem_join : ∀ {L : List (List α)}, l ∈ L → l <:+: join L - | l' :: _, h => - match h with - | List.Mem.head .. => infix_append [] _ _ - | List.Mem.tail _ hlMemL => - IsInfix.trans (infix_of_mem_join hlMemL) <| (suffix_append _ _).isInfix - -theorem prefix_append_right_inj (l) : l ++ l₁ <+: l ++ l₂ ↔ l₁ <+: l₂ := - exists_congr fun r => by rw [append_assoc, append_right_inj] - -@[simp] -theorem prefix_cons_inj (a) : a :: l₁ <+: a :: l₂ ↔ l₁ <+: l₂ := - prefix_append_right_inj [a] - -theorem take_prefix (n) (l : List α) : take n l <+: l := - ⟨_, take_append_drop _ _⟩ - -theorem drop_suffix (n) (l : List α) : drop n l <:+ l := - ⟨_, take_append_drop _ _⟩ - -theorem take_sublist (n) (l : List α) : take n l <+ l := - (take_prefix n l).sublist - -theorem drop_sublist (n) (l : List α) : drop n l <+ l := - (drop_suffix n l).sublist - -theorem take_subset (n) (l : List α) : take n l ⊆ l := - (take_sublist n l).subset - -theorem drop_subset (n) (l : List α) : drop n l ⊆ l := - (drop_sublist n l).subset - -theorem mem_of_mem_take {l : List α} (h : a ∈ l.take n) : a ∈ l := - take_subset n l h - -theorem IsPrefix.filter (p : α → Bool) ⦃l₁ l₂ : List α⦄ (h : l₁ <+: l₂) : - l₁.filter p <+: l₂.filter p := by - obtain ⟨xs, rfl⟩ := h - rw [filter_append]; apply prefix_append - -theorem IsSuffix.filter (p : α → Bool) ⦃l₁ l₂ : List α⦄ (h : l₁ <:+ l₂) : - l₁.filter p <:+ l₂.filter p := by - obtain ⟨xs, rfl⟩ := h - rw [filter_append]; apply suffix_append - -theorem IsInfix.filter (p : α → Bool) ⦃l₁ l₂ : List α⦄ (h : l₁ <:+: l₂) : - l₁.filter p <:+: l₂.filter p := by - obtain ⟨xs, ys, rfl⟩ := h - rw [filter_append, filter_append]; apply infix_append _ - /-! ### drop -/ -theorem mem_of_mem_drop {n} {l : List α} (h : a ∈ l.drop n) : a ∈ l := drop_subset _ _ h - theorem disjoint_take_drop : ∀ {l : List α}, l.Nodup → m ≤ n → Disjoint (l.take m) (l.drop n) | [], _, _ => by simp | x :: xs, hl, h => by @@ -1297,39 +482,6 @@ protected theorem Pairwise.chain (p : Pairwise R (a :: l)) : Chain R a l := by /-! ### range', range -/ -theorem range'_succ (s n step) : range' s (n + 1) step = s :: range' (s + step) n step := by - simp [range', Nat.add_succ, Nat.mul_succ] - -@[simp] theorem length_range' (s step) : ∀ n : Nat, length (range' s n step) = n - | 0 => rfl - | _ + 1 => congrArg succ (length_range' _ _ _) - -@[simp] theorem range'_eq_nil : range' s n step = [] ↔ n = 0 := by - rw [← length_eq_zero, length_range'] - -theorem mem_range' : ∀{n}, m ∈ range' s n step ↔ ∃ i < n, m = s + step * i - | 0 => by simp [range', Nat.not_lt_zero] - | n + 1 => by - have h (i) : i ≤ n ↔ i = 0 ∨ ∃ j, i = succ j ∧ j < n := by cases i <;> simp [Nat.succ_le] - simp [range', mem_range', Nat.lt_succ, h]; simp only [← exists_and_right, and_assoc] - rw [exists_comm]; simp [Nat.mul_succ, Nat.add_assoc, Nat.add_comm] - -@[simp] theorem mem_range'_1 : m ∈ range' s n ↔ s ≤ m ∧ m < s + n := by - simp [mem_range']; exact ⟨ - fun ⟨i, h, e⟩ => e ▸ ⟨Nat.le_add_right .., Nat.add_lt_add_left h _⟩, - fun ⟨h₁, h₂⟩ => ⟨m - s, Nat.sub_lt_left_of_lt_add h₁ h₂, (Nat.add_sub_cancel' h₁).symm⟩⟩ - -@[simp] -theorem map_add_range' (a) : ∀ s n step, map (a + ·) (range' s n step) = range' (a + s) n step - | _, 0, _ => rfl - | s, n + 1, step => by simp [range', map_add_range' _ (s + step) n step, Nat.add_assoc] - -theorem map_sub_range' (a s n : Nat) (h : a ≤ s) : - map (· - a) (range' s n step) = range' (s - a) n step := by - conv => lhs; rw [← Nat.add_sub_cancel' h] - rw [← map_add_range', map_map, (?_ : _∘_ = _), map_id] - funext x; apply Nat.add_sub_cancel_left - theorem chain_succ_range' : ∀ s n step : Nat, Chain (fun a b => b = a + step) s (range' (s + step) n step) | _, 0, _ => Chain.nil @@ -1339,41 +491,6 @@ theorem chain_lt_range' (s n : Nat) {step} (h : 0 < step) : Chain (· < ·) s (range' (s + step) n step) := (chain_succ_range' s n step).imp fun _ _ e => e.symm ▸ Nat.lt_add_of_pos_right h -theorem range'_append : ∀ s m n step : Nat, - range' s m step ++ range' (s + step * m) n step = range' s (n + m) step - | s, 0, n, step => rfl - | s, m + 1, n, step => by - simpa [range', Nat.mul_succ, Nat.add_assoc, Nat.add_comm] - using range'_append (s + step) m n step - -@[simp] theorem range'_append_1 (s m n : Nat) : - range' s m ++ range' (s + m) n = range' s (n + m) := by simpa using range'_append s m n 1 - -theorem range'_sublist_right {s m n : Nat} : range' s m step <+ range' s n step ↔ m ≤ n := - ⟨fun h => by simpa only [length_range'] using h.length_le, - fun h => by rw [← Nat.sub_add_cancel h, ← range'_append]; apply sublist_append_left⟩ - -theorem range'_subset_right {s m n : Nat} (step0 : 0 < step) : - range' s m step ⊆ range' s n step ↔ m ≤ n := by - refine ⟨fun h => Nat.le_of_not_lt fun hn => ?_, fun h => (range'_sublist_right.2 h).subset⟩ - have ⟨i, h', e⟩ := mem_range'.1 <| h <| mem_range'.2 ⟨_, hn, rfl⟩ - exact Nat.ne_of_gt h' (Nat.eq_of_mul_eq_mul_left step0 (Nat.add_left_cancel e)) - -theorem range'_subset_right_1 {s m n : Nat} : range' s m ⊆ range' s n ↔ m ≤ n := - range'_subset_right (by decide) - -theorem getElem?_range' (s step) : - ∀ {m n : Nat}, m < n → (range' s n step)[m]? = some (s + step * m) - | 0, n + 1, _ => by simp [range'_succ] - | m + 1, n + 1, h => by - simp only [range'_succ, getElem?_cons_succ] - exact (getElem?_range' (s + step) step (Nat.lt_of_add_lt_add_right h)).trans <| by - simp [Nat.mul_succ, Nat.add_assoc, Nat.add_comm] - -@[simp] theorem getElem_range' {n m step} (i) (H : i < (range' n m step).length) : - (range' n m step)[i] = n + step * i := - (getElem?_eq_some.1 <| getElem?_range' n step (by simpa using H)).2 - @[deprecated getElem?_range' (since := "2024-06-12")] theorem get?_range' (s step) {m n : Nat} (h : m < n) : get? (range' s n step) m = some (s + step * m) := by @@ -1384,54 +501,6 @@ theorem get_range' {n m step} (i) (H : i < (range' n m step).length) : get (range' n m step) ⟨i, H⟩ = n + step * i := by simp -theorem range'_concat (s n : Nat) : range' s (n + 1) step = range' s n step ++ [s + step * n] := by - rw [Nat.add_comm n 1]; exact (range'_append s n 1 step).symm - -theorem range'_1_concat (s n : Nat) : range' s (n + 1) = range' s n ++ [s + n] := by - simp [range'_concat] - -theorem range_loop_range' : ∀ s n : Nat, range.loop s (range' s n) = range' 0 (n + s) - | 0, n => rfl - | s + 1, n => by rw [← Nat.add_assoc, Nat.add_right_comm n s 1]; exact range_loop_range' s (n + 1) - -theorem range_eq_range' (n : Nat) : range n = range' 0 n := - (range_loop_range' n 0).trans <| by rw [Nat.zero_add] - -theorem range_succ_eq_map (n : Nat) : range (n + 1) = 0 :: map succ (range n) := by - rw [range_eq_range', range_eq_range', range', Nat.add_comm, ← map_add_range'] - congr; exact funext one_add - -theorem range'_eq_map_range (s n : Nat) : range' s n = map (s + ·) (range n) := by - rw [range_eq_range', map_add_range']; rfl - -@[simp] theorem length_range (n : Nat) : length (range n) = n := by - simp only [range_eq_range', length_range'] - -@[simp] theorem range_eq_nil {n : Nat} : range n = [] ↔ n = 0 := by - rw [← length_eq_zero, length_range] - -@[simp] -theorem range_sublist {m n : Nat} : range m <+ range n ↔ m ≤ n := by - simp only [range_eq_range', range'_sublist_right] - -@[simp] -theorem range_subset {m n : Nat} : range m ⊆ range n ↔ m ≤ n := by - simp only [range_eq_range', range'_subset_right, lt_succ_self] - -@[simp] -theorem mem_range {m n : Nat} : m ∈ range n ↔ m < n := by - simp only [range_eq_range', mem_range'_1, Nat.zero_le, true_and, Nat.zero_add] - -theorem not_mem_range_self {n : Nat} : n ∉ range n := by simp - -theorem self_mem_range_succ (n : Nat) : n ∈ range (n + 1) := by simp - -theorem getElem?_range {m n : Nat} (h : m < n) : (range n)[m]? = some m := by - simp [range_eq_range', getElem?_range' _ _ h] - -@[simp] theorem getElem_range {n : Nat} (m) (h : m < (range n).length) : (range n)[m] = m := by - simp [range_eq_range'] - @[deprecated getElem?_range (since := "2024-06-12")] theorem get?_range {m n : Nat} (h : m < n) : get? (range n) m = some m := by simp [getElem?_range, h] @@ -1440,41 +509,6 @@ theorem get?_range {m n : Nat} (h : m < n) : get? (range n) m = some m := by theorem get_range {n} (i) (H : i < (range n).length) : get (range n) ⟨i, H⟩ = i := by simp -theorem range_succ (n : Nat) : range (succ n) = range n ++ [n] := by - simp only [range_eq_range', range'_1_concat, Nat.zero_add] - -theorem range_add (a b : Nat) : range (a + b) = range a ++ (range b).map (a + ·) := by - rw [← range'_eq_map_range] - simpa [range_eq_range', Nat.add_comm] using (range'_append_1 0 a b).symm - -theorem iota_eq_reverse_range' : ∀ n : Nat, iota n = reverse (range' 1 n) - | 0 => rfl - | n + 1 => by simp [iota, range'_concat, iota_eq_reverse_range' n, reverse_append, Nat.add_comm] - -@[simp] theorem length_iota (n : Nat) : length (iota n) = n := by simp [iota_eq_reverse_range'] - -@[simp] -theorem mem_iota {m n : Nat} : m ∈ iota n ↔ 1 ≤ m ∧ m ≤ n := by - simp [iota_eq_reverse_range', Nat.add_comm, Nat.lt_succ] - -theorem reverse_range' : ∀ s n : Nat, reverse (range' s n) = map (s + n - 1 - ·) (range n) - | s, 0 => rfl - | s, n + 1 => by - rw [range'_1_concat, reverse_append, range_succ_eq_map, - show s + (n + 1) - 1 = s + n from rfl, map, map_map] - simp [reverse_range', Nat.sub_right_comm, Nat.sub_sub] - - -/-! ### enum, enumFrom -/ - -@[simp] theorem enumFrom_map_fst (n) : - ∀ (l : List α), map Prod.fst (enumFrom n l) = range' n l.length - | [] => rfl - | _ :: _ => congrArg (cons _) (enumFrom_map_fst _ _) - -@[simp] theorem enum_map_fst (l : List α) : map Prod.fst (enum l) = range l.length := by - simp only [enum, enumFrom_map_fst, range_eq_range'] - /-! ### indexOf and indexesOf -/ theorem foldrIdx_start : @@ -1514,13 +548,6 @@ theorem indexesOf_cons [BEq α] : (x :: xs : List α).indexesOf y = bif x == y then 0 :: (xs.indexesOf y).map (· + 1) else (xs.indexesOf y).map (· + 1) := by simp [indexesOf, findIdxs_cons] -@[simp] theorem indexOf_nil [BEq α] : ([] : List α).indexOf x = 0 := rfl - -theorem indexOf_cons [BEq α] : - (x :: xs : List α).indexOf y = bif x == y then 0 else xs.indexOf y + 1 := by - dsimp [indexOf] - simp [findIdx_cons] - theorem indexOf_mem_indexesOf [BEq α] [LawfulBEq α] {xs : List α} (m : x ∈ xs) : xs.indexOf x ∈ xs.indexesOf x := by induction xs with diff --git a/Batteries/Data/List/Pairwise.lean b/Batteries/Data/List/Pairwise.lean index 83b935f8ae..dd94875ca9 100644 --- a/Batteries/Data/List/Pairwise.lean +++ b/Batteries/Data/List/Pairwise.lean @@ -30,164 +30,6 @@ namespace List /-! ### Pairwise -/ -theorem rel_of_pairwise_cons (p : (a :: l).Pairwise R) : ∀ {a'}, a' ∈ l → R a a' := - (pairwise_cons.1 p).1 _ - -theorem Pairwise.of_cons (p : (a :: l).Pairwise R) : Pairwise R l := - (pairwise_cons.1 p).2 - -theorem Pairwise.tail : ∀ {l : List α} (_p : Pairwise R l), Pairwise R l.tail - | [], h => h - | _ :: _, h => h.of_cons - -theorem Pairwise.drop : ∀ {l : List α} {n : Nat}, List.Pairwise R l → List.Pairwise R (l.drop n) - | _, 0, h => h - | [], _ + 1, _ => List.Pairwise.nil - | _ :: _, n + 1, h => Pairwise.drop (n := n) (pairwise_cons.mp h).right - -theorem Pairwise.imp_of_mem {S : α → α → Prop} - (H : ∀ {a b}, a ∈ l → b ∈ l → R a b → S a b) (p : Pairwise R l) : Pairwise S l := by - induction p with - | nil => constructor - | @cons a l r _ ih => - constructor - · exact fun x h => H (mem_cons_self ..) (mem_cons_of_mem _ h) <| r x h - · exact ih fun m m' => H (mem_cons_of_mem _ m) (mem_cons_of_mem _ m') - -theorem Pairwise.and (hR : Pairwise R l) (hS : Pairwise S l) : - l.Pairwise fun a b => R a b ∧ S a b := by - induction hR with - | nil => simp only [Pairwise.nil] - | cons R1 _ IH => - simp only [Pairwise.nil, pairwise_cons] at hS ⊢ - exact ⟨fun b bl => ⟨R1 b bl, hS.1 b bl⟩, IH hS.2⟩ - -theorem pairwise_and_iff : l.Pairwise (fun a b => R a b ∧ S a b) ↔ Pairwise R l ∧ Pairwise S l := - ⟨fun h => ⟨h.imp fun h => h.1, h.imp fun h => h.2⟩, fun ⟨hR, hS⟩ => hR.and hS⟩ - -theorem Pairwise.imp₂ (H : ∀ a b, R a b → S a b → T a b) - (hR : Pairwise R l) (hS : l.Pairwise S) : l.Pairwise T := - (hR.and hS).imp fun ⟨h₁, h₂⟩ => H _ _ h₁ h₂ - -theorem Pairwise.iff_of_mem {S : α → α → Prop} {l : List α} - (H : ∀ {a b}, a ∈ l → b ∈ l → (R a b ↔ S a b)) : Pairwise R l ↔ Pairwise S l := - ⟨Pairwise.imp_of_mem fun m m' => (H m m').1, Pairwise.imp_of_mem fun m m' => (H m m').2⟩ - -theorem Pairwise.iff {S : α → α → Prop} (H : ∀ a b, R a b ↔ S a b) {l : List α} : - Pairwise R l ↔ Pairwise S l := - Pairwise.iff_of_mem fun _ _ => H .. - -theorem pairwise_of_forall {l : List α} (H : ∀ x y, R x y) : Pairwise R l := by - induction l <;> simp [*] - -theorem Pairwise.and_mem {l : List α} : - Pairwise R l ↔ Pairwise (fun x y => x ∈ l ∧ y ∈ l ∧ R x y) l := - Pairwise.iff_of_mem <| by simp (config := { contextual := true }) - -theorem Pairwise.imp_mem {l : List α} : - Pairwise R l ↔ Pairwise (fun x y => x ∈ l → y ∈ l → R x y) l := - Pairwise.iff_of_mem <| by simp (config := { contextual := true }) - -theorem Pairwise.forall_of_forall_of_flip (h₁ : ∀ x ∈ l, R x x) (h₂ : Pairwise R l) - (h₃ : l.Pairwise (flip R)) : ∀ ⦃x⦄, x ∈ l → ∀ ⦃y⦄, y ∈ l → R x y := by - induction l with - | nil => exact forall_mem_nil _ - | cons a l ih => - rw [pairwise_cons] at h₂ h₃ - simp only [mem_cons] - rintro x (rfl | hx) y (rfl | hy) - · exact h₁ _ (l.mem_cons_self _) - · exact h₂.1 _ hy - · exact h₃.1 _ hx - · exact ih (fun x hx => h₁ _ <| mem_cons_of_mem _ hx) h₂.2 h₃.2 hx hy - -theorem pairwise_singleton (R) (a : α) : Pairwise R [a] := by simp - -theorem pairwise_pair {a b : α} : Pairwise R [a, b] ↔ R a b := by simp - -theorem pairwise_append_comm {R : α → α → Prop} (s : ∀ {x y}, R x y → R y x) {l₁ l₂ : List α} : - Pairwise R (l₁ ++ l₂) ↔ Pairwise R (l₂ ++ l₁) := by - have (l₁ l₂ : List α) (H : ∀ x : α, x ∈ l₁ → ∀ y : α, y ∈ l₂ → R x y) - (x : α) (xm : x ∈ l₂) (y : α) (ym : y ∈ l₁) : R x y := s (H y ym x xm) - simp only [pairwise_append, and_left_comm]; rw [Iff.intro (this l₁ l₂) (this l₂ l₁)] - -theorem pairwise_middle {R : α → α → Prop} (s : ∀ {x y}, R x y → R y x) {a : α} {l₁ l₂ : List α} : - Pairwise R (l₁ ++ a :: l₂) ↔ Pairwise R (a :: (l₁ ++ l₂)) := by - show Pairwise R (l₁ ++ ([a] ++ l₂)) ↔ Pairwise R ([a] ++ l₁ ++ l₂) - rw [← append_assoc, pairwise_append, @pairwise_append _ _ ([a] ++ l₁), pairwise_append_comm s] - simp only [mem_append, or_comm] - -theorem Pairwise.of_map {S : β → β → Prop} (f : α → β) (H : ∀ a b : α, S (f a) (f b) → R a b) - (p : Pairwise S (map f l)) : Pairwise R l := - (pairwise_map.1 p).imp (H _ _) - -theorem Pairwise.map {S : β → β → Prop} (f : α → β) (H : ∀ a b : α, R a b → S (f a) (f b)) - (p : Pairwise R l) : Pairwise S (map f l) := - pairwise_map.2 <| p.imp (H _ _) - -theorem pairwise_filterMap (f : β → Option α) {l : List β} : - Pairwise R (filterMap f l) ↔ Pairwise (fun a a' : β => ∀ b ∈ f a, ∀ b' ∈ f a', R b b') l := by - let _S (a a' : β) := ∀ b ∈ f a, ∀ b' ∈ f a', R b b' - simp only [Option.mem_def] - induction l with - | nil => simp only [filterMap, Pairwise.nil] - | cons a l IH => ?_ - match e : f a with - | none => - rw [filterMap_cons_none _ _ e, pairwise_cons] - simp only [e, false_implies, implies_true, true_and, IH] - | some b => - rw [filterMap_cons_some _ _ _ e] - simpa [IH, e] using fun _ => - ⟨fun h a ha b hab => h _ _ ha hab, fun h a b ha hab => h _ ha _ hab⟩ - -theorem Pairwise.filter_map {S : β → β → Prop} (f : α → Option β) - (H : ∀ a a' : α, R a a' → ∀ b ∈ f a, ∀ b' ∈ f a', S b b') {l : List α} (p : Pairwise R l) : - Pairwise S (filterMap f l) := - (pairwise_filterMap _).2 <| p.imp (H _ _) - -theorem pairwise_filter (p : α → Prop) [DecidablePred p] {l : List α} : - Pairwise R (filter p l) ↔ Pairwise (fun x y => p x → p y → R x y) l := by - rw [← filterMap_eq_filter, pairwise_filterMap] - simp - -theorem Pairwise.filter (p : α → Bool) : Pairwise R l → Pairwise R (filter p l) := - Pairwise.sublist (filter_sublist _) - -theorem pairwise_join {L : List (List α)} : - Pairwise R (join L) ↔ - (∀ l ∈ L, Pairwise R l) ∧ Pairwise (fun l₁ l₂ => ∀ x ∈ l₁, ∀ y ∈ l₂, R x y) L := by - induction L with - | nil => simp - | cons l L IH => - simp only [join, pairwise_append, IH, mem_join, exists_imp, and_imp, forall_mem_cons, - pairwise_cons, and_assoc, and_congr_right_iff] - rw [and_comm, and_congr_left_iff] - intros; exact ⟨fun h a b c d e => h c d e a b, fun h c d e a b => h a b c d e⟩ - -theorem pairwise_bind {R : β → β → Prop} {l : List α} {f : α → List β} : - List.Pairwise R (l.bind f) ↔ - (∀ a ∈ l, Pairwise R (f a)) ∧ Pairwise (fun a₁ a₂ => ∀ x ∈ f a₁, ∀ y ∈ f a₂, R x y) l := by - simp [List.bind, pairwise_join, pairwise_map] - -theorem pairwise_iff_forall_sublist : l.Pairwise R ↔ (∀ {a b}, [a,b] <+ l → R a b) := by - induction l with - | nil => simp - | cons hd tl IH => - rw [List.pairwise_cons] - constructor <;> intro h - · intro - | a, b, .cons _ hab => exact IH.mp h.2 hab - | _, b, .cons₂ _ hab => refine h.1 _ (hab.subset ?_); simp - · constructor - · intro x hx - apply h - rw [List.cons_sublist_cons, List.singleton_sublist] - exact hx - · apply IH.mpr - intro a b hab - apply h; exact hab.cons _ - @[deprecated pairwise_iff_forall_sublist (since := "2023-09-18")] theorem pairwise_of_reflexive_on_dupl_of_forall_ne [DecidableEq α] {l : List α} {r : α → α → Prop} (hr : ∀ a, 1 < count a l → r a a) (h : ∀ a ∈ l, ∀ b ∈ l, a ≠ b → r a b) : l.Pairwise r := by @@ -200,53 +42,6 @@ theorem pairwise_of_reflexive_on_dupl_of_forall_ne [DecidableEq α] {l : List α apply h <;> try (apply hab.subset; simp) exact heq -/-- given a list `is` of monotonically increasing indices into `l`, getting each index - produces a sublist of `l`. -/ -theorem map_get_sublist {l : List α} {is : List (Fin l.length)} (h : is.Pairwise (·.val < ·.val)) : - is.map (get l) <+ l := by - suffices ∀ n l', l' = l.drop n → (∀ i ∈ is, n ≤ i) → map (get l) is <+ l' - from this 0 l (by simp) (by simp) - intro n l' hl' his - induction is generalizing n l' with - | nil => simp - | cons hd tl IH => - simp; cases hl' - have := IH h.of_cons (hd+1) _ rfl (pairwise_cons.mp h).1 - specialize his hd (.head _) - have := (drop_eq_getElem_cons ..).symm ▸ this.cons₂ (get l hd) - have := Sublist.append (nil_sublist (take hd l |>.drop n)) this - rwa [nil_append, ← (drop_append_of_le_length ?_), take_append_drop] at this - simp [Nat.min_eq_left (Nat.le_of_lt hd.isLt), his] - -/-- given a sublist `l' <+ l`, there exists a list of indices `is` such that - `l' = map (get l) is`. -/ -theorem sublist_eq_map_get (h : l' <+ l) : ∃ is : List (Fin l.length), - l' = map (get l) is ∧ is.Pairwise (· < ·) := by - induction h with - | slnil => exact ⟨[], by simp⟩ - | cons _ _ IH => - let ⟨is, IH⟩ := IH - refine ⟨is.map (·.succ), ?_⟩ - simp [comp, pairwise_map] - exact IH - | cons₂ _ _ IH => - rcases IH with ⟨is,IH⟩ - refine ⟨⟨0, by simp [Nat.zero_lt_succ]⟩ :: is.map (·.succ), ?_⟩ - simp [comp_def, pairwise_map, IH, ← get_eq_getElem] - -theorem pairwise_iff_getElem : Pairwise R l ↔ - ∀ (i j : Nat) (_hi : i < l.length) (_hj : j < l.length) (_hij : i < j), R l[i] l[j] := by - rw [pairwise_iff_forall_sublist] - constructor <;> intro h - · intros i j hi hj h' - apply h - simpa [h'] using map_get_sublist (is := [⟨i, hi⟩, ⟨j, hj⟩]) - · intros a b h' - have ⟨is, h', hij⟩ := sublist_eq_map_get h' - rcases is with ⟨⟩ | ⟨a', ⟨⟩ | ⟨b', ⟨⟩⟩⟩ <;> simp at h' - rcases h' with ⟨rfl, rfl⟩ - apply h; simpa using hij - theorem pairwise_iff_get : Pairwise R l ↔ ∀ (i j) (_hij : i < j), R (get l i) (get l j) := by rw [pairwise_iff_getElem] constructor <;> intro h @@ -255,11 +50,6 @@ theorem pairwise_iff_get : Pairwise R l ↔ ∀ (i j) (_hij : i < j), R (get l i · intros i j hi hj h' exact h ⟨i, hi⟩ ⟨j, hj⟩ h' -theorem pairwise_replicate {α : Type _} {r : α → α → Prop} {x : α} (hx : r x x) : - ∀ n : Nat, Pairwise r (List.replicate n x) - | 0 => by simp - | n + 1 => by simp [mem_replicate, hx, pairwise_replicate hx n, replicate_succ] - /-! ### Pairwise filtering -/ @[simp] theorem pwFilter_nil [DecidableRel R] : pwFilter R [] = [] := rfl diff --git a/Batteries/Data/List/Perm.lean b/Batteries/Data/List/Perm.lean index ab12db10b1..5c31720c7c 100644 --- a/Batteries/Data/List/Perm.lean +++ b/Batteries/Data/List/Perm.lean @@ -4,10 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Batteries.Tactic.Alias -import Batteries.Data.List.Init.Attach import Batteries.Data.List.Pairwise --- Adaptation note: nightly-2024-03-18. We should be able to remove this after nightly-2024-03-19. -import Lean.Elab.Tactic.Rfl /-! # List Permutations @@ -234,7 +231,7 @@ theorem Subperm.trans {l₁ l₂ l₃ : List α} (s₁₂ : l₁ <+~ l₂) (s₂ ⟨l₁', p₁, s₁.trans s₂⟩ theorem Subperm.cons_right {α : Type _} {l l' : List α} (x : α) (h : l <+~ l') : l <+~ x :: l' := - h.trans (sublist_cons x l').subperm + h.trans (sublist_cons_self x l').subperm theorem Subperm.length_le {l₁ l₂ : List α} : l₁ <+~ l₂ → length l₁ ≤ length l₂ | ⟨_l, p, s⟩ => p.length_eq ▸ s.length_le @@ -371,7 +368,7 @@ theorem perm_append_right_iff {l₁ l₂ : List α} (l) : l₁ ++ l ~ l₂ ++ l theorem subperm_cons (a : α) {l₁ l₂ : List α} : a :: l₁ <+~ a :: l₂ ↔ l₁ <+~ l₂ := by refine ⟨fun ⟨l, p, s⟩ => ?_, fun ⟨l, p, s⟩ => ⟨a :: l, p.cons a, s.cons₂ _⟩⟩ match s with - | .cons _ s' => exact (p.subperm_left.2 <| (sublist_cons _ _).subperm).trans s'.subperm + | .cons _ s' => exact (p.subperm_left.2 <| (sublist_cons_self _ _).subperm).trans s'.subperm | .cons₂ _ s' => exact ⟨_, p.cons_inv, s'⟩ /-- Weaker version of `Subperm.cons_left` -/ @@ -412,7 +409,7 @@ theorem Subperm.exists_of_length_lt {l₁ l₂ : List α} (s : l₁ <+~ l₂) (h | slnil => cases h | cons a s IH => match Nat.lt_or_eq_of_le (Nat.le_of_lt_succ h) with - | .inl h => exact (IH h).imp fun a s => s.trans (sublist_cons _ _).subperm + | .inl h => exact (IH h).imp fun a s => s.trans (sublist_cons_self _ _).subperm | .inr h => exact ⟨a, s.eq_of_length h ▸ .refl _⟩ | cons₂ b _ IH => exact (IH <| Nat.lt_of_succ_lt_succ h).imp fun a s => @@ -440,12 +437,12 @@ theorem Nodup.perm_iff_eq_of_sublist {l₁ l₂ l : List α} (d : Nodup l) | .cons _ s₁ => exact IH d.2 s₁ h | .cons₂ _ s₁ => have := Subperm.subset ⟨_, h.symm, s₂⟩ (.head _) - exact (d.1 _ this rfl).elim + exact (d.1 this).elim | cons₂ a _ IH => match s₁ with | .cons _ s₁ => have := Subperm.subset ⟨_, h, s₁⟩ (.head _) - exact (d.1 _ this rfl).elim + exact (d.1 this).elim | .cons₂ _ s₁ => rw [IH d.2 s₁ h.cons_inv] section DecidableEq @@ -464,7 +461,7 @@ theorem subperm_cons_erase (a : α) (l : List α) : l <+~ a :: l.erase a := if h : a ∈ l then (perm_cons_erase h).subperm else - (erase_of_not_mem h).symm ▸ (sublist_cons _ _).subperm + (erase_of_not_mem h).symm ▸ (sublist_cons_self _ _).subperm theorem erase_subperm (a : α) (l : List α) : l.erase a <+~ l := (erase_sublist _ _).subperm @@ -513,7 +510,7 @@ theorem erase_cons_subperm_cons_erase (a b : α) (l : List α) : rw [h, erase_cons_head]; apply subperm_cons_erase else have : ¬(a == b) = true := by simp only [beq_false_of_ne h, not_false_eq_true] - rw [erase_cons_tail _ this] + rw [erase_cons_tail this] theorem subperm_cons_diff {a : α} {l₁ l₂ : List α} : (a :: l₁).diff l₂ <+~ a :: l₁.diff l₂ := by induction l₂ with diff --git a/Batteries/Data/Nat/Basic.lean b/Batteries/Data/Nat/Basic.lean index a6e55a6ea2..1d9658f49e 100644 --- a/Batteries/Data/Nat/Basic.lean +++ b/Batteries/Data/Nat/Basic.lean @@ -94,9 +94,6 @@ protected def casesDiagOn {motive : Nat → Nat → Sort _} (m n : Nat) Nat.recDiag zero_zero (fun _ _ => zero_succ _) (fun _ _ => succ_zero _) (fun _ _ _ => succ_succ _ _) m n -/-- Sum of a list of natural numbers. -/ -protected def sum (l : List Nat) : Nat := l.foldr (·+·) 0 - /-- Integer square root function. Implemented via Newton's method. -/ diff --git a/Batteries/Data/Nat/Lemmas.lean b/Batteries/Data/Nat/Lemmas.lean index 842ca09e95..1c3edddeaf 100644 --- a/Batteries/Data/Nat/Lemmas.lean +++ b/Batteries/Data/Nat/Lemmas.lean @@ -177,10 +177,6 @@ protected alias le_of_le_of_sub_le_sub_left := Nat.le_of_sub_le_sub_left /-! ### sum -/ -@[simp] theorem sum_nil : Nat.sum [] = 0 := rfl - -@[simp] theorem sum_cons : Nat.sum (a :: l) = a + Nat.sum l := rfl - @[simp] theorem sum_append : Nat.sum (l₁ ++ l₂) = Nat.sum l₁ + Nat.sum l₂ := by induction l₁ <;> simp [*, Nat.add_assoc] diff --git a/Batteries/Data/RBMap/Lemmas.lean b/Batteries/Data/RBMap/Lemmas.lean index 37f24ebea9..5577b3a1a7 100644 --- a/Batteries/Data/RBMap/Lemmas.lean +++ b/Batteries/Data/RBMap/Lemmas.lean @@ -152,7 +152,6 @@ theorem min?_eq_toList_head? {t : RBNode α} : t.min? = t.toList.head? := by | nil => rfl | node _ l _ _ ih => cases l <;> simp [RBNode.min?, ih] - next ll _ _ => cases toList ll <;> rfl theorem max?_eq_toList_getLast? {t : RBNode α} : t.max? = t.toList.getLast? := by rw [← min?_reverse, min?_eq_toList_head?]; simp diff --git a/Batteries/Data/Range/Lemmas.lean b/Batteries/Data/Range/Lemmas.lean index 586f32a7e8..d71082848f 100644 --- a/Batteries/Data/Range/Lemmas.lean +++ b/Batteries/Data/Range/Lemmas.lean @@ -5,7 +5,6 @@ Authors: Mario Carneiro -/ import Batteries.Tactic.SeqFocus import Batteries.Data.List.Lemmas -import Batteries.Data.List.Init.Attach namespace Std.Range diff --git a/Batteries/Data/String/Lemmas.lean b/Batteries/Data/String/Lemmas.lean index e046783fee..9984d0fb52 100644 --- a/Batteries/Data/String/Lemmas.lean +++ b/Batteries/Data/String/Lemmas.lean @@ -12,7 +12,8 @@ import Batteries.Tactic.SqueezeScope namespace String -attribute [ext] ext +-- TODO(kmill): add `@[ext]` attribute to `String.ext` in core. +attribute [ext (iff := false)] ext theorem lt_trans {s₁ s₂ s₃ : String} : s₁ < s₂ → s₂ < s₃ → s₁ < s₃ := List.lt_trans' (α := Char) Nat.lt_trans @@ -87,7 +88,8 @@ end namespace Pos -attribute [ext] ext +-- TODO(kmill): add `@[ext]` attribute to `String.Pos.ext` in core. +attribute [ext (iff := false)] ext theorem lt_addChar (p : Pos) (c : Char) : p < p + c := Nat.lt_add_of_pos_right (Char.utf8Size_pos _) @@ -335,7 +337,7 @@ theorem firstDiffPos_loop_eq (l₁ l₂ r₁ r₂ stop p) · next h => rw [dif_neg] <;> simp [hstop, ← hl₁, ← hl₂, -Nat.not_lt, Nat.lt_min] intro h₁ h₂ - have : ∀ {cs}, p < p + utf8Len cs → cs ≠ [] := by rintro _ h rfl; simp at h + have : ∀ {cs}, 0 < utf8Len cs → cs ≠ [] := by rintro _ h rfl; simp at h obtain ⟨a, as, e₁⟩ := List.exists_cons_of_ne_nil (this h₁) obtain ⟨b, bs, e₂⟩ := List.exists_cons_of_ne_nil (this h₂) exact h _ _ _ _ e₁ e₂ @@ -420,8 +422,9 @@ theorem splitAux_of_valid (p l m r acc) : splitAux ⟨l ++ m ++ r⟩ p ⟨utf8Len l⟩ ⟨utf8Len l + utf8Len m⟩ acc = acc.reverse ++ (List.splitOnP.go p r m.reverse).map mk := by unfold splitAux - simp only [List.append_assoc, atEnd_iff, endPos_eq, utf8Len_append, Pos.mk_le_mk, by - simpa using atEnd_of_valid (l ++ m) r, List.reverse_cons, dite_eq_ite] + simp only [List.append_assoc, atEnd_iff, endPos_eq, utf8Len_append, Pos.mk_le_mk, + Nat.add_le_add_iff_left, by simpa using atEnd_of_valid (l ++ m) r, List.reverse_cons, + dite_eq_ite] split · subst r; simpa [List.splitOnP.go] using extract_of_valid l m [] · obtain ⟨c, r, rfl⟩ := r.exists_cons_of_ne_nil ‹_› @@ -674,7 +677,7 @@ theorem foldrAux_of_valid (f : Char → α → α) (l m r a) : rw [← m.reverse_reverse] induction m.reverse generalizing r a with (unfold foldrAux; simp) | cons c m IH => - rw [if_pos (by exact Nat.lt_add_of_pos_right add_utf8Size_pos)] + rw [if_pos add_utf8Size_pos] simp only [← Nat.add_assoc, by simpa using prev_of_valid (l ++ m.reverse) c r] simp only [by simpa using get_of_valid (l ++ m.reverse) (c :: r)] simpa using IH (c::r) (f c a) diff --git a/Batteries/Data/UInt.lean b/Batteries/Data/UInt.lean index a9ce1b561a..44ebfd9b80 100644 --- a/Batteries/Data/UInt.lean +++ b/Batteries/Data/UInt.lean @@ -10,8 +10,6 @@ import Batteries.Classes.Order @[ext] theorem UInt8.ext : {x y : UInt8} → x.toNat = y.toNat → x = y | ⟨⟨_,_⟩⟩, ⟨⟨_,_⟩⟩, rfl => rfl -theorem UInt8.ext_iff {x y : UInt8} : x = y ↔ x.toNat = y.toNat := ⟨congrArg _, UInt8.ext⟩ - @[simp] theorem UInt8.val_val_eq_toNat (x : UInt8) : x.val.val = x.toNat := rfl @[simp] theorem UInt8.val_ofNat (n) : @@ -57,8 +55,6 @@ instance : Batteries.LawfulOrd UInt8 := .compareOfLessAndEq @[ext] theorem UInt16.ext : {x y : UInt16} → x.toNat = y.toNat → x = y | ⟨⟨_,_⟩⟩, ⟨⟨_,_⟩⟩, rfl => rfl -theorem UInt16.ext_iff {x y : UInt16} : x = y ↔ x.toNat = y.toNat := ⟨congrArg _, UInt16.ext⟩ - @[simp] theorem UInt16.val_val_eq_toNat (x : UInt16) : x.val.val = x.toNat := rfl @[simp] theorem UInt16.val_ofNat (n) : @@ -104,8 +100,6 @@ instance : Batteries.LawfulOrd UInt16 := .compareOfLessAndEq @[ext] theorem UInt32.ext : {x y : UInt32} → x.toNat = y.toNat → x = y | ⟨⟨_,_⟩⟩, ⟨⟨_,_⟩⟩, rfl => rfl -theorem UInt32.ext_iff {x y : UInt32} : x = y ↔ x.toNat = y.toNat := ⟨congrArg _, UInt32.ext⟩ - @[simp] theorem UInt32.val_val_eq_toNat (x : UInt32) : x.val.val = x.toNat := rfl @[simp] theorem UInt32.val_ofNat (n) : @@ -151,8 +145,6 @@ instance : Batteries.LawfulOrd UInt32 := .compareOfLessAndEq @[ext] theorem UInt64.ext : {x y : UInt64} → x.toNat = y.toNat → x = y | ⟨⟨_,_⟩⟩, ⟨⟨_,_⟩⟩, rfl => rfl -theorem UInt64.ext_iff {x y : UInt64} : x = y ↔ x.toNat = y.toNat := ⟨congrArg _, UInt64.ext⟩ - @[simp] theorem UInt64.val_val_eq_toNat (x : UInt64) : x.val.val = x.toNat := rfl @[simp] theorem UInt64.val_ofNat (n) : @@ -198,8 +190,6 @@ instance : Batteries.LawfulOrd UInt64 := .compareOfLessAndEq @[ext] theorem USize.ext : {x y : USize} → x.toNat = y.toNat → x = y | ⟨⟨_,_⟩⟩, ⟨⟨_,_⟩⟩, rfl => rfl -theorem USize.ext_iff {x y : USize} : x = y ↔ x.toNat = y.toNat := ⟨congrArg _, USize.ext⟩ - @[simp] theorem USize.val_val_eq_toNat (x : USize) : x.val.val = x.toNat := rfl @[simp] theorem USize.val_ofNat (n) : diff --git a/Batteries/Lean/SMap.lean b/Batteries/Lean/SMap.lean deleted file mode 100644 index 44a660e2e0..0000000000 --- a/Batteries/Lean/SMap.lean +++ /dev/null @@ -1,17 +0,0 @@ -/- -Copyright (c) 2023 Scott Morrison. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison --/ -import Lean.Data.SMap - -/-! -# Extra functions on Lean.SMap --/ - -set_option autoImplicit true - -/-- Monadic fold over a staged map. -/ -def Lean.SMap.foldM {m : Type w → Type w} [Monad m] [BEq α] [Hashable α] - (f : σ → α → β → m σ) (init : σ) (map : SMap α β) : m σ := do - map.map₂.foldlM f (← map.map₁.foldM f init) diff --git a/Batteries/Lean/Util/EnvSearch.lean b/Batteries/Lean/Util/EnvSearch.lean index 9143c89511..d9419b8cbf 100644 --- a/Batteries/Lean/Util/EnvSearch.lean +++ b/Batteries/Lean/Util/EnvSearch.lean @@ -24,8 +24,7 @@ where /-- Check constant should be returned -/ @[nolint unusedArguments] check matches_ (_name : Name) cinfo := do - let include ← p cinfo - if include then + if ← p cinfo then pure $ matches_.push cinfo else pure matches_ diff --git a/Batteries/Lean/Util/Path.lean b/Batteries/Lean/Util/Path.lean deleted file mode 100644 index d9488c0dc8..0000000000 --- a/Batteries/Lean/Util/Path.lean +++ /dev/null @@ -1,34 +0,0 @@ -/- -Copyright (c) 2022 Gabriel Ebner. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Gabriel Ebner --/ -import Lean.Elab.Term - -/-! -# `compile_time_search_path%` term elaborator. - -Use this as `searchPathRef.set compile_time_search_path%`. --/ - -open Lean System - --- Ideally this instance would be constructed simply by `deriving instance ToExpr for FilePath` --- but for now we have decided not to upstream the `ToExpr` derive handler from `Mathlib`. --- https://leanprover.zulipchat.com/#narrow/stream/348111-std4/topic/ToExpr.20derive.20handler/near/386476438 -instance : ToExpr FilePath where - toTypeExpr := mkConst ``FilePath - toExpr path := mkApp (mkConst ``FilePath.mk) (toExpr path.1) - -/-- -Term elaborator that retrieves the current `SearchPath`. - -Typical usage is `searchPathRef.set compile_time_search_path%`. - -This must not be used in files that are potentially compiled on another machine and then -imported. -(That is, if used in an imported file it will embed the search path from whichever machine -compiled the `.olean`.) --/ -elab "compile_time_search_path%" : term => - return toExpr (← searchPathRef.get) diff --git a/Batteries/Logic.lean b/Batteries/Logic.lean index d0fd84e688..5c6efe108a 100644 --- a/Batteries/Logic.lean +++ b/Batteries/Logic.lean @@ -16,10 +16,6 @@ instance {f : α → β} [DecidablePred p] : DecidablePred (p ∘ f) := theorem Function.id_def : @id α = fun x => x := rfl -/-! ## exists and forall -/ - -alias ⟨forall_not_of_not_exists, not_exists_of_forall_not⟩ := not_exists - /-! ## decidable -/ protected alias ⟨Decidable.exists_not_of_not_forall, _⟩ := Decidable.not_forall @@ -60,8 +56,7 @@ theorem funext₃ {β : α → Sort _} {γ : ∀ a, β a → Sort _} {δ : ∀ a {f g : ∀ a b c, δ a b c} (h : ∀ a b c, f a b c = g a b c) : f = g := funext fun _ => funext₂ <| h _ -theorem Function.funext_iff {β : α → Sort u} {f₁ f₂ : ∀ x : α, β x} : f₁ = f₂ ↔ ∀ a, f₁ a = f₂ a := - ⟨congrFun, funext⟩ +protected alias Function.funext_iff := funext_iff theorem ne_of_apply_ne {α β : Sort _} (f : α → β) {x y : α} : f x ≠ f y → x ≠ y := mt <| congrArg _ @@ -128,9 +123,7 @@ end Mem -- instance (priority := 10) {α} [Subsingleton α] : DecidableEq α -- | a, b => isTrue (Subsingleton.elim a b) --- @[simp] -- TODO(Mario): profile -theorem eq_iff_true_of_subsingleton [Subsingleton α] (x y : α) : x = y ↔ True := - iff_true_intro (Subsingleton.elim ..) +-- TODO(Mario): profile adding `@[simp]` to `eq_iff_true_of_subsingleton` /-- If all points are equal to a given point `x`, then `α` is a subsingleton. -/ theorem subsingleton_of_forall_eq (x : α) (h : ∀ y, y = x) : Subsingleton α := diff --git a/Batteries/StdDeprecations.lean b/Batteries/StdDeprecations.lean index c705ee839d..f82b1739a7 100644 --- a/Batteries/StdDeprecations.lean +++ b/Batteries/StdDeprecations.lean @@ -22,7 +22,6 @@ Let's hope that people using the tactic implementations can work this out themse -/ @[deprecated (since := "2024-05-07")] alias Std.AssocList := Batteries.AssocList -@[deprecated (since := "2024-05-07")] alias Std.HashMap := Batteries.HashMap @[deprecated (since := "2024-05-07")] alias Std.mkHashMap := Batteries.mkHashMap @[deprecated (since := "2024-05-07")] alias Std.DList := Batteries.DList @[deprecated (since := "2024-05-07")] alias Std.PairingHeapImp.Heap := Batteries.PairingHeapImp.Heap diff --git a/Batteries/Tactic/Lint/Basic.lean b/Batteries/Tactic/Lint/Basic.lean index 87cc925a06..948bab5456 100644 --- a/Batteries/Tactic/Lint/Basic.lean +++ b/Batteries/Tactic/Lint/Basic.lean @@ -42,6 +42,8 @@ def isAutoDecl (decl : Name) : CoreM Bool := do if env.isConstructor n && ["injEq", "inj", "sizeOf_spec"].any (· == s) then return true if let ConstantInfo.inductInfo _ := (← getEnv).find? n then + if s.startsWith "brecOn_" || s.startsWith "below_" || s.startsWith "binductionOn_" + || s.startsWith "ibelow_" then return true if [casesOnSuffix, recOnSuffix, brecOnSuffix, binductionOnSuffix, belowSuffix, "ibelow", "ndrec", "ndrecOn", "noConfusionType", "noConfusion", "ofNat", "toCtorIdx" ].any (· == s) then diff --git a/Batteries/Util/CheckTactic.lean b/Batteries/Util/CheckTactic.lean deleted file mode 100644 index bc97e681cc..0000000000 --- a/Batteries/Util/CheckTactic.lean +++ /dev/null @@ -1,124 +0,0 @@ -/- -Copyright (c) 2024 Lean FRO. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Joe Hendrix --/ -import Lean.Elab.Tactic.ElabTerm -import Lean.Elab.Term - -/- -This file is the home for commands to tactics behave as expected. - -It currently includes two tactics: - -#check_tactic t ~> res - - -only a #check_simp command that applies simp -IT --/ - -namespace Batteries.Tactic - -open Lean Elab Tactic -open Meta - -/-- -Type used to lift an arbitrary value into a type parameter so it can -appear in a proof goal. - -It is used by the #check_tactic command. --/ -private inductive CheckGoalType {α : Sort u} : (val : α) → Prop where -| intro : (val : α) → CheckGoalType val - -private def matchCheckGoalType (stx : Syntax) (goalType : Expr) : MetaM (Expr × Expr × Level) := do - let u ← mkFreshLevelMVar - let type ← mkFreshExprMVar (some (.sort u)) - let val ← mkFreshExprMVar (some type) - let extType := mkAppN (.const ``CheckGoalType [u]) #[type, val] - if !(← isDefEq goalType extType) then - throwErrorAt stx "Goal{indentExpr goalType}\nis expected to match {indentExpr extType}" - pure (val, type, u) - -/-- -`check_tactic_goal t` verifies that the goal has is equal to -`CheckGoalType t` with reducible transparency. It closes the goal if so -and otherwise reports an error. - -It is used by #check_tactic. --/ -local syntax (name := check_tactic_goal) "check_tactic_goal " term : tactic - -/-- -Implementation of `check_tactic_goal` --/ -@[tactic check_tactic_goal] private def evalCheckTacticGoal : Tactic := fun stx => - match stx with - | `(tactic| check_tactic_goal $exp) => - closeMainGoalUsing (checkUnassigned := false) fun goalType => do - let (val, type, u) ← matchCheckGoalType stx goalType - let expTerm ← elabTermEnsuringType exp type - if !(← Meta.withReducible <| isDefEq val expTerm) then - throwErrorAt stx - m!"Term reduces to{indentExpr val}\nbut is expected to reduce to {indentExpr expTerm}" - return mkAppN (.const ``CheckGoalType.intro [u]) #[type, val] - | _ => throwErrorAt stx "check_goal syntax error" - -/-- -`check_tactic_goal t` verifies that the goal has is equal to -`CheckGoalType t` with reducible transparency. It closes the goal if so -and otherwise reports an error. - -It is used by #check_tactic. --/ -local syntax (name := check_tactic_fails) "check_tactic_fails " tactic : tactic - -@[tactic check_tactic_fails] private def evalCheckTacticFails : Tactic := fun stx => do - let `(tactic| check_tactic_fails $tactic) := stx - | throwUnsupportedSyntax - closeMainGoalUsing (checkUnassigned := false) fun goalType => do - let (val, type, u) ← matchCheckGoalType stx goalType - Term.withoutErrToSorry <| withoutRecover do - match (← try (some <$> evalTactic tactic) catch _ => pure none) with - | none => - return mkAppN (.const ``CheckGoalType.intro [u]) #[type, val] - | some () => - let gls ← getGoals - let ppGoal (g : MVarId) := do - let (val, _type, _u) ← matchCheckGoalType stx (← g.getType) - pure m!"{indentExpr val}" - let msg ← - match gls with - | [] => pure m!"{tactic} expected to fail on {val}, but closed goal." - | [g] => - pure <| m!"{tactic} expected to fail on {val}, but returned: {←ppGoal g}" - | gls => - let app m g := do pure <| m ++ (←ppGoal g) - let init := m!"{tactic} expected to fail on {val}, but returned goals:" - gls.foldlM (init := init) app - throwErrorAt stx msg - -/-- -`#check_tactic t ~> r by commands` runs the tactic sequence `commands` -on a goal with t in the type and sees if the resulting expression has -reduced it to `r`. --/ -macro "#check_tactic " t:term "~>" result:term "by" tac:tactic : command => - `(command|example : CheckGoalType $t := by $tac; check_tactic_goal $result) - -/-- -`#check_simp t ~> r` checks `simp` reduces `t` to `r`. --/ -macro "#check_simp " t:term "~>" exp:term : command => - `(command|#check_tactic $t ~> $exp by simp) - -example (x : Nat) : CheckGoalType ((x + z) = x) := by - fail_if_success simp [] - exact (CheckGoalType.intro ((x + z) = x)) - -/-- -`#check_simp t !~>` checks `simp` fails to reduce `t`. --/ -macro "#check_simp " t:term "!~>" : command => - `(command|example : CheckGoalType $t := by check_tactic_fails simp) diff --git a/lean-toolchain b/lean-toolchain index 7f0ea50aaa..64981ae5f5 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.10.0 +leanprover/lean4:v4.11.0-rc1 diff --git a/scripts/runLinter.lean b/scripts/runLinter.lean index 3d98a1c902..4f18be6b50 100644 --- a/scripts/runLinter.lean +++ b/scripts/runLinter.lean @@ -1,6 +1,6 @@ +import Lean.Util.SearchPath import Batteries.Tactic.Lint import Batteries.Data.Array.Basic -import Batteries.Lean.Util.Path open Lean Core Elab Command Std.Tactic.Lint open System (FilePath) diff --git a/test/array.lean b/test/array.lean index 41338a8bd4..2ba3a29577 100644 --- a/test/array.lean +++ b/test/array.lean @@ -1,4 +1,3 @@ -import Batteries.Util.CheckTactic import Batteries.Data.Array section diff --git a/test/congr.lean b/test/congr.lean index a71aed0315..065a899f99 100644 --- a/test/congr.lean +++ b/test/congr.lean @@ -19,8 +19,6 @@ example {α β γ δ} {F : ∀ {α β}, (α → β) → γ → δ} {f g : α → -- apply_assumption -- FIXME apply h -attribute [ext] Subtype.eq - example {α β : Type _} {f : _ → β} {x y : { x : { x : α // x = x } // x = x }} (h : x.1 = y.1) : f x = f y := by congr with : 1 From dc167d260ff7ee9849b436037add06bed15104be Mon Sep 17 00:00:00 2001 From: Austin Letson Date: Sun, 4 Aug 2024 22:57:37 -0400 Subject: [PATCH 040/177] chore: bump lean-action to v1 (#904) --- .github/workflows/build.yml | 3 +-- lakefile.lean | 2 +- 2 files changed, 2 insertions(+), 3 deletions(-) diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 58ca29d3d1..15d72c6876 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -18,10 +18,9 @@ jobs: - id: lean-action name: build, test, and lint batteries - uses: leanprover/lean-action@v1-beta + uses: leanprover/lean-action@v1 with: build-args: '-Kwerror' - lint-module: 'Batteries' - name: Check that all files are imported run: lake env lean scripts/check_imports.lean diff --git a/lakefile.lean b/lakefile.lean index 7ee5efafe6..e55e8b2639 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -13,7 +13,7 @@ package batteries where @[default_target] lean_lib Batteries -@[default_target] +@[default_target, lint_driver] lean_exe runLinter where srcDir := "scripts" supportInterpreter := true From 7b244eafac9dd0d47735d0a9a3a906bfb4910246 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 12 Aug 2024 12:35:52 +1000 Subject: [PATCH 041/177] chore: shake Batteries (#903) --- Batteries/CodeAction/Attr.lean | 2 +- Batteries/CodeAction/Deprecated.lean | 1 - Batteries/Data/ByteArray.lean | 1 - Batteries/Data/HashMap/Lemmas.lean | 1 - Batteries/Data/HashMap/WF.lean | 3 ++- Batteries/Data/List/Pairwise.lean | 1 - Batteries/Data/Rat/Basic.lean | 1 - Batteries/Data/String/Lemmas.lean | 1 - Batteries/Data/UnionFind/Basic.lean | 2 +- Batteries/Lean/Meta/AssertHypotheses.lean | 2 +- Batteries/Lean/Meta/Basic.lean | 4 +--- Batteries/Lean/Meta/Clear.lean | 1 + Batteries/Lean/Meta/Inaccessible.lean | 2 +- Batteries/Lean/Position.lean | 1 - Batteries/Lean/System/IO.lean | 1 - Batteries/Logic.lean | 2 -- Batteries/StdDeprecations.lean | 8 ++++---- Batteries/Tactic/Case.lean | 2 +- Batteries/Tactic/Classical.lean | 2 +- Batteries/Tactic/Init.lean | 6 ++---- Batteries/Tactic/NoMatch.lean | 2 +- Batteries/Tactic/PrintDependents.lean | 1 - Batteries/Tactic/PrintPrefix.lean | 1 - Batteries/Tactic/SeqFocus.lean | 2 +- Batteries/Tactic/ShowUnused.lean | 1 - Batteries/Tactic/Unreachable.lean | 2 +- Batteries/Util/Cache.lean | 2 +- Batteries/Util/ExtendedBinder.lean | 2 -- Batteries/Util/LibraryNote.lean | 2 +- test/MLList.lean | 1 + 30 files changed, 22 insertions(+), 38 deletions(-) diff --git a/Batteries/CodeAction/Attr.lean b/Batteries/CodeAction/Attr.lean index ad5e0f62d3..79802d6932 100644 --- a/Batteries/CodeAction/Attr.lean +++ b/Batteries/CodeAction/Attr.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Lean.Server.CodeActions +import Lean.Server.CodeActions.Basic /-! # Initial setup for code action attributes diff --git a/Batteries/CodeAction/Deprecated.lean b/Batteries/CodeAction/Deprecated.lean index ec6a2f11c2..ce98bcc5c6 100644 --- a/Batteries/CodeAction/Deprecated.lean +++ b/Batteries/CodeAction/Deprecated.lean @@ -5,7 +5,6 @@ Authors: Mario Carneiro -/ import Lean.Server.CodeActions import Batteries.CodeAction.Basic -import Batteries.Lean.Position /-! # Code action for @[deprecated] replacements diff --git a/Batteries/Data/ByteArray.lean b/Batteries/Data/ByteArray.lean index f147eaf47b..585563ec5a 100644 --- a/Batteries/Data/ByteArray.lean +++ b/Batteries/Data/ByteArray.lean @@ -3,7 +3,6 @@ Copyright (c) 2023 François G. Dorais. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: François G. Dorais -/ -import Batteries.Data.Array.Lemmas namespace ByteArray diff --git a/Batteries/Data/HashMap/Lemmas.lean b/Batteries/Data/HashMap/Lemmas.lean index 290f4c0cff..79107fc0bf 100644 --- a/Batteries/Data/HashMap/Lemmas.lean +++ b/Batteries/Data/HashMap/Lemmas.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import Batteries.Data.HashMap.Basic -import Batteries.Data.Array.Lemmas /-! # Lemmas for `Batteries.HashMap` diff --git a/Batteries/Data/HashMap/WF.lean b/Batteries/Data/HashMap/WF.lean index e83e32e567..ae51625133 100644 --- a/Batteries/Data/HashMap/WF.lean +++ b/Batteries/Data/HashMap/WF.lean @@ -3,9 +3,10 @@ Copyright (c) 2022 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Batteries.Tactic.SeqFocus import Batteries.Data.HashMap.Basic -import Batteries.Data.Array.Lemmas import Batteries.Data.Nat.Lemmas +import Batteries.Data.List.Lemmas namespace Batteries.HashMap namespace Imp diff --git a/Batteries/Data/List/Pairwise.lean b/Batteries/Data/List/Pairwise.lean index dd94875ca9..9df947bf4c 100644 --- a/Batteries/Data/List/Pairwise.lean +++ b/Batteries/Data/List/Pairwise.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, James Gallicchio -/ import Batteries.Data.List.Count -import Batteries.Data.Fin.Lemmas /-! # Pairwise relations on a list diff --git a/Batteries/Data/Rat/Basic.lean b/Batteries/Data/Rat/Basic.lean index 600e67981f..f96651d34c 100644 --- a/Batteries/Data/Rat/Basic.lean +++ b/Batteries/Data/Rat/Basic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Batteries.Data.Nat.Gcd -import Batteries.Data.Int.DivMod import Batteries.Lean.Float /-! # Basics for the Rational Numbers -/ diff --git a/Batteries/Data/String/Lemmas.lean b/Batteries/Data/String/Lemmas.lean index 9984d0fb52..df338e03ff 100644 --- a/Batteries/Data/String/Lemmas.lean +++ b/Batteries/Data/String/Lemmas.lean @@ -8,7 +8,6 @@ import Batteries.Data.List.Lemmas import Batteries.Data.String.Basic import Batteries.Tactic.Lint.Misc import Batteries.Tactic.SeqFocus -import Batteries.Tactic.SqueezeScope namespace String diff --git a/Batteries/Data/UnionFind/Basic.lean b/Batteries/Data/UnionFind/Basic.lean index 292f4285eb..e0b4174112 100644 --- a/Batteries/Data/UnionFind/Basic.lean +++ b/Batteries/Data/UnionFind/Basic.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Batteries.Data.Array.Lemmas import Batteries.Tactic.Lint.Misc +import Batteries.Tactic.SeqFocus namespace Batteries diff --git a/Batteries/Lean/Meta/AssertHypotheses.lean b/Batteries/Lean/Meta/AssertHypotheses.lean index a0905c6718..6275032063 100644 --- a/Batteries/Lean/Meta/AssertHypotheses.lean +++ b/Batteries/Lean/Meta/AssertHypotheses.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jannis Limperg -/ -import Batteries.Lean.Meta.Basic +import Lean.Meta.Tactic.Assert open Lean Lean.Meta diff --git a/Batteries/Lean/Meta/Basic.lean b/Batteries/Lean/Meta/Basic.lean index 3e72809893..d2c55af7ea 100644 --- a/Batteries/Lean/Meta/Basic.lean +++ b/Batteries/Lean/Meta/Basic.lean @@ -3,9 +3,7 @@ Copyright (c) 2022 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Jannis Limperg -/ -import Lean.Elab.Term -import Lean.Meta.Tactic.Apply -import Lean.Meta.Tactic.Replace +import Lean.Meta.Tactic.Intro open Lean Lean.Meta diff --git a/Batteries/Lean/Meta/Clear.lean b/Batteries/Lean/Meta/Clear.lean index 47caaaa4d3..66f38b7798 100644 --- a/Batteries/Lean/Meta/Clear.lean +++ b/Batteries/Lean/Meta/Clear.lean @@ -5,6 +5,7 @@ Authors: Jannis Limperg -/ import Batteries.Lean.Meta.Basic +import Lean.Meta.Tactic.Clear open Lean Lean.Meta diff --git a/Batteries/Lean/Meta/Inaccessible.lean b/Batteries/Lean/Meta/Inaccessible.lean index e308364616..e63c4ec4f7 100644 --- a/Batteries/Lean/Meta/Inaccessible.lean +++ b/Batteries/Lean/Meta/Inaccessible.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Jannis Limperg. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jannis Limperg -/ -import Lean.Meta.InferType +import Lean.Meta.Basic open Lean Lean.Meta diff --git a/Batteries/Lean/Position.lean b/Batteries/Lean/Position.lean index f2bc155886..49fdad203e 100644 --- a/Batteries/Lean/Position.lean +++ b/Batteries/Lean/Position.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Lean.Syntax -import Lean.Meta.Tactic.TryThis import Lean.Data.Lsp.Utf16 /-- Gets the LSP range of syntax `stx`. -/ diff --git a/Batteries/Lean/System/IO.lean b/Batteries/Lean/System/IO.lean index c4d6b6ce84..40dcda5365 100644 --- a/Batteries/Lean/System/IO.lean +++ b/Batteries/Lean/System/IO.lean @@ -3,7 +3,6 @@ Copyright (c) 2023 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import Batteries.Data.List.Lemmas /-! # Functions for manipulating a list of tasks diff --git a/Batteries/Logic.lean b/Batteries/Logic.lean index 5c6efe108a..03c82eb5ac 100644 --- a/Batteries/Logic.lean +++ b/Batteries/Logic.lean @@ -3,9 +3,7 @@ Copyright (c) 2014 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Floris van Doorn, Mario Carneiro -/ -import Batteries.Tactic.Init import Batteries.Tactic.Alias -import Batteries.Tactic.Lint.Misc instance {f : α → β} [DecidablePred p] : DecidablePred (p ∘ f) := inferInstanceAs <| DecidablePred fun x => p (f x) diff --git a/Batteries/StdDeprecations.lean b/Batteries/StdDeprecations.lean index f82b1739a7..cb49a3bdbf 100644 --- a/Batteries/StdDeprecations.lean +++ b/Batteries/StdDeprecations.lean @@ -4,12 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Batteries.Tactic.Alias -import Batteries.Data.HashMap import Batteries.Data.DList import Batteries.Data.PairingHeap -import Batteries.Data.RBMap -import Batteries.Data.BinomialHeap -import Batteries.Data.UnionFind +import Batteries.Data.BinomialHeap.Basic +import Batteries.Data.HashMap.Basic +import Batteries.Data.RBMap.Basic +import Batteries.Data.UnionFind.Basic /-! # We set up deprecations for identifiers formerly in the `Std` namespace. diff --git a/Batteries/Tactic/Case.lean b/Batteries/Tactic/Case.lean index 60bd585244..375d351b05 100644 --- a/Batteries/Tactic/Case.lean +++ b/Batteries/Tactic/Case.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ -import Lean.Elab.Tactic.Conv.Pattern +import Lean.Elab.Tactic.BuiltinTactic /-! # Extensions to the `case` tactic diff --git a/Batteries/Tactic/Classical.lean b/Batteries/Tactic/Classical.lean index 7d3e6eb3c0..a3bc78b0a4 100644 --- a/Batteries/Tactic/Classical.lean +++ b/Batteries/Tactic/Classical.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Lean.Elab.ElabRules +import Lean.Elab.Tactic.Basic /-! # `classical` tactic -/ diff --git a/Batteries/Tactic/Init.lean b/Batteries/Tactic/Init.lean index fed3c1ad99..db7894f633 100644 --- a/Batteries/Tactic/Init.lean +++ b/Batteries/Tactic/Init.lean @@ -3,16 +3,14 @@ Copyright (c) 2021 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Batteries.Lean.Meta.Basic -import Lean.Elab.Command -import Lean.Elab.Tactic.BuiltinTactic +import Lean.Elab.Tactic.ElabTerm /-! # Simple tactics that are used throughout Batteries. -/ namespace Batteries.Tactic -open Lean Parser.Tactic Elab Command Elab.Tactic Meta +open Lean Parser.Tactic Elab Elab.Tactic Meta /-- `_` in tactic position acts like the `done` tactic: it fails and gives the list diff --git a/Batteries/Tactic/NoMatch.lean b/Batteries/Tactic/NoMatch.lean index 7efe68e7d2..ee6acefaf0 100644 --- a/Batteries/Tactic/NoMatch.lean +++ b/Batteries/Tactic/NoMatch.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Lean.Elab.ElabRules import Lean.DocString +import Lean.Elab.Tactic.Basic /-! Deprecation warnings for `match ⋯ with.`, `fun.`, `λ.`, and `intro.`. diff --git a/Batteries/Tactic/PrintDependents.lean b/Batteries/Tactic/PrintDependents.lean index 9c55a04ecd..e927713886 100644 --- a/Batteries/Tactic/PrintDependents.lean +++ b/Batteries/Tactic/PrintDependents.lean @@ -5,7 +5,6 @@ Authors: Mario Carneiro -/ import Lean.Elab.Command import Lean.Util.FoldConsts -import Batteries.Lean.Delaborator /-! # `#print dependents` command diff --git a/Batteries/Tactic/PrintPrefix.lean b/Batteries/Tactic/PrintPrefix.lean index 63d2c3bbf5..b83e7b0ad8 100644 --- a/Batteries/Tactic/PrintPrefix.lean +++ b/Batteries/Tactic/PrintPrefix.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Shing Tak Lam, Daniel Selsam, Mario Carneiro -/ import Batteries.Lean.Util.EnvSearch -import Batteries.Lean.Delaborator import Lean.Elab.Tactic.Config namespace Batteries.Tactic diff --git a/Batteries/Tactic/SeqFocus.lean b/Batteries/Tactic/SeqFocus.lean index 8972f77891..cbbe70a44a 100644 --- a/Batteries/Tactic/SeqFocus.lean +++ b/Batteries/Tactic/SeqFocus.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad -/ -import Lean.Elab.ElabRules +import Lean.Elab.Tactic.Basic open Lean Elab Meta Tactic diff --git a/Batteries/Tactic/ShowUnused.lean b/Batteries/Tactic/ShowUnused.lean index 25989a3da9..3797801cd8 100644 --- a/Batteries/Tactic/ShowUnused.lean +++ b/Batteries/Tactic/ShowUnused.lean @@ -5,7 +5,6 @@ Authors: Mario Carneiro -/ import Lean.Util.FoldConsts import Lean.Linter.UnusedVariables -import Batteries.Lean.Delaborator /-! # The `#show_unused` command diff --git a/Batteries/Tactic/Unreachable.lean b/Batteries/Tactic/Unreachable.lean index 76571dc850..4d6e6cf0a1 100644 --- a/Batteries/Tactic/Unreachable.lean +++ b/Batteries/Tactic/Unreachable.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Lean.Elab.ElabRules +import Lean.Elab.Tactic.Basic namespace Batteries.Tactic diff --git a/Batteries/Util/Cache.lean b/Batteries/Util/Cache.lean index edfac4fe97..866325df5a 100644 --- a/Batteries/Util/Cache.lean +++ b/Batteries/Util/Cache.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Gabriel Ebner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner -/ -import Batteries.Lean.Meta.DiscrTree +import Lean.Meta.DiscrTree /-! # Once-per-file cache for tactics diff --git a/Batteries/Util/ExtendedBinder.lean b/Batteries/Util/ExtendedBinder.lean index d577e1d785..c489c0a432 100644 --- a/Batteries/Util/ExtendedBinder.lean +++ b/Batteries/Util/ExtendedBinder.lean @@ -3,8 +3,6 @@ Copyright (c) 2021 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner -/ -import Lean.Elab.MacroArgUtil -import Lean.Linter.MissingDocs /-! Defines an extended binder syntax supporting `∀ ε > 0, ...` etc. diff --git a/Batteries/Util/LibraryNote.lean b/Batteries/Util/LibraryNote.lean index eb08bf0b27..8c1d45f786 100644 --- a/Batteries/Util/LibraryNote.lean +++ b/Batteries/Util/LibraryNote.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Gabriel Ebner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner -/ -import Lean.Elab.ElabRules +import Lean.Elab.Command /-! # Define the `library_note` command. diff --git a/test/MLList.lean b/test/MLList.lean index 278d6dc83d..78f0299717 100644 --- a/test/MLList.lean +++ b/test/MLList.lean @@ -1,4 +1,5 @@ import Batteries.Data.MLList.IO +import Batteries.Data.List.Basic set_option linter.missingDocs false From d4dfc4ed5d432668b39f03eb69b7415d17156f05 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 12 Aug 2024 12:36:18 +1000 Subject: [PATCH 042/177] chore: squeeze a simp in HashMap/WF (#909) --- Batteries/Data/HashMap/WF.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Batteries/Data/HashMap/WF.lean b/Batteries/Data/HashMap/WF.lean index ae51625133..9f69994638 100644 --- a/Batteries/Data/HashMap/WF.lean +++ b/Batteries/Data/HashMap/WF.lean @@ -225,7 +225,7 @@ private theorem pairwise_replaceF [BEq α] [PartialEquivBEq α] | cons a l ih => simp only [List.pairwise_cons, List.replaceF] at H ⊢ generalize e : cond .. = z; unfold cond at e; revert e - split <;> (intro h; subst h; simp) + split <;> (intro h; subst h; simp only [List.pairwise_cons]) · next e => exact ⟨(H.1 · · ∘ PartialEquivBEq.trans e), H.2⟩ · next e => refine ⟨fun a h => ?_, ih H.2⟩ From ad26fe1ebccc9d5b7ca9111d5daf9b4488374415 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 12 Aug 2024 17:19:11 +1000 Subject: [PATCH 043/177] chore: bump toolchain to v4.11.0-rc2 (#911) --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 64981ae5f5..e7a4f40b89 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.11.0-rc1 +leanprover/lean4:v4.11.0-rc2 From 1d25ec7ec98d6d9fb526c997aa014bcabbad8b72 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Tue, 13 Aug 2024 15:34:09 -0400 Subject: [PATCH 044/177] fix: naming convention (#914) Many theorems have the reverse naming convention for `Array.data` and `ByteArray.data`. This fixes many of them but there are plenty more in Lean core. An unexpected upshot is that some corrected names are actually shorter! --- Batteries/Data/Array/Lemmas.lean | 20 +++++++++++-------- Batteries/Data/ByteArray.lean | 33 ++++++++++++++++++++------------ 2 files changed, 33 insertions(+), 20 deletions(-) diff --git a/Batteries/Data/Array/Lemmas.lean b/Batteries/Data/Array/Lemmas.lean index 753cbb4955..dd4fe12424 100644 --- a/Batteries/Data/Array/Lemmas.lean +++ b/Batteries/Data/Array/Lemmas.lean @@ -11,7 +11,7 @@ import Batteries.Util.ProofWanted namespace Array -theorem forIn_eq_data_forIn [Monad m] +theorem forIn_eq_forIn_data [Monad m] (as : Array α) (b : β) (f : α → β → m (ForInStep β)) : forIn as b f = forIn as.data b f := by let rec loop : ∀ {i h b j}, j + i = as.size → @@ -27,10 +27,11 @@ theorem forIn_eq_data_forIn [Monad m] rw [loop (i := i)]; rw [← ij, Nat.succ_add]; rfl conv => lhs; simp only [forIn, Array.forIn] rw [loop (Nat.zero_add _)]; rfl +@[deprecated (since := "2024-08-13")] alias forIn_eq_data_forIn := forIn_eq_forIn_data /-! ### zipWith / zip -/ -theorem zipWith_eq_zipWith_data (f : α → β → γ) (as : Array α) (bs : Array β) : +theorem data_zipWith (f : α → β → γ) (as : Array α) (bs : Array β) : (as.zipWith bs f).data = as.data.zipWith f bs.data := by let rec loop : ∀ (i : Nat) cs, i ≤ as.size → i ≤ bs.size → (zipWithAux f as bs i cs).data = cs.data ++ (as.data.drop i).zipWith f (bs.data.drop i) := by @@ -70,14 +71,16 @@ theorem zipWith_eq_zipWith_data (f : α → β → γ) (as : Array α) (bs : Arr List.zipWith f (List.drop i as.data) (List.drop i bs.data) simp only [data_length, Fin.getElem_fin, List.getElem_cons_drop, List.get_eq_getElem] simp [zipWith, loop 0 #[] (by simp) (by simp)] +@[deprecated (since := "2024-08-13")] alias zipWith_eq_zipWith_data := data_zipWith theorem size_zipWith (as : Array α) (bs : Array β) (f : α → β → γ) : (as.zipWith bs f).size = min as.size bs.size := by - rw [size_eq_length_data, zipWith_eq_zipWith_data, List.length_zipWith] + rw [size_eq_length_data, data_zipWith, List.length_zipWith] -theorem zip_eq_zip_data (as : Array α) (bs : Array β) : +theorem data_zip (as : Array α) (bs : Array β) : (as.zip bs).data = as.data.zip bs.data := - zipWith_eq_zipWith_data Prod.mk as bs + data_zipWith Prod.mk as bs +@[deprecated (since := "2024-08-13")] alias zip_eq_zip_data := data_zip theorem size_zip (as : Array α) (bs : Array β) : (as.zip bs).size = min as.size bs.size := @@ -92,7 +95,7 @@ theorem size_filter_le (p : α → Bool) (l : Array α) : /-! ### join -/ -@[simp] theorem join_data {l : Array (Array α)} : l.join.data = (l.data.map data).join := by +@[simp] theorem data_join {l : Array (Array α)} : l.join.data = (l.data.map data).join := by dsimp [join] simp only [foldl_eq_foldl_data] generalize l.data = l @@ -101,9 +104,10 @@ theorem size_filter_le (p : α → Bool) (l : Array α) : induction l with | nil => simp | cons h => induction h.data <;> simp [*] +@[deprecated (since := "2024-08-13")] alias join_data := data_join theorem mem_join : ∀ {L : Array (Array α)}, a ∈ L.join ↔ ∃ l, l ∈ L ∧ a ∈ l := by - simp only [mem_def, join_data, List.mem_join, List.mem_map] + simp only [mem_def, data_join, List.mem_join, List.mem_map] intro l constructor · rintro ⟨_, ⟨s, m, rfl⟩, h⟩ @@ -113,7 +117,7 @@ theorem mem_join : ∀ {L : Array (Array α)}, a ∈ L.join ↔ ∃ l, l ∈ L /-! ### erase -/ -@[simp] proof_wanted erase_data [BEq α] {l : Array α} {a : α} : (l.erase a).data = l.data.erase a +@[simp] proof_wanted data_erase [BEq α] {l : Array α} {a : α} : (l.erase a).data = l.data.erase a /-! ### shrink -/ diff --git a/Batteries/Data/ByteArray.lean b/Batteries/Data/ByteArray.lean index 585563ec5a..a9653bf10c 100644 --- a/Batteries/Data/ByteArray.lean +++ b/Batteries/Data/ByteArray.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 François G. Dorais. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: François G. Dorais -/ +import Batteries.Tactic.Alias namespace ByteArray @@ -18,15 +19,18 @@ theorem getElem_eq_data_getElem (a : ByteArray) (h : i < a.size) : a[i] = a.data /-! ### empty -/ -@[simp] theorem mkEmpty_data (cap) : (mkEmpty cap).data = #[] := rfl +@[simp] theorem data_mkEmpty (cap) : (mkEmpty cap).data = #[] := rfl +@[deprecated (since := "2024-08-13")] alias mkEmpty_data := data_mkEmpty -@[simp] theorem empty_data : empty.data = #[] := rfl +@[simp] theorem data_empty : empty.data = #[] := rfl +@[deprecated (since := "2024-08-13")] alias empty_data := data_empty @[simp] theorem size_empty : empty.size = 0 := rfl /-! ### push -/ -@[simp] theorem push_data (a : ByteArray) (b : UInt8) : (a.push b).data = a.data.push b := rfl +@[simp] theorem data_push (a : ByteArray) (b : UInt8) : (a.push b).data = a.data.push b := rfl +@[deprecated (since := "2024-08-13")] alias push_data := data_push @[simp] theorem size_push (a : ByteArray) (b : UInt8) : (a.push b).size = a.size + 1 := Array.size_push .. @@ -40,8 +44,9 @@ theorem get_push_lt (a : ByteArray) (x : UInt8) (i : Nat) (h : i < a.size) : /-! ### set -/ -@[simp] theorem set_data (a : ByteArray) (i : Fin a.size) (v : UInt8) : +@[simp] theorem data_set (a : ByteArray) (i : Fin a.size) (v : UInt8) : (a.set i v).data = a.data.set i v := rfl +@[deprecated (since := "2024-08-13")] alias set_data := data_set @[simp] theorem size_set (a : ByteArray) (i : Fin a.size) (v : UInt8) : (a.set i v).size = a.size := @@ -60,20 +65,22 @@ theorem set_set (a : ByteArray) (i : Fin a.size) (v v' : UInt8) : /-! ### copySlice -/ -@[simp] theorem copySlice_data (a i b j len exact) : +@[simp] theorem data_copySlice (a i b j len exact) : (copySlice a i b j len exact).data = b.data.extract 0 j ++ a.data.extract i (i + len) ++ b.data.extract (j + min len (a.data.size - i)) b.data.size := rfl +@[deprecated (since := "2024-08-13")] alias copySlice_data := data_copySlice /-! ### append -/ @[simp] theorem append_eq (a b) : ByteArray.append a b = a ++ b := rfl -@[simp] theorem append_data (a b : ByteArray) : (a ++ b).data = a.data ++ b.data := by +@[simp] theorem data_append (a b : ByteArray) : (a ++ b).data = a.data ++ b.data := by rw [←append_eq]; simp [ByteArray.append, size] rw [Array.extract_empty_of_stop_le_start (h:=Nat.le_add_right ..), Array.append_nil] +@[deprecated (since := "2024-08-13")] alias append_data := data_append theorem size_append (a b : ByteArray) : (a ++ b).size = a.size + b.size := by - simp only [size, append_eq, append_data]; exact Array.size_append .. + simp only [size, append_eq, data_append]; exact Array.size_append .. theorem get_append_left {a b : ByteArray} (hlt : i < a.size) (h : i < (a ++ b).size := size_append .. ▸ Nat.lt_of_lt_of_le hlt (Nat.le_add_right ..)) : @@ -87,12 +94,13 @@ theorem get_append_right {a b : ByteArray} (hle : a.size ≤ i) (h : i < (a ++ b /-! ### extract -/ -@[simp] theorem extract_data (a : ByteArray) (start stop) : +@[simp] theorem data_extract (a : ByteArray) (start stop) : (a.extract start stop).data = a.data.extract start stop := by simp [extract] match Nat.le_total start stop with | .inl h => simp [h, Nat.add_sub_cancel'] | .inr h => simp [h, Nat.sub_eq_zero_of_le, Array.extract_empty_of_stop_le_start] +@[deprecated (since := "2024-08-13")] alias extract_data := data_extract @[simp] theorem size_extract (a : ByteArray) (start stop) : (a.extract start stop).size = min stop a.size - start := by @@ -113,7 +121,8 @@ theorem get_extract_aux {a : ByteArray} {start stop} (h : i < (a.extract start s def ofFn (f : Fin n → UInt8) : ByteArray where data := .ofFn f -@[simp] theorem ofFn_data (f : Fin n → UInt8) : (ofFn f).data = .ofFn f := rfl +@[simp] theorem data_ofFn (f : Fin n → UInt8) : (ofFn f).data = .ofFn f := rfl +@[deprecated (since := "2024-08-13")] alias ofFn_data := data_ofFn @[simp] theorem size_ofFn (f : Fin n → UInt8) : (ofFn f).size = n := by simp [size] @@ -131,9 +140,9 @@ private def ofFnAux (f : Fin n → UInt8) : ByteArray := go 0 (mkEmpty n) where termination_by n - i @[csimp] private theorem ofFn_eq_ofFnAux : @ofFn = @ofFnAux := by - funext n f; ext; simp [ofFnAux, Array.ofFn, ofFnAux_data, mkEmpty] + funext n f; ext; simp [ofFnAux, Array.ofFn, data_ofFnAux, mkEmpty] where - ofFnAux_data {n} (f : Fin n → UInt8) (i) {acc} : + data_ofFnAux {n} (f : Fin n → UInt8) (i) {acc} : (ofFnAux.go f i acc).data = Array.ofFn.go f i acc.data := by - rw [ofFnAux.go, Array.ofFn.go]; split; rw [ofFnAux_data f (i+1), push_data]; rfl + rw [ofFnAux.go, Array.ofFn.go]; split; rw [data_ofFnAux f (i+1), data_push]; rfl termination_by n - i From 5e5e54c4028f7b6bd086ebb72e5822794c64c35d Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 14 Aug 2024 18:08:56 +1000 Subject: [PATCH 045/177] chore: cleanup lexOrd_def (#915) --- Batteries/Classes/Order.lean | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/Batteries/Classes/Order.lean b/Batteries/Classes/Order.lean index 8cbfae25b0..1296bd7a9d 100644 --- a/Batteries/Classes/Order.lean +++ b/Batteries/Classes/Order.lean @@ -278,11 +278,8 @@ instance [Ord β] [OrientedOrd β] (f : α → β) : OrientedCmp (compareOn f) w instance [Ord β] [TransOrd β] (f : α → β) : TransCmp (compareOn f) where le_trans := TransCmp.le_trans (α := β) --- FIXME: remove after lean4#3882 is merged theorem _root_.lexOrd_def [Ord α] [Ord β] : - (lexOrd : Ord (α × β)).compare = compareLex (compareOn (·.1)) (compareOn (·.2)) := by - funext a b - simp [lexOrd, compareLex, compareOn] + (lexOrd : Ord (α × β)).compare = compareLex (compareOn (·.1)) (compareOn (·.2)) := rfl section «non-canonical instances» -- Note: the following instances seem to cause lean to fail, see: From 91cf60cee20d61bd2012cae20ea6589432f5bf5a Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Fri, 16 Aug 2024 09:35:05 +0200 Subject: [PATCH 046/177] fix: flaky test (#889) --- test/MLList.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/test/MLList.lean b/test/MLList.lean index 78f0299717..ad53377237 100644 --- a/test/MLList.lean +++ b/test/MLList.lean @@ -19,7 +19,7 @@ def g (n : Nat) : MLList Lean.Meta.MetaM Nat := do /-! ### Test `MLList.ofTaskList`. -We generate three tasks which sleep for `100`, `10`, and `1` milliseconds respectively, +We generate three tasks which sleep for `100`, `50`, and `1` milliseconds respectively, and then verify that `MLList.ofTaskList` return their results in the order they complete. -/ @@ -29,9 +29,9 @@ def sleep (n : UInt32) : BaseIO (Task UInt32) := | .error _ => 0 def sleeps : MLList BaseIO UInt32 := .squash fun _ => do - let r ← [100,10,1].map sleep |>.traverse id + let r ← [100,50,1].map sleep |>.traverse id return .ofTaskList r -/-- info: [1, 10, 100] -/ +/-- info: [1, 50, 100] -/ #guard_msgs in #eval sleeps.force From b0587b2dc108d23519735a4595de4db92e2cf2ca Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 16 Aug 2024 17:35:32 +1000 Subject: [PATCH 047/177] chore: remove makefile (#918) --- GNUmakefile | 12 ------------ README.md | 4 +++- 2 files changed, 3 insertions(+), 13 deletions(-) delete mode 100644 GNUmakefile diff --git a/GNUmakefile b/GNUmakefile deleted file mode 100644 index 123c8af1a1..0000000000 --- a/GNUmakefile +++ /dev/null @@ -1,12 +0,0 @@ -.PHONY: all build test lint - -all: build test - -build: - lake build - -test: - lake test - -lint: - lake exe runLinter diff --git a/README.md b/README.md index f5f241fe22..ab8606d854 100644 --- a/README.md +++ b/README.md @@ -23,7 +23,9 @@ Additionally, please make sure that you're using the version of Lean that the cu curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh ``` If this also fails, follow the instructions under `Regular install` [here](https://leanprover-community.github.io/get_started.html). -* To build `batteries` run `lake build`. To build and run all tests, run `make`. +* To build `batteries` run `lake build`. +* To build and run all tests, run `lake test`. +* To run the environment linter, run `lake lint`. * If you added a new file, run the command `scripts/updateBatteries.sh` to update the imports. From a36e34fb3d00117ccde6ad6c3d5774b1e4f1f36e Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 16 Aug 2024 17:38:33 +1000 Subject: [PATCH 048/177] feat: faster splitAt (#919) --- Batteries/Data/List/Basic.lean | 22 +++++++++++----------- 1 file changed, 11 insertions(+), 11 deletions(-) diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index 43915cd733..7e6e034dc4 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -116,13 +116,13 @@ Split a list at an index. splitAt 2 [a, b, c] = ([a, b], [c]) ``` -/ -def splitAt (n : Nat) (l : List α) : List α × List α := go l n #[] where - /-- Auxiliary for `splitAt`: `splitAt.go l n xs acc = (acc.toList ++ take n xs, drop n xs)` +def splitAt (n : Nat) (l : List α) : List α × List α := go l n [] where + /-- Auxiliary for `splitAt`: `splitAt.go l n xs acc = (acc.reverse ++ take n xs, drop n xs)` if `n < length xs`, else `(l, [])`. -/ - go : List α → Nat → Array α → List α × List α + go : List α → Nat → List α → List α × List α | [], _, _ => (l, []) - | x :: xs, n+1, acc => go xs n (acc.push x) - | xs, _, acc => (acc.toList, xs) + | x :: xs, n+1, acc => go xs n (x :: acc) + | xs, _, acc => (acc.reverse, xs) /-- Split a list at an index. Ensures the left list always has the specified length @@ -132,13 +132,13 @@ splitAtD 2 [a, b, c] x = ([a, b], [c]) splitAtD 4 [a, b, c] x = ([a, b, c, x], []) ``` -/ -def splitAtD (n : Nat) (l : List α) (dflt : α) : List α × List α := go n l #[] where - /-- Auxiliary for `splitAtD`: `splitAtD.go dflt n l acc = (acc.toList ++ left, right)` +def splitAtD (n : Nat) (l : List α) (dflt : α) : List α × List α := go n l [] where + /-- Auxiliary for `splitAtD`: `splitAtD.go dflt n l acc = (acc.reverse ++ left, right)` if `splitAtD n l dflt = (left, right)`. -/ - go : Nat → List α → Array α → List α × List α - | n+1, x :: xs, acc => go n xs (acc.push x) - | 0, xs, acc => (acc.toList, xs) - | n, [], acc => (acc.toListAppend (replicate n dflt), []) + go : Nat → List α → List α → List α × List α + | n+1, x :: xs, acc => go n xs (x :: acc) + | 0, xs, acc => (acc, xs) + | n, [], acc => (acc.reverseAux (replicate n dflt), []) /-- Split a list at every element satisfying a predicate. The separators are not in the result. From 1f4c8fab70847b92df7efa5a35a0e7613217bd2b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Fri, 16 Aug 2024 17:16:08 -0400 Subject: [PATCH 049/177] feat: add `ByteSubarray` (#851) --- Batteries.lean | 1 + Batteries/Data/ByteSubarray.lean | 105 +++++++++++++++++++++++++++++++ 2 files changed, 106 insertions(+) create mode 100644 Batteries/Data/ByteSubarray.lean diff --git a/Batteries.lean b/Batteries.lean index 8e1e6f1828..a585b2c6cf 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -19,6 +19,7 @@ import Batteries.Data.BinomialHeap import Batteries.Data.BitVec import Batteries.Data.Bool import Batteries.Data.ByteArray +import Batteries.Data.ByteSubarray import Batteries.Data.Char import Batteries.Data.DList import Batteries.Data.Fin diff --git a/Batteries/Data/ByteSubarray.lean b/Batteries/Data/ByteSubarray.lean new file mode 100644 index 0000000000..3844d1b1bc --- /dev/null +++ b/Batteries/Data/ByteSubarray.lean @@ -0,0 +1,105 @@ + +/- +Copyright (c) 2021 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, François G. Dorais +-/ + +namespace Batteries + +/-- A subarray of a `ByteArray`. -/ +structure ByteSubarray where + /-- `O(1)`. Get data array of a `ByteSubarray`. -/ + array : ByteArray + /-- `O(1)`. Get start index of a `ByteSubarray`. -/ + start : Nat + /-- `O(1)`. Get stop index of a `ByteSubarray`. -/ + stop : Nat + /-- Start index is before stop index. -/ + start_le_stop : start ≤ stop + /-- Stop index is before end of data array. -/ + stop_le_array_size : stop ≤ array.size + +namespace ByteSubarray + +/-- `O(1)`. Get the size of a `ByteSubarray`. -/ +protected def size (self : ByteSubarray) := self.stop - self.start + +/-- `O(1)`. Test if a `ByteSubarray` is empty. -/ +protected def isEmpty (self : ByteSubarray) := self.start != self.stop + +theorem stop_eq_start_add_size (self : ByteSubarray) : self.stop = self.start + self.size := by + rw [ByteSubarray.size, Nat.add_sub_cancel' self.start_le_stop] + +/-- `O(n)`. Extract a `ByteSubarray` to a `ByteArray`. -/ +def toByteArray (self : ByteSubarray) : ByteArray := + self.array.extract self.start self.stop + +/-- `O(1)`. Get the element at index `i` from the start of a `ByteSubarray`. -/ +@[inline] def get (self : ByteSubarray) (i : Fin self.size) : UInt8 := + have : self.start + i.1 < self.array.size := by + apply Nat.lt_of_lt_of_le _ self.stop_le_array_size + rw [stop_eq_start_add_size] + apply Nat.add_lt_add_left i.is_lt self.start + self.array[self.start + i.1] + +instance : GetElem ByteSubarray Nat UInt8 fun self i => i < self.size where + getElem self i h := self.get ⟨i, h⟩ + +/-- `O(1)`. Pop the last element of a `ByteSubarray`. -/ +@[inline] def pop (self : ByteSubarray) : ByteSubarray := + if h : self.start = self.stop then self else + {self with + stop := self.stop - 1 + start_le_stop := Nat.le_pred_of_lt (Nat.lt_of_le_of_ne self.start_le_stop h) + stop_le_array_size := Nat.le_trans (Nat.pred_le _) self.stop_le_array_size + } + +/-- `O(1)`. Pop the first element of a `ByteSubarray`. -/ +@[inline] def popFront (self : ByteSubarray) : ByteSubarray := + if h : self.start = self.stop then self else + {self with + start := self.start + 1 + start_le_stop := Nat.succ_le_of_lt (Nat.lt_of_le_of_ne self.start_le_stop h) + } + +/-- Folds a monadic function over a `ByteSubarray` from left to right. -/ +@[inline] def foldlM [Monad m] (self : ByteSubarray) (f : β → UInt8 → m β) (init : β) : m β := + self.array.foldlM f init self.start self.stop + +/-- Folds a function over a `ByteSubarray` from left to right. -/ +@[inline] def foldl (self : ByteSubarray) (f : β → UInt8 → β) (init : β) : β := + self.foldlM (m:=Id) f init + +/-- Implementation of `forIn` for a `ByteSubarray`. -/ +@[specialize] +protected def forIn [Monad m] (self : ByteSubarray) (init : β) (f : UInt8 → β → m (ForInStep β)) : + m β := loop self.size (Nat.le_refl _) init +where + /-- Inner loop of the `forIn` implementation for `ByteSubarray`. -/ + loop (i : Nat) (h : i ≤ self.size) (b : β) : m β := do + match i, h with + | 0, _ => pure b + | i+1, h => + match (← f self[self.size - 1 - i] b) with + | ForInStep.done b => pure b + | ForInStep.yield b => loop i (Nat.le_of_succ_le h) b + +instance : ForIn m ByteSubarray UInt8 where + forIn := ByteSubarray.forIn + +instance : Stream ByteSubarray UInt8 where + next? s := s[0]? >>= fun x => (x, s.popFront) + +end Batteries.ByteSubarray + +/-- `O(1)`. Coerce a byte array into a byte slice. -/ +def ByteArray.toByteSubarray (array : ByteArray) : Batteries.ByteSubarray where + array := array + start := 0 + stop := array.size + start_le_stop := Nat.zero_le _ + stop_le_array_size := Nat.le_refl _ + +instance : Coe ByteArray Batteries.ByteSubarray where + coe := ByteArray.toByteSubarray From a2ef715a780564dfea52464d573f3ce5c7d6cf95 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Fri, 16 Aug 2024 17:16:29 -0400 Subject: [PATCH 050/177] feat: lemmas `List.foldlM_map` and `List.foldrM_map` (#827) --- Batteries/Data/List/Lemmas.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 2e97cb5264..3eea690ad6 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -637,3 +637,13 @@ theorem mem_merge_left (s : α → α → Bool) (h : x ∈ l) : x ∈ merge s l theorem mem_merge_right (s : α → α → Bool) (h : x ∈ r) : x ∈ merge s l r := mem_merge.2 <| .inr h + +/-! ### foldlM and foldrM -/ + +theorem foldlM_map [Monad m] (f : β₁ → β₂) (g : α → β₂ → m α) (l : List β₁) (init : α) : + (l.map f).foldlM g init = l.foldlM (fun x y => g x (f y)) init := by + induction l generalizing g init <;> simp [*] + +theorem foldrM_map [Monad m] [LawfulMonad m] (f : β₁ → β₂) (g : β₂ → α → m α) (l : List β₁) + (init : α) : (l.map f).foldrM g init = l.foldrM (fun x y => g (f x) y) init := by + induction l generalizing g init <;> simp [*] From 5f405b1c0d74619b34501ff0c711cb882406800e Mon Sep 17 00:00:00 2001 From: Shrys Date: Sat, 17 Aug 2024 00:45:53 +0200 Subject: [PATCH 051/177] feat: statically sized Vector type (#793) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds a static vector type to batteries akin to `Mathlib.Data.Vector` with a key difference: It is based on `Array` instead of `List` and is designed for efficient operations with while offering a static size. Co-authored-by: François G. Dorais Co-authored-by: Mario Carneiro --- Batteries.lean | 1 + Batteries/Data/Vector.lean | 2 + Batteries/Data/Vector/Basic.lean | 321 ++++++++++++++++++++++++++++++ Batteries/Data/Vector/Lemmas.lean | 47 +++++ 4 files changed, 371 insertions(+) create mode 100644 Batteries/Data/Vector.lean create mode 100644 Batteries/Data/Vector/Basic.lean create mode 100644 Batteries/Data/Vector/Lemmas.lean diff --git a/Batteries.lean b/Batteries.lean index a585b2c6cf..5fd8bc72ad 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -39,6 +39,7 @@ import Batteries.Data.Sum import Batteries.Data.Thunk import Batteries.Data.UInt import Batteries.Data.UnionFind +import Batteries.Data.Vector import Batteries.Lean.AttributeExtra import Batteries.Lean.Delaborator import Batteries.Lean.Except diff --git a/Batteries/Data/Vector.lean b/Batteries/Data/Vector.lean new file mode 100644 index 0000000000..f543f121ef --- /dev/null +++ b/Batteries/Data/Vector.lean @@ -0,0 +1,2 @@ +import Batteries.Data.Vector.Basic +import Batteries.Data.Vector.Lemmas diff --git a/Batteries/Data/Vector/Basic.lean b/Batteries/Data/Vector/Basic.lean new file mode 100644 index 0000000000..6fc8166b6d --- /dev/null +++ b/Batteries/Data/Vector/Basic.lean @@ -0,0 +1,321 @@ +/- +Copyright (c) 2024 Shreyas Srinivas. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Shreyas Srinivas, François G. Dorais +-/ + +import Batteries.Data.Array +import Batteries.Data.List.Basic +import Batteries.Data.List.Lemmas +import Batteries.Tactic.Lint.Misc + +/-! +## Vectors +`Vector α n` is an array with a statically fixed size `n`. +It combines the benefits of Lean's special support for `Arrays` +that offer `O(1)` accesses and in-place mutations for arrays with no more than one reference, +with static guarantees about the size of the underlying array. +-/ + +namespace Batteries + +/-- `Vector α n` is an `Array α` whose size is statically fixed to `n` -/ +structure Vector (α : Type u) (n : Nat) where + /-- Internally, a vector is stored as an array for fast access -/ + toArray : Array α + /-- `size_eq` fixes the size of `toArray` statically -/ + size_eq : toArray.size = n +deriving Repr, DecidableEq + +namespace Vector + +/-- Syntax for `Vector α n` -/ +syntax "#v[" withoutPosition(sepBy(term, ", ")) "]" : term + +open Lean in +macro_rules + | `(#v[ $elems,* ]) => `(Vector.mk (n := $(quote elems.getElems.size)) #[$elems,*] rfl) + +/-- Custom eliminator for `Vector α n` through `Array α` -/ +@[elab_as_elim] +def elimAsArray {motive : ∀ {n}, Vector α n → Sort u} (mk : ∀ a : Array α, motive ⟨a, rfl⟩) : + {n : Nat} → (v : Vector α n) → motive v + | _, ⟨a, rfl⟩ => mk a + +/-- Custom eliminator for `Vector α n` through `List α` -/ +@[elab_as_elim] +def elimAsList {motive : ∀ {n}, Vector α n → Sort u} (mk : ∀ a : List α, motive ⟨⟨a⟩, rfl⟩) : + {n : Nat} → (v : Vector α n) → motive v + | _, ⟨⟨a⟩, rfl⟩ => mk a + +/-- `Vector.size` gives the size of a vector. -/ +@[nolint unusedArguments] +def size (_ : Vector α n) : Nat := n + +/-- `Vector.empty` produces an empty vector -/ +def empty : Vector α 0 := ⟨Array.empty, rfl⟩ + +/-- Make an empty vector with pre-allocated capacity-/ +def mkEmpty (capacity : Nat) : Vector α 0 := ⟨Array.mkEmpty capacity, rfl⟩ + +/-- Makes a vector of size `n` with all cells containing `v` -/ +def mkVector (n : Nat) (v : α) : Vector α n := ⟨mkArray n v, Array.size_mkArray ..⟩ + +/-- Returns a vector of size `1` with a single element `v` -/ +def singleton (v : α) : Vector α 1 := mkVector 1 v + +/-- +The Inhabited instance for `Vector α n` with `[Inhabited α]` produces a vector of size `n` +with all its elements equal to the `default` element of type `α` +-/ +instance [Inhabited α] : Inhabited (Vector α n) where + default := mkVector n default + +/-- The list obtained from a vector. -/ +def toList (v : Vector α n) : List α := v.toArray.toList + +/-- nth element of a vector, indexed by a `Fin` type. -/ +def get (v : Vector α n) (i : Fin n) : α := v.toArray.get <| i.cast v.size_eq.symm + +/-- Vector lookup function that takes an index `i` of type `USize` -/ +def uget (v : Vector α n) (i : USize) (h : i.toNat < n) : α := v.toArray.uget i (v.size_eq.symm ▸ h) + +/-- `Vector α n` nstance for the `GetElem` typeclass. -/ +instance : GetElem (Vector α n) Nat α fun _ i => i < n where + getElem := fun x i h => get x ⟨i, h⟩ + +/-- +`getD v i v₀` gets the `iᵗʰ` element of v if valid. +Otherwise it returns `v₀` by default +-/ +def getD (v : Vector α n) (i : Nat) (v₀ : α) : α := Array.getD v.toArray i v₀ + +/-- +`v.back! v` gets the last element of the vector. +panics if `v` is empty. +-/ +abbrev back! [Inhabited α] (v : Vector α n) : α := v[n - 1]! + +/-- +`v.back?` gets the last element `x` of the array as `some x` +if it exists. Else the vector is empty and it returns `none` +-/ +abbrev back? (v : Vector α n) : Option α := v[n - 1]? + +/-- `Vector.head` produces the head of a vector -/ +abbrev head (v : Vector α (n+1)) := v[0] + +/-- `push v x` pushes `x` to the end of vector `v` in O(1) time -/ +def push (x : α) (v : Vector α n) : Vector α (n + 1) := + ⟨v.toArray.push x, by simp [v.size_eq]⟩ + +/-- `pop v` returns the vector with the last element removed -/ +def pop (v : Vector α n) : Vector α (n - 1) := + ⟨Array.pop v.toArray, by simp [v.size_eq]⟩ + +/-- +Sets an element in a vector using a Fin index. + +This will perform the update destructively provided that a has a reference count of 1 when called. +-/ +def set (v : Vector α n) (i : Fin n) (x : α) : Vector α n := + ⟨v.toArray.set (Fin.cast v.size_eq.symm i) x, by simp [v.size_eq]⟩ + +/-- +`setN v i h x` sets an element in a vector using a Nat index which is provably valid. +By default a proof by `get_elem_tactic` is provided. + +This will perform the update destructively provided that a has a reference count of 1 when called. +-/ +def setN (v : Vector α n) (i : Nat) (x : α) (h : i < n := by get_elem_tactic) : Vector α n := + v.set ⟨i, h⟩ x + +/-- +Sets an element in a vector, or do nothing if the index is out of bounds. + +This will perform the update destructively provided that a has a reference count of 1 when called. +-/ +def setD (v : Vector α n) (i : Nat) (x : α) : Vector α n := + ⟨v.toArray.setD i x, by simp [v.size_eq]⟩ + +/-- +Sets an element in an array, or panic if the index is out of bounds. + +This will perform the update destructively provided that a has a reference count of 1 when called. +-/ +def set! (v : Vector α n) (i : Nat) (x : α) : Vector α n := + ⟨v.toArray.set! i x, by simp [v.size_eq]⟩ + +/-- Appends a vector to another. -/ +def append : Vector α n → Vector α m → Vector α (n + m) + | ⟨a₁, _⟩, ⟨a₂, _⟩ => ⟨a₁ ++ a₂, by simp [Array.size_append, *]⟩ + +instance : HAppend (Vector α n) (Vector α m) (Vector α (n + m)) where + hAppend := append + +/-- Creates a vector from another with a provably equal length. -/ +protected def cast {n m : Nat} (h : n = m) : Vector α n → Vector α m + | ⟨x, p⟩ => ⟨x, h ▸ p⟩ + +/-- +`extract v start halt` Returns the slice of `v` from indices `start` to `stop` (exclusive). +If `start` is greater or equal to `stop`, the result is empty. +If `stop` is greater than the size of `v`, the size is used instead. +-/ +def extract (v : Vector α n) (start stop : Nat) : Vector α (min stop n - start) := + ⟨Array.extract v.toArray start stop, by simp [v.size_eq]⟩ + +/-- Maps a vector under a function. -/ +def map (f : α → β) (v : Vector α n) : Vector β n := + ⟨v.toArray.map f, by simp [v.size_eq]⟩ + +/-- Maps two vectors under a curried function of two variables. -/ +def zipWith : (a : Vector α n) → (b : Vector β n) → (f : α → β → φ) → Vector φ n + | ⟨a, h₁⟩, ⟨b, h₂⟩, f => ⟨Array.zipWith a b f, by simp [Array.size_zipWith, h₁, h₂]⟩ + +/-- Returns a vector of length `n` from a function on `Fin n`. -/ +def ofFn (f : Fin n → α) : Vector α n := ⟨Array.ofFn f, Array.size_ofFn ..⟩ + +/-- +Swaps two entries in a Vector. + +This will perform the update destructively provided that `v` has a reference count of 1 when called. +-/ +def swap (v : Vector α n) (i j : Fin n) : Vector α n := + ⟨v.toArray.swap (Fin.cast v.size_eq.symm i) (Fin.cast v.size_eq.symm j), by simp [v.size_eq]⟩ + +/-- +`swapN v i j hi hj` swaps two `Nat` indexed entries in a `Vector α n`. +Uses `get_elem_tactic` to supply proofs `hi` and `hj` respectively +that the indices `i` and `j` are in range. + +This will perform the update destructively provided that `v` has a reference count of 1 when called. +-/ +def swapN (v : Vector α n) (i j : Nat) + (hi : i < n := by get_elem_tactic) (hj : j < n := by get_elem_tactic) : Vector α n := + v.swap ⟨i, hi⟩ ⟨j, hj⟩ + +/-- +Swaps two entries in a `Vector α n`, or panics if either index is out of bounds. + +This will perform the update destructively provided that `v` has a reference count of 1 when called. +-/ +def swap! (v : Vector α n) (i j : Nat) : Vector α n := + ⟨Array.swap! v.toArray i j, by simp [v.size_eq]⟩ + +/-- +Swaps the entry with index `i : Fin n` in the vector for a new entry. +The old entry is returned with the modified vector. +-/ +def swapAt (v : Vector α n) (i : Fin n) (x : α) : α × Vector α n := + let res := v.toArray.swapAt (Fin.cast v.size_eq.symm i) x + (res.1, ⟨res.2, by simp [Array.swapAt_def, res, v.size_eq]⟩) + +/-- +Swaps the entry with index `i : Nat` in the vector for a new entry `x`. +The old entry is returned alongwith the modified vector. + +Automatically generates a proof of `i < n` with `get_elem_tactic` where feasible. +-/ +def swapAtN (v : Vector α n) (i : Nat) (x : α) (h : i < n := by get_elem_tactic) : α × Vector α n := + swapAt v ⟨i, h⟩ x + +/-- +`swapAt! v i x` swaps out the entry at index `i : Nat` in the vector for `x`, if the index is valid. +Otherwise it panics The old entry is returned with the modified vector. +-/ +@[inline] def swapAt! (v : Vector α n) (i : Nat) (x : α) : α × Vector α n := + if h : i < n then + swapAt v ⟨i, h⟩ x + else + have : Inhabited α := ⟨x⟩ + panic! s!"Index {i} is out of bounds" + +/-- `range n` returns the vector `#v[0,1,2,...,n-1]`. -/ +def range (n : Nat) : Vector Nat n := ⟨Array.range n, Array.size_range ..⟩ + +/-- +`shrink v m` shrinks the vector to the first `m` elements if `m < n`. +Returns `v` unchanged if `m ≥ n`. +-/ +def shrink (v : Vector α n) (m : Nat) : Vector α (min n m) := + ⟨v.toArray.shrink m, by simp [Array.size_shrink, v.size_eq]⟩ + +/-- +Drops the first (up to) `i` elements from a vector of length `n`. +If `m ≥ n`, the return value is empty. +-/ +def drop (i : Nat) (v : Vector α n) : Vector α (n - i) := + have : min n n - i = n - i := by rw [Nat.min_self] + Vector.cast this (extract v i n) + +/-- +Takes the first (up to) `i` elements from a vector of length `n`. + +-/ +alias take := shrink + +/-- +`isEqv` takes a given boolean property `p`. It returns `true` +if and only if `p a[i] b[i]` holds true for all valid indices `i`. +-/ +@[inline] def isEqv (a b : Vector α n) (p : α → α → Bool) : Bool := + Array.isEqvAux a.toArray b.toArray (a.size_eq.trans b.size_eq.symm) p 0 + +instance [BEq α] : BEq (Vector α n) := + ⟨fun a b => isEqv a b BEq.beq⟩ + +proof_wanted lawfulBEq [BEq α] [LawfulBEq α] : LawfulBEq (Vector α n) + +/-- `reverse v` reverses the vector `v`. -/ +def reverse (v : Vector α n) : Vector α n := + ⟨v.toArray.reverse, by simp [v.size_eq]⟩ + +/-- `O(|v| - i)`. `feraseIdx v i` removes the element at position `i` in vector `v`. -/ +def feraseIdx (v : Vector α n) (i : Fin n) : Vector α (n-1) := + ⟨v.toArray.feraseIdx (Fin.cast v.size_eq.symm i), by simp [Array.size_feraseIdx, v.size_eq]⟩ + +/-- `Vector.tail` produces the tail of the vector `v`. -/ +@[inline] def tail (v : Vector α n) : Vector α (n-1) := + match n with + | 0 => v + | _ + 1 => Vector.feraseIdx v 0 + +/-- +`O(|v| - i)`. `eraseIdx! v i` removes the element at position `i` from vector `v` of length `n`. +Panics if `i` is not less than `n`. +-/ +@[inline] def eraseIdx! (v : Vector α n) (i : Nat) : Vector α (n-1) := + if h : i < n then + feraseIdx v ⟨i,h⟩ + else + have : Inhabited (Vector α (n-1)) := ⟨v.tail⟩ + panic! s!"Index {i} is out of bounds" + +/-- +`eraseIdxN v i h` removes the element at position `i` from a vector of length `n`. +`h : i < n` has a default argument `by get_elem_tactic` which tries to supply a proof +that the index is valid. + +This function takes worst case O(n) time because it has to backshift all elements at positions +greater than i. +-/ +abbrev eraseIdxN (v : Vector α n) (i : Nat) (h : i < n := by get_elem_tactic) : Vector α (n - 1) := + v.feraseIdx ⟨i, h⟩ + +/-- +If `x` is an element of vector `v` at index `j`, then `indexOf? v x` returns `some j`. +Otherwise it returns `none`. +-/ +def indexOf? [BEq α] (v : Vector α n) (x : α) : Option (Fin n) := + match Array.indexOf? v.toArray x with + | some res => some (Fin.cast v.size_eq res) + | none => none + +/-- `isPrefixOf as bs` returns true iff vector `as` is a prefix of vector`bs` -/ +def isPrefixOf [BEq α] (as : Vector α m) (bs : Vector α n) : Bool := + Array.isPrefixOf as.toArray bs.toArray + +/-- `allDiff as i` returns `true` when all elements of `v` are distinct from each other` -/ +def allDiff [BEq α] (as : Vector α n) : Bool := + Array.allDiff as.toArray diff --git a/Batteries/Data/Vector/Lemmas.lean b/Batteries/Data/Vector/Lemmas.lean new file mode 100644 index 0000000000..02acccb8ce --- /dev/null +++ b/Batteries/Data/Vector/Lemmas.lean @@ -0,0 +1,47 @@ +/- +Copyright (c) 2024 Shreyas Srinivas. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Shreyas Srinivas, Francois Dorais +-/ + +import Batteries.Data.Vector.Basic +import Batteries.Data.List.Basic +import Batteries.Data.List.Lemmas +import Batteries.Data.Array.Lemmas + +/-! +## Vectors +Lemmas about `Vector α n` +-/ + +namespace Batteries + +namespace Vector + + +/-- An `empty` vector maps to a `empty` vector. -/ +@[simp] +theorem map_empty (f : α → β) : map f empty = empty := by + rw [map, empty, mk.injEq] + exact Array.map_empty f + +theorem toArray_injective : ∀ {v w : Vector α n}, v.toArray = w.toArray → v = w + | {..}, {..}, rfl => rfl + +/-- A vector of length `0` is an `empty` vector. -/ +protected theorem eq_empty (v : Vector α 0) : v = empty := by + apply Vector.toArray_injective + apply Array.eq_empty_of_size_eq_zero v.2 + +/-- +`Vector.ext` is an extensionality theorem. +Vectors `a` and `b` are equal to each other if their elements are equal for each valid index. +-/ +@[ext] +protected theorem ext {a b : Vector α n} (h : (i : Nat) → (_ : i < n) → a[i] = b[i]) : a = b := by + apply Vector.toArray_injective + apply Array.ext + · rw [a.size_eq, b.size_eq] + · intro i hi _ + rw [a.size_eq] at hi + exact h i hi From f2c939c3986ee3a6480b3811588a11ee008ab3b2 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Sat, 17 Aug 2024 00:50:34 +0200 Subject: [PATCH 052/177] feat: RBMap lemmas (#739) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: François G. Dorais --- Batteries/Data/RBMap/Lemmas.lean | 123 +++++++++++++++++++++++++++++++ 1 file changed, 123 insertions(+) diff --git a/Batteries/Data/RBMap/Lemmas.lean b/Batteries/Data/RBMap/Lemmas.lean index 5577b3a1a7..43d31af8f7 100644 --- a/Batteries/Data/RBMap/Lemmas.lean +++ b/Batteries/Data/RBMap/Lemmas.lean @@ -1103,3 +1103,126 @@ theorem lowerBound?_lt {t : RBSet α cmp} [TransCmp cmp] lowerBoundP?_lt H hz end RBSet + +namespace RBMap + +-- @[simp] -- FIXME: RBSet.val_toList already triggers here, seems bad? +theorem val_toList {t : RBMap α β cmp} : t.1.toList = t.toList := rfl + +@[simp] theorem mkRBSet_eq : mkRBMap α β cmp = ∅ := rfl +@[simp] theorem empty_eq : @RBMap.empty α β cmp = ∅ := rfl +@[simp] theorem default_eq : (default : RBMap α β cmp) = ∅ := rfl +@[simp] theorem empty_toList : toList (∅ : RBMap α β cmp) = [] := rfl +@[simp] theorem single_toList : toList (single a b : RBMap α β cmp) = [(a, b)] := rfl + +theorem mem_toList {t : RBMap α β cmp} : x ∈ toList t ↔ x ∈ t.1 := RBNode.mem_toList + +theorem foldl_eq_foldl_toList {t : RBMap α β cmp} : + t.foldl f init = t.toList.foldl (fun r p => f r p.1 p.2) init := + RBNode.foldl_eq_foldl_toList + +theorem foldr_eq_foldr_toList {t : RBMap α β cmp} : + t.foldr f init = t.toList.foldr (fun p r => f p.1 p.2 r) init := + RBNode.foldr_eq_foldr_toList + +theorem foldlM_eq_foldlM_toList [Monad m] [LawfulMonad m] {t : RBMap α β cmp} : + t.foldlM (m := m) f init = t.toList.foldlM (fun r p => f r p.1 p.2) init := + RBNode.foldlM_eq_foldlM_toList + +theorem forM_eq_forM_toList [Monad m] [LawfulMonad m] {t : RBMap α β cmp} : + t.forM (m := m) f = t.toList.forM (fun p => f p.1 p.2) := + RBNode.forM_eq_forM_toList + +theorem forIn_eq_forIn_toList [Monad m] [LawfulMonad m] {t : RBMap α β cmp} : + forIn (m := m) t init f = forIn t.toList init f := RBNode.forIn_eq_forIn_toList + +theorem toStream_eq {t : RBMap α β cmp} : toStream t = t.1.toStream .nil := rfl + +@[simp] theorem toStream_toList {t : RBMap α β cmp} : (toStream t).toList = t.toList := + RBSet.toStream_toList + +theorem toList_sorted {t : RBMap α β cmp} : t.toList.Pairwise (RBNode.cmpLT (cmp ·.1 ·.1)) := + RBSet.toList_sorted + +theorem findEntry?_some_eq_eq {t : RBMap α β cmp} : t.findEntry? x = some (y, v) → cmp x y = .eq := + RBSet.findP?_some_eq_eq + +theorem findEntry?_some_mem_toList {t : RBMap α β cmp} (h : t.findEntry? x = some y) : + y ∈ toList t := RBSet.findP?_some_mem_toList h + +theorem find?_some_mem_toList {t : RBMap α β cmp} (h : t.find? x = some v) : + ∃ y, (y, v) ∈ toList t ∧ cmp x y = .eq := by + obtain ⟨⟨y, v⟩, h', rfl⟩ := Option.map_eq_some'.1 h + exact ⟨_, findEntry?_some_mem_toList h', findEntry?_some_eq_eq h'⟩ + +theorem mem_toList_unique [@TransCmp α cmp] {t : RBMap α β cmp} + (hx : x ∈ toList t) (hy : y ∈ toList t) (e : cmp x.1 y.1 = .eq) : x = y := + RBSet.mem_toList_unique hx hy e + +/-- A "representable cut" is one generated by `cmp a` for some `a`. This is always a valid cut. -/ +instance (cmp) (a : α) : IsStrictCut cmp (cmp a) where + le_lt_trans h₁ h₂ := TransCmp.lt_le_trans h₂ h₁ + le_gt_trans h₁ := Decidable.not_imp_not.1 (TransCmp.le_trans · h₁) + exact h := (TransCmp.cmp_congr_left h).symm + +instance (f : α → β) (cmp) [@TransCmp β cmp] (x : β) : + IsStrictCut (Ordering.byKey f cmp) (fun y => cmp x (f y)) where + le_lt_trans h₁ h₂ := TransCmp.lt_le_trans h₂ h₁ + le_gt_trans h₁ := Decidable.not_imp_not.1 (TransCmp.le_trans · h₁) + exact h := (TransCmp.cmp_congr_left h).symm + +theorem findEntry?_some [@TransCmp α cmp] {t : RBMap α β cmp} : + t.findEntry? x = some y ↔ y ∈ toList t ∧ cmp x y.1 = .eq := RBSet.findP?_some + +theorem find?_some [@TransCmp α cmp] {t : RBMap α β cmp} : + t.find? x = some v ↔ ∃ y, (y, v) ∈ toList t ∧ cmp x y = .eq := by + simp only [find?, findEntry?_some, Option.map_eq_some']; constructor + · rintro ⟨_, h, rfl⟩; exact ⟨_, h⟩ + · rintro ⟨b, h⟩; exact ⟨_, h, rfl⟩ + +theorem contains_iff_findEntry? {t : RBMap α β cmp} : + t.contains x ↔ ∃ v, t.findEntry? x = some v := Option.isSome_iff_exists + +theorem contains_iff_find? {t : RBMap α β cmp} : + t.contains x ↔ ∃ v, t.find? x = some v := by + simp [contains_iff_findEntry?, find?, and_comm, exists_comm] + +theorem size_eq (t : RBMap α β cmp) : t.size = t.toList.length := RBNode.size_eq + +theorem mem_toList_insert_self (v) (t : RBMap α β cmp) : (k, v) ∈ toList (t.insert k v) := + RBSet.mem_toList_insert_self .. + +theorem mem_toList_insert_of_mem (v) {t : RBMap α β cmp} (h : y ∈ toList t) : + y ∈ toList (t.insert k v) ∨ cmp k y.1 = .eq := RBSet.mem_toList_insert_of_mem _ h + +theorem mem_toList_insert [@TransCmp α cmp] {t : RBMap α β cmp} : + y ∈ toList (t.insert k v) ↔ (y ∈ toList t ∧ t.findEntry? k ≠ some y) ∨ y = (k, v) := + RBSet.mem_toList_insert + +theorem findEntry?_congr [@TransCmp α cmp] (t : RBMap α β cmp) (h : cmp k₁ k₂ = .eq) : + t.findEntry? k₁ = t.findEntry? k₂ := by simp [findEntry?, TransCmp.cmp_congr_left' h] + +theorem find?_congr [@TransCmp α cmp] (t : RBMap α β cmp) (h : cmp k₁ k₂ = .eq) : + t.find? k₁ = t.find? k₂ := by simp [find?, findEntry?_congr _ h] + +theorem findEntry?_insert_of_eq [@TransCmp α cmp] (t : RBMap α β cmp) (h : cmp k' k = .eq) : + (t.insert k v).findEntry? k' = some (k, v) := RBSet.findP?_insert_of_eq _ h + +theorem find?_insert_of_eq [@TransCmp α cmp] (t : RBMap α β cmp) (h : cmp k' k = .eq) : + (t.insert k v).find? k' = some v := by rw [find?, findEntry?_insert_of_eq _ h]; rfl + +theorem findEntry?_insert_of_ne [@TransCmp α cmp] (t : RBMap α β cmp) (h : cmp k' k ≠ .eq) : + (t.insert k v).findEntry? k' = t.findEntry? k' := RBSet.findP?_insert_of_ne _ h + +theorem find?_insert_of_ne [@TransCmp α cmp] (t : RBMap α β cmp) (h : cmp k' k ≠ .eq) : + (t.insert k v).find? k' = t.find? k' := by simp [find?, findEntry?_insert_of_ne _ h] + +theorem findEntry?_insert [@TransCmp α cmp] (t : RBMap α β cmp) (k v k') : + (t.insert k v).findEntry? k' = if cmp k' k = .eq then some (k, v) else t.findEntry? k' := + RBSet.findP?_insert .. + +theorem find?_insert [@TransCmp α cmp] (t : RBMap α β cmp) (k v k') : + (t.insert k v).find? k' = if cmp k' k = .eq then some v else t.find? k' := by + split <;> [exact find?_insert_of_eq t ‹_›; exact find?_insert_of_ne t ‹_›] + +end RBMap From f7aebca12535aff8577e927e7f5bfc185c26e0c4 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Fri, 16 Aug 2024 19:04:06 -0400 Subject: [PATCH 053/177] fix: splitAtD reverse bug from #919 --- Batteries/Data/List/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index 7e6e034dc4..679e8c505d 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -137,7 +137,7 @@ def splitAtD (n : Nat) (l : List α) (dflt : α) : List α × List α := go n l if `splitAtD n l dflt = (left, right)`. -/ go : Nat → List α → List α → List α × List α | n+1, x :: xs, acc => go n xs (x :: acc) - | 0, xs, acc => (acc, xs) + | 0, xs, acc => (acc.reverse, xs) | n, [], acc => (acc.reverseAux (replicate n dflt), []) /-- From d3091570b4a34a4f77ef2cdc227838edf625dc85 Mon Sep 17 00:00:00 2001 From: damiano Date: Sat, 17 Aug 2024 01:37:12 +0200 Subject: [PATCH 054/177] feat: add lemma code-action (#625) Co-authored-by: F. G. Dorais --- Batteries.lean | 1 + Batteries/Tactic/Lemma.lean | 50 +++++++++++++++++++++++++++++++++++++ test/lemma_cmd.lean | 34 +++++++++++++++++++++++++ 3 files changed, 85 insertions(+) create mode 100644 Batteries/Tactic/Lemma.lean create mode 100644 test/lemma_cmd.lean diff --git a/Batteries.lean b/Batteries.lean index 5fd8bc72ad..d607662196 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -82,6 +82,7 @@ import Batteries.Tactic.Congr import Batteries.Tactic.Exact import Batteries.Tactic.Init import Batteries.Tactic.Instances +import Batteries.Tactic.Lemma import Batteries.Tactic.Lint import Batteries.Tactic.Lint.Basic import Batteries.Tactic.Lint.Frontend diff --git a/Batteries/Tactic/Lemma.lean b/Batteries/Tactic/Lemma.lean new file mode 100644 index 0000000000..10fdec7a0e --- /dev/null +++ b/Batteries/Tactic/Lemma.lean @@ -0,0 +1,50 @@ +/- +Copyright (c) 2024 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin, Damiano Testa +-/ +import Lean.Meta.Tactic.TryThis + +/-! +# Control for `lemma` command + +The `lemma` command exists in `Mathlib`, but not in `Std`. + +This file enforces the convention by introducing a code-action +to replace `lemma` by `theorem`. +-/ + +namespace Batteries.Tactic.Lemma + +open Lean Elab.Command Meta.Tactic + +/-- Enables the use of `lemma` as a synonym for `theorem` -/ +register_option lang.lemmaCmd : Bool := { + defValue := false + descr := "enable the use of the `lemma` command as a synonym for `theorem`" +} + +/-- Check whether `lang.lemmaCmd` option is enabled -/ +def checkLangLemmaCmd (o : Options) : Bool := o.get `lang.lemmaCmd lang.lemmaCmd.defValue + +/-- `lemma` is not supported, please use `theorem` instead -/ +syntax (name := lemmaCmd) declModifiers + group("lemma " declId ppIndent(declSig) declVal) : command + +/-- Elaborator for the `lemma` command, if the option `lang.lemmaCmd` is false the command +emits a warning and code action instructing the user to use `theorem` instead.-/ +@[command_elab «lemmaCmd»] def elabLemma : CommandElab := fun stx => do + unless checkLangLemmaCmd (← getOptions) do + let lemmaStx := stx[1][0] + Elab.Command.liftTermElabM <| + TryThis.addSuggestion lemmaStx { suggestion := "theorem" } + logErrorAt lemmaStx + "`lemma` is not supported by default, please use `theorem` instead.\n\ + Use `set_option lang.lemmaCmd true` to enable the use of the `lemma` command in a file.\n\ + Use the command line option `-Dlang.lemmaCmd=true` to enable the use of `lemma` globally." + let out ← Elab.liftMacroM <| do + let stx := stx.modifyArg 1 fun stx => + let stx := stx.modifyArg 0 (mkAtomFrom · "theorem" (canonical := true)) + stx.setKind ``Parser.Command.theorem + pure <| stx.setKind ``Parser.Command.declaration + Elab.Command.elabCommand out diff --git a/test/lemma_cmd.lean b/test/lemma_cmd.lean new file mode 100644 index 0000000000..7952f5766f --- /dev/null +++ b/test/lemma_cmd.lean @@ -0,0 +1,34 @@ +import Batteries.Tactic.Lemma + +-- lemma disabled by default +/-- +info: Try this: theorem +--- +error: `lemma` is not supported by default, please use `theorem` instead. +Use `set_option lang.lemmaCmd true` to enable the use of the `lemma` command in a file. +Use the command line option `-Dlang.lemmaCmd=true` to enable the use of `lemma` globally. +-/ +#guard_msgs in +lemma test1 : 3 < 7 := by decide + +-- lemma enabled for one command +set_option lang.lemmaCmd true in +lemma test2 : 3 < 7 := by decide + +-- lemma disabled again +/-- +info: Try this: theorem +--- +error: `lemma` is not supported by default, please use `theorem` instead. +Use `set_option lang.lemmaCmd true` to enable the use of the `lemma` command in a file. +Use the command line option `-Dlang.lemmaCmd=true` to enable the use of `lemma` globally. +-/ +#guard_msgs in +lemma test3 : 3 < 7 := by decide + +-- lemma enabled for rest of file +set_option lang.lemmaCmd true + +lemma test4 : 3 < 7 := by decide + +lemma test5 : 3 < 7 := by decide From b1a2ec94f9752a45ca655ab48286a1d8d9df78b8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Fri, 16 Aug 2024 20:27:21 -0400 Subject: [PATCH 055/177] feat: `List.insertP` and lemmas (#569) Co-authored-by: Mario Carneiro --- Batteries/Data/List/Basic.lean | 12 ++++++++++++ Batteries/Data/List/Lemmas.lean | 27 +++++++++++++++++++++++++++ Batteries/Data/List/Perm.lean | 11 +++++++++++ 3 files changed, 50 insertions(+) diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index 679e8c505d..2d6e73a14b 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -1180,6 +1180,18 @@ See `isSubperm_iff` for a characterization in terms of `List.Subperm`. -/ def isSubperm [BEq α] (l₁ l₂ : List α) : Bool := ∀ x ∈ l₁, count x l₁ ≤ count x l₂ +/-- +`O(|l|)`. Inserts `a` in `l` right before the first element such that `p` is true, or at the end of +the list if `p` always false on `l`. +-/ +def insertP (p : α → Bool) (a : α) (l : List α) : List α := + loop l [] +where + /-- Inner loop for `insertP`. Tail recursive. -/ + loop : List α → List α → List α + | [], r => reverseAux (a :: r) [] -- Note: `reverseAux` is tail recursive. + | b :: l, r => bif p b then reverseAux (a :: r) (b :: l) else loop l (b :: r) + /-- `O(|l| + |r|)`. Merge two lists using `s` as a switch. -/ diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 3eea690ad6..8d8b3e1cd6 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -562,10 +562,37 @@ theorem indexOf_mem_indexesOf [BEq α] [LawfulBEq α] {xs : List α} (m : x ∈ specialize ih m simpa +/-! ### insertP -/ + +theorem insertP_loop (a : α) (l r : List α) : + insertP.loop p a l r = reverseAux r (insertP p a l) := by + induction l generalizing r with simp [insertP, insertP.loop, cond] + | cons b l ih => rw [ih (b :: r), ih [b]]; split <;> rfl + +@[simp] theorem insertP_nil (p : α → Bool) (a) : insertP p a [] = [a] := rfl + +@[simp] theorem insertP_cons_pos (p : α → Bool) (a b l) (h : p b) : + insertP p a (b :: l) = a :: b :: l := by + simp only [insertP, insertP.loop, cond, h]; rfl + +@[simp] theorem insertP_cons_neg (p : α → Bool) (a b l) (h : ¬ p b) : + insertP p a (b :: l) = b :: insertP p a l := by + simp only [insertP, insertP.loop, cond, h]; exact insertP_loop .. + +@[simp] theorem length_insertP (p : α → Bool) (a l) : (insertP p a l).length = l.length + 1 := by + induction l with simp [insertP, insertP.loop, cond] + | cons _ _ ih => split <;> simp [insertP_loop, ih] + +@[simp] theorem mem_insertP (p : α → Bool) (a l) : a ∈ insertP p a l := by + induction l with simp [insertP, insertP.loop, cond] + | cons _ _ ih => split <;> simp [insertP_loop, ih] + theorem merge_loop_nil_left (s : α → α → Bool) (r t) : merge.loop s [] r t = reverseAux t r := by rw [merge.loop] +/-! ### merge -/ + theorem merge_loop_nil_right (s : α → α → Bool) (l t) : merge.loop s l [] t = reverseAux t l := by cases l <;> rw [merge.loop]; intro; contradiction diff --git a/Batteries/Data/List/Perm.lean b/Batteries/Data/List/Perm.lean index 5c31720c7c..ed2a18f5cf 100644 --- a/Batteries/Data/List/Perm.lean +++ b/Batteries/Data/List/Perm.lean @@ -700,6 +700,17 @@ theorem Perm.eraseP (f : α → Bool) {l₁ l₂ : List α} refine (IH₁ H).trans (IH₂ ((p₁.pairwise_iff ?_).1 H)) exact fun h h₁ h₂ => h h₂ h₁ +theorem perm_insertP (p : α → Bool) (a l) : insertP p a l ~ a :: l := by + induction l with simp [insertP, insertP.loop, cond] + | cons _ _ ih => + split + · exact Perm.refl .. + · rw [insertP_loop, reverseAux, reverseAux] + exact Perm.trans (Perm.cons _ ih) (Perm.swap ..) + +theorem Perm.insertP (p : α → Bool) (a) (h : l₁ ~ l₂) : insertP p a l₁ ~ insertP p a l₂ := + Perm.trans (perm_insertP ..) <| Perm.trans (Perm.cons _ h) <| Perm.symm (perm_insertP ..) + theorem perm_merge (s : α → α → Bool) (l r) : merge s l r ~ l ++ r := by match l, r with | [], r => simp From 65f464e6a5f30b2c2a58569f530c4677b371cec6 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 17 Aug 2024 12:34:20 +1000 Subject: [PATCH 056/177] chore: remove unnecessary case for splitAt.go (#921) --- Batteries/Data/List/Basic.lean | 4 +--- Batteries/Data/List/Lemmas.lean | 16 ++++++++++++++++ 2 files changed, 17 insertions(+), 3 deletions(-) diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index 2d6e73a14b..fc2cecae44 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -117,10 +117,8 @@ splitAt 2 [a, b, c] = ([a, b], [c]) ``` -/ def splitAt (n : Nat) (l : List α) : List α × List α := go l n [] where - /-- Auxiliary for `splitAt`: `splitAt.go l n xs acc = (acc.reverse ++ take n xs, drop n xs)` - if `n < length xs`, else `(l, [])`. -/ + /-- Auxiliary for `splitAt`: `splitAt.go xs n acc = (acc.reverse ++ take n xs, drop n xs)`. -/ go : List α → Nat → List α → List α × List α - | [], _, _ => (l, []) | x :: xs, n+1, acc => go xs n (x :: acc) | xs, _, acc => (acc.reverse, xs) diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 8d8b3e1cd6..9e5f4e3ace 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -189,6 +189,22 @@ theorem get?_set_of_lt' (a : α) {m n} (l : List α) (h : m < length l) : @[deprecated (since := "2024-05-06")] alias length_removeNth := length_eraseIdx +/-! ### splitAt -/ + +theorem splitAt_go (n : Nat) (l acc : List α) : + splitAt.go l n acc = (acc.reverse ++ l.take n, l.drop n) := by + induction l generalizing n acc with + | nil => simp [splitAt.go] + | cons x xs ih => + cases n with + | zero => simp [splitAt.go] + | succ n => + rw [splitAt.go, take_succ_cons, drop_succ_cons, ih n (x :: acc), + reverse_cons, append_assoc, singleton_append] + +theorem splitAt_eq (n : Nat) (l : List α) : splitAt n l = (l.take n, l.drop n) := by + rw [splitAt, splitAt_go, reverse_nil, nil_append] + /-! ### eraseP -/ @[simp] theorem extractP_eq_find?_eraseP From a975dea2c4d8258a55b4f9861c537e2bb0f9ef63 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 17 Aug 2024 16:00:57 +1000 Subject: [PATCH 057/177] chore: restore splitAt pointer equality behaviour (#922) --- Batteries/Data/List/Basic.lean | 8 +++++++- Batteries/Data/List/Lemmas.lean | 9 ++++++--- 2 files changed, 13 insertions(+), 4 deletions(-) diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index fc2cecae44..ca8d6ca04b 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -117,8 +117,14 @@ splitAt 2 [a, b, c] = ([a, b], [c]) ``` -/ def splitAt (n : Nat) (l : List α) : List α × List α := go l n [] where - /-- Auxiliary for `splitAt`: `splitAt.go xs n acc = (acc.reverse ++ take n xs, drop n xs)`. -/ + /-- + Auxiliary for `splitAt`: + `splitAt.go l xs n acc = (acc.reverse ++ take n xs, drop n xs)` if `n < xs.length`, + and `(l, [])` otherwise. + -/ go : List α → Nat → List α → List α × List α + | [], _, _ => (l, []) -- This branch ensures the pointer equality of the result with the input + -- without any runtime branching cost. | x :: xs, n+1, acc => go xs n (x :: acc) | xs, _, acc => (acc.reverse, xs) diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 9e5f4e3ace..74d44af907 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -192,18 +192,21 @@ theorem get?_set_of_lt' (a : α) {m n} (l : List α) (h : m < length l) : /-! ### splitAt -/ theorem splitAt_go (n : Nat) (l acc : List α) : - splitAt.go l n acc = (acc.reverse ++ l.take n, l.drop n) := by - induction l generalizing n acc with + splitAt.go l xs n acc = + if n < xs.length then (acc.reverse ++ xs.take n, xs.drop n) else (l, []) := by + induction xs generalizing n acc with | nil => simp [splitAt.go] | cons x xs ih => cases n with | zero => simp [splitAt.go] | succ n => rw [splitAt.go, take_succ_cons, drop_succ_cons, ih n (x :: acc), - reverse_cons, append_assoc, singleton_append] + reverse_cons, append_assoc, singleton_append, length_cons] + simp only [Nat.succ_lt_succ_iff] theorem splitAt_eq (n : Nat) (l : List α) : splitAt n l = (l.take n, l.drop n) := by rw [splitAt, splitAt_go, reverse_nil, nil_append] + split <;> simp_all [take_of_length_le, drop_of_length_le] /-! ### eraseP -/ From d747f070e42dd21e2649b75090f5b0d45c6ec8e0 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 21 Aug 2024 12:44:54 +1000 Subject: [PATCH 058/177] chore: make some proofs more robust (#928) --- Batteries/Data/String/Lemmas.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Batteries/Data/String/Lemmas.lean b/Batteries/Data/String/Lemmas.lean index df338e03ff..cfb2058c24 100644 --- a/Batteries/Data/String/Lemmas.lean +++ b/Batteries/Data/String/Lemmas.lean @@ -174,7 +174,7 @@ theorem utf8GetAux?_of_valid (cs cs' : List Char) {i p : Nat} (hp : i + utf8Len | c::cs, cs' => simp only [utf8GetAux?, List.append_eq] rw [if_neg] - case hnc => simp [← hp, Pos.ext_iff]; exact ne_self_add_add_utf8Size + case hnc => simp only [← hp, Pos.ext_iff]; exact ne_self_add_add_utf8Size refine utf8GetAux?_of_valid cs cs' ?_ simpa [Nat.add_assoc, Nat.add_comm] using hp @@ -189,7 +189,7 @@ theorem utf8SetAux_of_valid (c' : Char) (cs cs' : List Char) {i p : Nat} (hp : i | c::cs, cs' => simp only [utf8SetAux, List.append_eq, List.cons_append] rw [if_neg] - case hnc => simp [← hp, Pos.ext_iff]; exact ne_self_add_add_utf8Size + case hnc => simp only [← hp, Pos.ext_iff]; exact ne_self_add_add_utf8Size refine congrArg (c::·) (utf8SetAux_of_valid c' cs cs' ?_) simpa [Nat.add_assoc, Nat.add_comm] using hp From 1851cf1aa9aa386aee92311d406c80840d7e7e22 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 26 Aug 2024 16:53:14 +1000 Subject: [PATCH 059/177] chore: robustify a proof (#934) --- test/lint_lean.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/test/lint_lean.lean b/test/lint_lean.lean index 2b2f8b89e1..46eb6591ff 100644 --- a/test/lint_lean.lean +++ b/test/lint_lean.lean @@ -18,11 +18,11 @@ but it is useful to run locally to see what the linters would catch. /-! Failing lints that need work. -/ --- #lint only explicitVarsOfIff in all -- Found 109 errors +-- #lint only explicitVarsOfIff in all -- Found 156 errors -- Many fixed in https://github.com/leanprover/lean4/pull/4620 -- and should be checked again. --- #lint only simpNF in all -- Found 34 errors +-- #lint only simpNF in all -- Found 12 errors /-! Lints that fail, but that we're not intending to do anything about. -/ From 31fb27d6b89dc94cf7349df247fc44d2a1d130af Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 26 Aug 2024 17:01:34 +1000 Subject: [PATCH 060/177] chore: robustify a proof (#935) --- Batteries/Data/List/Lemmas.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 74d44af907..279174f7b1 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -160,7 +160,7 @@ theorem exists_of_set' {l : List α} (h : n < l.length) : @[deprecated getElem?_set_eq' (since := "2024-06-12")] theorem get?_set_eq (a : α) (n) (l : List α) : (set l n a).get? n = (fun _ => a) <$> l.get? n := by - simp + simp only [get?_eq_getElem?, getElem?_set_eq', Option.map_eq_map] theorem getElem?_set_eq_of_lt (a : α) {n} {l : List α} (h : n < length l) : (set l n a)[n]? = some a := by rw [getElem?_set_eq', getElem?_eq_getElem h]; rfl From e6d3a32d66252a70fda1d56463e1da975b3b8f53 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 27 Aug 2024 11:45:40 +1000 Subject: [PATCH 061/177] chore: make proofs robust to incoming changes (#929) --- Batteries/Data/List/Lemmas.lean | 2 +- Batteries/Data/String/Lemmas.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 279174f7b1..fd82629a8a 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -586,7 +586,7 @@ theorem indexOf_mem_indexesOf [BEq α] [LawfulBEq α] {xs : List α} (m : x ∈ theorem insertP_loop (a : α) (l r : List α) : insertP.loop p a l r = reverseAux r (insertP p a l) := by induction l generalizing r with simp [insertP, insertP.loop, cond] - | cons b l ih => rw [ih (b :: r), ih [b]]; split <;> rfl + | cons b l ih => rw [ih (b :: r), ih [b]]; split <;> simp @[simp] theorem insertP_nil (p : α → Bool) (a) : insertP p a [] = [a] := rfl diff --git a/Batteries/Data/String/Lemmas.lean b/Batteries/Data/String/Lemmas.lean index cfb2058c24..a8d2bfa7b4 100644 --- a/Batteries/Data/String/Lemmas.lean +++ b/Batteries/Data/String/Lemmas.lean @@ -59,7 +59,7 @@ private theorem ne_self_add_add_utf8Size : i ≠ i + (n + Char.utf8Size c) := @[simp] theorem utf8Len_reverseAux (cs₁ cs₂) : utf8Len (cs₁.reverseAux cs₂) = utf8Len cs₁ + utf8Len cs₂ := by - induction cs₁ generalizing cs₂ <;> simp [*, ← Nat.add_assoc, Nat.add_right_comm] + induction cs₁ generalizing cs₂ <;> simp_all [← Nat.add_assoc, Nat.add_right_comm] @[simp] theorem utf8Len_reverse (cs) : utf8Len cs.reverse = utf8Len cs := utf8Len_reverseAux .. From a7fd140a94bbbfa40cf10839227bbb9e8492be2d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Fri, 30 Aug 2024 17:35:26 -0400 Subject: [PATCH 062/177] chore: bump toolchain to v4.11.0-rc3 (#939) --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index e7a4f40b89..28b8e55a50 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.11.0-rc2 +leanprover/lean4:v4.11.0-rc3 From e776546fa4ffd590bbc788ad165d46f1609efa8f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 2 Sep 2024 10:08:54 +1000 Subject: [PATCH 063/177] feat: runLinter uses ``withImportModules #[module, Batteries.Tactic.Lint]`` (#931) --- scripts/runLinter.lean | 14 +++++++++++++- 1 file changed, 13 insertions(+), 1 deletion(-) diff --git a/scripts/runLinter.lean b/scripts/runLinter.lean index 4f18be6b50..3fab25f834 100644 --- a/scripts/runLinter.lean +++ b/scripts/runLinter.lean @@ -47,12 +47,24 @@ unsafe def main (args : List String) : IO Unit := do stdin := .null } _ ← child.wait + -- If the linter is being run on a target that doesn't import `Batteries.Tactic.List`, + -- the linters are ineffective. So we import it here. + let lintModule := `Batteries.Tactic.Lint + let lintFile ← findOLean lintModule + unless (← lintFile.pathExists) do + -- run `lake build +Batteries.Tactic.Lint` (and ignore result) if the file hasn't been built yet + let child ← IO.Process.spawn { + cmd := (← IO.getEnv "LAKE").getD "lake" + args := #["build", s!"+{lintModule}"] + stdin := .null + } + _ ← child.wait let nolintsFile : FilePath := "scripts/nolints.json" let nolints ← if ← nolintsFile.pathExists then readJsonFile NoLints nolintsFile else pure #[] - withImportModules #[{module}] {} (trustLevel := 1024) fun env => + withImportModules #[module, lintModule] {} (trustLevel := 1024) fun env => let ctx := { fileName := "", fileMap := default } let state := { env } Prod.fst <$> (CoreM.toIO · ctx state) do From 9c6c2d647e57b2b7a0b42dd8080c698bd33a1b6f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 2 Sep 2024 11:36:48 +1000 Subject: [PATCH 064/177] chore: bump toolchain to v4.11.0 (#940) --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 28b8e55a50..5a9c76dc98 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.11.0-rc3 +leanprover/lean4:v4.11.0 From fa571ea02a804b52a59e58012b1e21fe4f0514f2 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 3 Sep 2024 13:37:15 +1000 Subject: [PATCH 065/177] chore: merge bump/v4.12.0, moving toolchain to v4.12.0-rc1 (#942) Co-authored-by: leanprover-community-mathlib4-bot Co-authored-by: Wojciech Nawrocki Co-authored-by: L Co-authored-by: Joachim Breitner Co-authored-by: Kyle Miller Co-authored-by: Matthew Ballard Co-authored-by: Jeremy Tan Jie Rui --- Batteries.lean | 1 - Batteries/Data/Array/Monadic.lean | 2 +- Batteries/Data/ByteArray.lean | 2 +- Batteries/Data/Fin/Lemmas.lean | 6 - Batteries/Data/LazyList.lean | 8 +- Batteries/Data/List/Basic.lean | 56 +--- Batteries/Data/List/Lemmas.lean | 71 +--- Batteries/Data/List/Perm.lean | 390 +--------------------- Batteries/Data/Nat/Basic.lean | 7 - Batteries/Data/Nat/Lemmas.lean | 2 +- Batteries/Data/RBMap/Basic.lean | 6 +- Batteries/Data/RBMap/Lemmas.lean | 8 +- Batteries/Data/String/Lemmas.lean | 2 +- Batteries/Data/Sum/Lemmas.lean | 8 +- Batteries/Data/Thunk.lean | 10 - Batteries/Data/UInt.lean | 40 --- Batteries/Data/UnionFind/Basic.lean | 6 +- Batteries/Lean/AttributeExtra.lean | 8 +- Batteries/Lean/HashMap.lean | 19 +- Batteries/Lean/HashSet.lean | 11 +- Batteries/Lean/Meta/Inaccessible.lean | 4 +- Batteries/Linter/UnnecessarySeqFocus.lean | 6 +- Batteries/Linter/UnreachableTactic.lean | 6 +- Batteries/Tactic/Lint/Frontend.lean | 6 +- Batteries/Tactic/Lint/Misc.lean | 4 +- lean-toolchain | 2 +- scripts/check_imports.lean | 2 +- test/where.lean | 1 + 28 files changed, 60 insertions(+), 634 deletions(-) delete mode 100644 Batteries/Data/Thunk.lean diff --git a/Batteries.lean b/Batteries.lean index d607662196..38e7617093 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -36,7 +36,6 @@ import Batteries.Data.Range import Batteries.Data.Rat import Batteries.Data.String import Batteries.Data.Sum -import Batteries.Data.Thunk import Batteries.Data.UInt import Batteries.Data.UnionFind import Batteries.Data.Vector diff --git a/Batteries/Data/Array/Monadic.lean b/Batteries/Data/Array/Monadic.lean index a9821eb689..316813972a 100644 --- a/Batteries/Data/Array/Monadic.lean +++ b/Batteries/Data/Array/Monadic.lean @@ -102,7 +102,7 @@ theorem SatisfiesM_anyM_iff_exists [Monad m] [LawfulMonad m] · intro | true, h => simp only [true_iff]; exact h | false, h => - simp only [false_iff] + simp only [false_iff, reduceCtorEq] exact h.2.imp fun ⟨j, h₁, h₂, hq⟩ => ⟨j, h₁, Nat.lt_min.2 ⟨h₂, j.2⟩, hq⟩ | inr hstart => rw [anyM_stop_le_start (h := hstart)] diff --git a/Batteries/Data/ByteArray.lean b/Batteries/Data/ByteArray.lean index a9653bf10c..e051e01fd2 100644 --- a/Batteries/Data/ByteArray.lean +++ b/Batteries/Data/ByteArray.lean @@ -140,7 +140,7 @@ private def ofFnAux (f : Fin n → UInt8) : ByteArray := go 0 (mkEmpty n) where termination_by n - i @[csimp] private theorem ofFn_eq_ofFnAux : @ofFn = @ofFnAux := by - funext n f; ext; simp [ofFnAux, Array.ofFn, data_ofFnAux, mkEmpty] + funext n f; ext1; simp [ofFnAux, Array.ofFn, data_ofFnAux, mkEmpty] where data_ofFnAux {n} (f : Fin n → UInt8) (i) {acc} : (ofFnAux.go f i acc).data = Array.ofFn.go f i acc.data := by diff --git a/Batteries/Data/Fin/Lemmas.lean b/Batteries/Data/Fin/Lemmas.lean index 4d4eb59a1a..bc1eb0d2c9 100644 --- a/Batteries/Data/Fin/Lemmas.lean +++ b/Batteries/Data/Fin/Lemmas.lean @@ -9,12 +9,6 @@ namespace Fin attribute [norm_cast] val_last -protected theorem le_antisymm_iff {x y : Fin n} : x = y ↔ x ≤ y ∧ y ≤ x := - Fin.ext_iff.trans Nat.le_antisymm_iff - -protected theorem le_antisymm {x y : Fin n} (h1 : x ≤ y) (h2 : y ≤ x) : x = y := - Fin.le_antisymm_iff.2 ⟨h1, h2⟩ - /-! ### clamp -/ @[simp] theorem coe_clamp (n m : Nat) : (clamp n m : Nat) = min n m := rfl diff --git a/Batteries/Data/LazyList.lean b/Batteries/Data/LazyList.lean index 426cdd2248..7a3074d1b6 100644 --- a/Batteries/Data/LazyList.lean +++ b/Batteries/Data/LazyList.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Simon Hudon -/ -import Batteries.Data.Thunk + /-! # Lazy lists @@ -229,8 +229,8 @@ protected def Mem {α} (x : α) : LazyList α → Prop | nil => False | cons y ys => x = y ∨ ys.get.Mem x -instance {α} : Membership α (LazyList α) := - ⟨LazyList.Mem⟩ +instance {α} : Membership α (LazyList α) where + mem l a := LazyList.Mem a l instance Mem.decidable {α} [DecidableEq α] (x : α) : ∀ xs : LazyList α, Decidable (x ∈ xs) | LazyList.nil => by @@ -243,7 +243,7 @@ instance Mem.decidable {α} [DecidableEq α] (x : α) : ∀ xs : LazyList α, De exact Or.inl h else by have := Mem.decidable x ys.get - have : (x ∈ ys.get) ↔ (x ∈ cons y ys) := by simp [(· ∈ ·), LazyList.Mem, h] + have : (x ∈ ys.get) ↔ (x ∈ cons y ys) := by simp [Membership.mem, LazyList.Mem, h] exact decidable_of_decidable_of_iff this @[simp] diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index ca8d6ca04b..a44f430c8d 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -110,24 +110,6 @@ Unlike `bagInter` this does not preserve multiplicity: `[1, 1].inter [1]` is `[1 instance [BEq α] : Inter (List α) := ⟨List.inter⟩ -/-- -Split a list at an index. -``` -splitAt 2 [a, b, c] = ([a, b], [c]) -``` --/ -def splitAt (n : Nat) (l : List α) : List α × List α := go l n [] where - /-- - Auxiliary for `splitAt`: - `splitAt.go l xs n acc = (acc.reverse ++ take n xs, drop n xs)` if `n < xs.length`, - and `(l, [])` otherwise. - -/ - go : List α → Nat → List α → List α × List α - | [], _, _ => (l, []) -- This branch ensures the pointer equality of the result with the input - -- without any runtime branching cost. - | x :: xs, n+1, acc => go xs n (x :: acc) - | xs, _, acc => (acc.reverse, xs) - /-- Split a list at an index. Ensures the left list always has the specified length by right padding with the provided default element. @@ -548,7 +530,7 @@ theorem sections_eq_nil_of_isEmpty : ∀ {L}, L.any isEmpty → @sections α L = | l :: L, h => by simp only [any, foldr, Bool.or_eq_true] at h match l, h with - | [], .inl rfl => simp; induction sections L <;> simp [*] + | [], .inl rfl => simp | l, .inr h => simp [sections, sections_eq_nil_of_isEmpty h] @[csimp] theorem sections_eq_sectionsTR : @sections = @sectionsTR := by @@ -1145,30 +1127,6 @@ protected def traverse [Applicative F] (f : α → F β) : List α → F (List | [] => pure [] | x :: xs => List.cons <$> f x <*> List.traverse f xs -/-- -`Perm l₁ l₂` or `l₁ ~ l₂` asserts that `l₁` and `l₂` are permutations -of each other. This is defined by induction using pairwise swaps. --/ -inductive Perm : List α → List α → Prop - /-- `[] ~ []` -/ - | nil : Perm [] [] - /-- `l₁ ~ l₂ → x::l₁ ~ x::l₂` -/ - | cons (x : α) {l₁ l₂ : List α} : Perm l₁ l₂ → Perm (x :: l₁) (x :: l₂) - /-- `x::y::l ~ y::x::l` -/ - | swap (x y : α) (l : List α) : Perm (y :: x :: l) (x :: y :: l) - /-- `Perm` is transitive. -/ - | trans {l₁ l₂ l₃ : List α} : Perm l₁ l₂ → Perm l₂ l₃ → Perm l₁ l₃ - -@[inherit_doc] scoped infixl:50 " ~ " => Perm - -/-- -`O(|l₁| * |l₂|)`. Computes whether `l₁` is a permutation of `l₂`. See `isPerm_iff` for a -characterization in terms of `List.Perm`. --/ -def isPerm [BEq α] : List α → List α → Bool - | [], l₂ => l₂.isEmpty - | a :: l₁, l₂ => l₂.contains a && l₁.isPerm (l₂.erase a) - /-- `Subperm l₁ l₂`, denoted `l₁ <+~ l₂`, means that `l₁` is a sublist of a permutation of `l₂`. This is an analogue of `l₁ ⊆ l₂` which respects @@ -1195,15 +1153,3 @@ where loop : List α → List α → List α | [], r => reverseAux (a :: r) [] -- Note: `reverseAux` is tail recursive. | b :: l, r => bif p b then reverseAux (a :: r) (b :: l) else loop l (b :: r) - -/-- -`O(|l| + |r|)`. Merge two lists using `s` as a switch. --/ -def merge (s : α → α → Bool) (l r : List α) : List α := - loop l r [] -where - /-- Inner loop for `List.merge`. Tail recursive. -/ - loop : List α → List α → List α → List α - | [], r, t => reverseAux t r - | l, [], t => reverseAux t l - | a::l, b::r, t => bif s a b then loop l (b::r) (a::t) else loop (a::l) r (b::t) diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index fd82629a8a..bb7255f4d0 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -161,6 +161,7 @@ theorem exists_of_set' {l : List α} (h : n < l.length) : @[deprecated getElem?_set_eq' (since := "2024-06-12")] theorem get?_set_eq (a : α) (n) (l : List α) : (set l n a).get? n = (fun _ => a) <$> l.get? n := by simp only [get?_eq_getElem?, getElem?_set_eq', Option.map_eq_map] + rfl theorem getElem?_set_eq_of_lt (a : α) {n} {l : List α} (h : n < length l) : (set l n a)[n]? = some a := by rw [getElem?_set_eq', getElem?_eq_getElem h]; rfl @@ -189,25 +190,6 @@ theorem get?_set_of_lt' (a : α) {m n} (l : List α) (h : m < length l) : @[deprecated (since := "2024-05-06")] alias length_removeNth := length_eraseIdx -/-! ### splitAt -/ - -theorem splitAt_go (n : Nat) (l acc : List α) : - splitAt.go l xs n acc = - if n < xs.length then (acc.reverse ++ xs.take n, xs.drop n) else (l, []) := by - induction xs generalizing n acc with - | nil => simp [splitAt.go] - | cons x xs ih => - cases n with - | zero => simp [splitAt.go] - | succ n => - rw [splitAt.go, take_succ_cons, drop_succ_cons, ih n (x :: acc), - reverse_cons, append_assoc, singleton_append, length_cons] - simp only [Nat.succ_lt_succ_iff] - -theorem splitAt_eq (n : Nat) (l : List α) : splitAt n l = (l.take n, l.drop n) := by - rw [splitAt, splitAt_go, reverse_nil, nil_append] - split <;> simp_all [take_of_length_le, drop_of_length_le] - /-! ### eraseP -/ @[simp] theorem extractP_eq_find?_eraseP @@ -606,47 +588,11 @@ theorem insertP_loop (a : α) (l r : List α) : induction l with simp [insertP, insertP.loop, cond] | cons _ _ ih => split <;> simp [insertP_loop, ih] -theorem merge_loop_nil_left (s : α → α → Bool) (r t) : - merge.loop s [] r t = reverseAux t r := by - rw [merge.loop] - /-! ### merge -/ -theorem merge_loop_nil_right (s : α → α → Bool) (l t) : - merge.loop s l [] t = reverseAux t l := by - cases l <;> rw [merge.loop]; intro; contradiction - -theorem merge_loop (s : α → α → Bool) (l r t) : - merge.loop s l r t = reverseAux t (merge s l r) := by - rw [merge]; generalize hn : l.length + r.length = n - induction n using Nat.recAux generalizing l r t with - | zero => - rw [eq_nil_of_length_eq_zero (Nat.eq_zero_of_add_eq_zero_left hn)] - rw [eq_nil_of_length_eq_zero (Nat.eq_zero_of_add_eq_zero_right hn)] - simp only [merge.loop, reverseAux] - | succ n ih => - match l, r with - | [], r => simp only [merge_loop_nil_left]; rfl - | l, [] => simp only [merge_loop_nil_right]; rfl - | a::l, b::r => - simp only [merge.loop, cond] - split - · have hn : l.length + (b :: r).length = n := by - apply Nat.add_right_cancel (m:=1) - rw [←hn]; simp only [length_cons, Nat.add_succ, Nat.succ_add] - rw [ih _ _ (a::t) hn, ih _ _ [] hn, ih _ _ [a] hn]; rfl - · have hn : (a::l).length + r.length = n := by - apply Nat.add_right_cancel (m:=1) - rw [←hn]; simp only [length_cons, Nat.add_succ, Nat.succ_add] - rw [ih _ _ (b::t) hn, ih _ _ [] hn, ih _ _ [b] hn]; rfl - -@[simp] theorem merge_nil (s : α → α → Bool) (l) : merge s l [] = l := merge_loop_nil_right .. - -@[simp] theorem nil_merge (s : α → α → Bool) (r) : merge s [] r = r := merge_loop_nil_left .. - theorem cons_merge_cons (s : α → α → Bool) (a b l r) : - merge s (a::l) (b::r) = if s a b then a :: merge s l (b::r) else b :: merge s (a::l) r := by - simp only [merge, merge.loop, cond]; split <;> (next hs => rw [hs, merge_loop]; rfl) + merge s (a::l) (b::r) = if s a b then a :: merge s l (b::r) else b :: merge s (a::l) r := by + simp only [merge] @[simp] theorem cons_merge_cons_pos (s : α → α → Bool) (l r) (h : s a b) : merge s (a::l) (b::r) = a :: merge s l (b::r) := by @@ -667,17 +613,6 @@ theorem cons_merge_cons (s : α → α → Bool) (a b l r) : · simp_arith [length_merge s l (b::r)] · simp_arith [length_merge s (a::l) r] -@[simp] -theorem mem_merge {s : α → α → Bool} : x ∈ merge s l r ↔ x ∈ l ∨ x ∈ r := by - match l, r with - | l, [] => simp - | [], l => simp - | a::l, b::r => - rw [cons_merge_cons] - split - · simp [mem_merge (l := l) (r := b::r), or_assoc] - · simp [mem_merge (l := a::l) (r := r), or_assoc, or_left_comm] - theorem mem_merge_left (s : α → α → Bool) (h : x ∈ l) : x ∈ merge s l r := mem_merge.2 <| .inl h diff --git a/Batteries/Data/List/Perm.lean b/Batteries/Data/List/Perm.lean index ed2a18f5cf..42c4bbc9c2 100644 --- a/Batteries/Data/List/Perm.lean +++ b/Batteries/Data/List/Perm.lean @@ -23,189 +23,6 @@ namespace List open Perm (swap) -@[simp, refl] protected theorem Perm.refl : ∀ l : List α, l ~ l - | [] => .nil - | x :: xs => (Perm.refl xs).cons x - -protected theorem Perm.rfl {l : List α} : l ~ l := .refl _ - -theorem Perm.of_eq (h : l₁ = l₂) : l₁ ~ l₂ := h ▸ .rfl - -protected theorem Perm.symm {l₁ l₂ : List α} (h : l₁ ~ l₂) : l₂ ~ l₁ := by - induction h with - | nil => exact nil - | cons _ _ ih => exact cons _ ih - | swap => exact swap .. - | trans _ _ ih₁ ih₂ => exact trans ih₂ ih₁ - -theorem perm_comm {l₁ l₂ : List α} : l₁ ~ l₂ ↔ l₂ ~ l₁ := ⟨Perm.symm, Perm.symm⟩ - -theorem Perm.swap' (x y : α) {l₁ l₂ : List α} (p : l₁ ~ l₂) : y :: x :: l₁ ~ x :: y :: l₂ := - (swap ..).trans <| p.cons _ |>.cons _ - -/-- -Similar to `Perm.recOn`, but the `swap` case is generalized to `Perm.swap'`, -where the tail of the lists are not necessarily the same. --/ -@[elab_as_elim] theorem Perm.recOnSwap' - {motive : (l₁ : List α) → (l₂ : List α) → l₁ ~ l₂ → Prop} {l₁ l₂ : List α} (p : l₁ ~ l₂) - (nil : motive [] [] .nil) - (cons : ∀ x {l₁ l₂}, (h : l₁ ~ l₂) → motive l₁ l₂ h → motive (x :: l₁) (x :: l₂) (.cons x h)) - (swap' : ∀ x y {l₁ l₂}, (h : l₁ ~ l₂) → motive l₁ l₂ h → - motive (y :: x :: l₁) (x :: y :: l₂) (.swap' _ _ h)) - (trans : ∀ {l₁ l₂ l₃}, (h₁ : l₁ ~ l₂) → (h₂ : l₂ ~ l₃) → motive l₁ l₂ h₁ → motive l₂ l₃ h₂ → - motive l₁ l₃ (.trans h₁ h₂)) : motive l₁ l₂ p := - have motive_refl l : motive l l (.refl l) := - List.recOn l nil fun x xs ih => cons x (.refl xs) ih - Perm.recOn p nil cons (fun x y l => swap' x y (.refl l) (motive_refl l)) trans - -theorem Perm.eqv (α) : Equivalence (@Perm α) := ⟨.refl, .symm, .trans⟩ - -instance isSetoid (α) : Setoid (List α) := .mk Perm (Perm.eqv α) - -theorem Perm.mem_iff {a : α} {l₁ l₂ : List α} (p : l₁ ~ l₂) : a ∈ l₁ ↔ a ∈ l₂ := by - induction p with - | nil => rfl - | cons _ _ ih => simp only [mem_cons, ih] - | swap => simp only [mem_cons, or_left_comm] - | trans _ _ ih₁ ih₂ => simp only [ih₁, ih₂] - -theorem Perm.subset {l₁ l₂ : List α} (p : l₁ ~ l₂) : l₁ ⊆ l₂ := fun _ => p.mem_iff.mp - -theorem Perm.append_right {l₁ l₂ : List α} (t₁ : List α) (p : l₁ ~ l₂) : l₁ ++ t₁ ~ l₂ ++ t₁ := by - induction p with - | nil => rfl - | cons _ _ ih => exact cons _ ih - | swap => exact swap .. - | trans _ _ ih₁ ih₂ => exact trans ih₁ ih₂ - -theorem Perm.append_left {t₁ t₂ : List α} : ∀ l : List α, t₁ ~ t₂ → l ++ t₁ ~ l ++ t₂ - | [], p => p - | x :: xs, p => (p.append_left xs).cons x - -theorem Perm.append {l₁ l₂ t₁ t₂ : List α} (p₁ : l₁ ~ l₂) (p₂ : t₁ ~ t₂) : l₁ ++ t₁ ~ l₂ ++ t₂ := - (p₁.append_right t₁).trans (p₂.append_left l₂) - -theorem Perm.append_cons (a : α) {h₁ h₂ t₁ t₂ : List α} (p₁ : h₁ ~ h₂) (p₂ : t₁ ~ t₂) : - h₁ ++ a :: t₁ ~ h₂ ++ a :: t₂ := p₁.append (p₂.cons a) - -@[simp] theorem perm_middle {a : α} : ∀ {l₁ l₂ : List α}, l₁ ++ a :: l₂ ~ a :: (l₁ ++ l₂) - | [], _ => .refl _ - | b :: _, _ => (Perm.cons _ perm_middle).trans (swap a b _) - -@[simp] theorem perm_append_singleton (a : α) (l : List α) : l ++ [a] ~ a :: l := - perm_middle.trans <| by rw [append_nil] - -theorem perm_append_comm : ∀ {l₁ l₂ : List α}, l₁ ++ l₂ ~ l₂ ++ l₁ - | [], l₂ => by simp - | a :: t, l₂ => (perm_append_comm.cons _).trans perm_middle.symm - -theorem concat_perm (l : List α) (a : α) : concat l a ~ a :: l := by simp - -theorem Perm.length_eq {l₁ l₂ : List α} (p : l₁ ~ l₂) : length l₁ = length l₂ := by - induction p with - | nil => rfl - | cons _ _ ih => simp only [length_cons, ih] - | swap => rfl - | trans _ _ ih₁ ih₂ => simp only [ih₁, ih₂] - -theorem Perm.eq_nil {l : List α} (p : l ~ []) : l = [] := eq_nil_of_length_eq_zero p.length_eq - -theorem Perm.nil_eq {l : List α} (p : [] ~ l) : [] = l := p.symm.eq_nil.symm - -@[simp] theorem perm_nil {l₁ : List α} : l₁ ~ [] ↔ l₁ = [] := - ⟨fun p => p.eq_nil, fun e => e ▸ .rfl⟩ - -@[simp] theorem nil_perm {l₁ : List α} : [] ~ l₁ ↔ l₁ = [] := perm_comm.trans perm_nil - -theorem not_perm_nil_cons (x : α) (l : List α) : ¬[] ~ x :: l := (nomatch ·.symm.eq_nil) - -@[simp] theorem reverse_perm : ∀ l : List α, reverse l ~ l - | [] => .nil - | a :: l => reverse_cons .. ▸ (perm_append_singleton _ _).trans ((reverse_perm l).cons a) - -theorem perm_cons_append_cons {l l₁ l₂ : List α} (a : α) (p : l ~ l₁ ++ l₂) : - a :: l ~ l₁ ++ a :: l₂ := (p.cons a).trans perm_middle.symm - -@[simp] theorem perm_replicate {n : Nat} {a : α} {l : List α} : - l ~ replicate n a ↔ l = replicate n a := by - refine ⟨fun p => eq_replicate.2 ?_, fun h => h ▸ .rfl⟩ - exact ⟨p.length_eq.trans <| length_replicate .., fun _b m => eq_of_mem_replicate <| p.subset m⟩ - -@[simp] theorem replicate_perm {n : Nat} {a : α} {l : List α} : - replicate n a ~ l ↔ replicate n a = l := (perm_comm.trans perm_replicate).trans eq_comm - -@[simp] theorem perm_singleton {a : α} {l : List α} : l ~ [a] ↔ l = [a] := perm_replicate (n := 1) - -@[simp] theorem singleton_perm {a : α} {l : List α} : [a] ~ l ↔ [a] = l := replicate_perm (n := 1) - -alias ⟨Perm.eq_singleton,_⟩ := perm_singleton -alias ⟨Perm.singleton_eq,_⟩ := singleton_perm - -theorem singleton_perm_singleton {a b : α} : [a] ~ [b] ↔ a = b := by simp - -theorem perm_cons_erase [DecidableEq α] {a : α} {l : List α} (h : a ∈ l) : l ~ a :: l.erase a := - let ⟨_l₁, _l₂, _, e₁, e₂⟩ := exists_erase_eq h - e₂ ▸ e₁ ▸ perm_middle - -theorem Perm.filterMap (f : α → Option β) {l₁ l₂ : List α} (p : l₁ ~ l₂) : - filterMap f l₁ ~ filterMap f l₂ := by - induction p with - | nil => simp - | cons x _p IH => cases h : f x <;> simp [h, filterMap_cons, IH, Perm.cons] - | swap x y l₂ => cases hx : f x <;> cases hy : f y <;> simp [hx, hy, filterMap_cons, swap] - | trans _p₁ _p₂ IH₁ IH₂ => exact IH₁.trans IH₂ - -theorem Perm.map (f : α → β) {l₁ l₂ : List α} (p : l₁ ~ l₂) : map f l₁ ~ map f l₂ := - filterMap_eq_map f ▸ p.filterMap _ - -theorem Perm.pmap {p : α → Prop} (f : ∀ a, p a → β) {l₁ l₂ : List α} (p : l₁ ~ l₂) {H₁ H₂} : - pmap f l₁ H₁ ~ pmap f l₂ H₂ := by - induction p with - | nil => simp - | cons x _p IH => simp [IH, Perm.cons] - | swap x y => simp [swap] - | trans _p₁ p₂ IH₁ IH₂ => exact IH₁.trans (IH₂ (H₁ := fun a m => H₂ a (p₂.subset m))) - -theorem Perm.filter (p : α → Bool) {l₁ l₂ : List α} (s : l₁ ~ l₂) : - filter p l₁ ~ filter p l₂ := by rw [← filterMap_eq_filter]; apply s.filterMap - -theorem filter_append_perm (p : α → Bool) (l : List α) : - filter p l ++ filter (fun x => !p x) l ~ l := by - induction l with - | nil => rfl - | cons x l ih => - by_cases h : p x <;> simp [h] - · exact ih.cons x - · exact Perm.trans (perm_append_comm.trans (perm_append_comm.cons _)) (ih.cons x) - -theorem exists_perm_sublist {l₁ l₂ l₂' : List α} (s : l₁ <+ l₂) (p : l₂ ~ l₂') : - ∃ l₁', l₁' ~ l₁ ∧ l₁' <+ l₂' := by - induction p generalizing l₁ with - | nil => exact ⟨[], sublist_nil.mp s ▸ .rfl, nil_sublist _⟩ - | cons x _ IH => - match s with - | .cons _ s => let ⟨l₁', p', s'⟩ := IH s; exact ⟨l₁', p', s'.cons _⟩ - | .cons₂ _ s => let ⟨l₁', p', s'⟩ := IH s; exact ⟨x :: l₁', p'.cons x, s'.cons₂ _⟩ - | swap x y l' => - match s with - | .cons _ (.cons _ s) => exact ⟨_, .rfl, (s.cons _).cons _⟩ - | .cons _ (.cons₂ _ s) => exact ⟨x :: _, .rfl, (s.cons _).cons₂ _⟩ - | .cons₂ _ (.cons _ s) => exact ⟨y :: _, .rfl, (s.cons₂ _).cons _⟩ - | .cons₂ _ (.cons₂ _ s) => exact ⟨x :: y :: _, .swap .., (s.cons₂ _).cons₂ _⟩ - | trans _ _ IH₁ IH₂ => - let ⟨m₁, pm, sm⟩ := IH₁ s - let ⟨r₁, pr, sr⟩ := IH₂ sm - exact ⟨r₁, pr.trans pm, sr⟩ - -theorem Perm.sizeOf_eq_sizeOf [SizeOf α] {l₁ l₂ : List α} (h : l₁ ~ l₂) : - sizeOf l₁ = sizeOf l₂ := by - induction h with - | nil => rfl - | cons _ _ h_sz₁₂ => simp [h_sz₁₂] - | swap => simp [Nat.add_left_comm] - | trans _ _ h_sz₁₂ h_sz₂₃ => simp [h_sz₁₂, h_sz₂₃] - section Subperm theorem nil_subperm {l : List α} : [] <+~ l := ⟨[], Perm.nil, by simp⟩ @@ -256,115 +73,12 @@ theorem Subperm.filter (p : α → Bool) ⦃l l' : List α⦄ (h : l <+~ l') : end Subperm -theorem Sublist.exists_perm_append {l₁ l₂ : List α} : l₁ <+ l₂ → ∃ l, l₂ ~ l₁ ++ l - | Sublist.slnil => ⟨nil, .rfl⟩ - | Sublist.cons a s => - let ⟨l, p⟩ := Sublist.exists_perm_append s - ⟨a :: l, (p.cons a).trans perm_middle.symm⟩ - | Sublist.cons₂ a s => - let ⟨l, p⟩ := Sublist.exists_perm_append s - ⟨l, p.cons a⟩ - -theorem Perm.countP_eq (p : α → Bool) {l₁ l₂ : List α} (s : l₁ ~ l₂) : - countP p l₁ = countP p l₂ := by - simp only [countP_eq_length_filter] - exact (s.filter _).length_eq - theorem Subperm.countP_le (p : α → Bool) {l₁ l₂ : List α} : l₁ <+~ l₂ → countP p l₁ ≤ countP p l₂ | ⟨_l, p', s⟩ => p'.countP_eq p ▸ s.countP_le p -theorem Perm.countP_congr {l₁ l₂ : List α} (s : l₁ ~ l₂) {p p' : α → Bool} - (hp : ∀ x ∈ l₁, p x = p' x) : l₁.countP p = l₂.countP p' := by - rw [← s.countP_eq p'] - clear s - induction l₁ with - | nil => rfl - | cons y s hs => - simp only [mem_cons, forall_eq_or_imp] at hp - simp only [countP_cons, hs hp.2, hp.1] - -theorem countP_eq_countP_filter_add (l : List α) (p q : α → Bool) : - l.countP p = (l.filter q).countP p + (l.filter fun a => !q a).countP p := - countP_append .. ▸ Perm.countP_eq _ (filter_append_perm _ _).symm - -theorem Perm.count_eq [DecidableEq α] {l₁ l₂ : List α} (p : l₁ ~ l₂) (a) : - count a l₁ = count a l₂ := p.countP_eq _ - theorem Subperm.count_le [DecidableEq α] {l₁ l₂ : List α} (s : l₁ <+~ l₂) (a) : count a l₁ ≤ count a l₂ := s.countP_le _ -theorem Perm.foldl_eq' {f : β → α → β} {l₁ l₂ : List α} (p : l₁ ~ l₂) - (comm : ∀ x ∈ l₁, ∀ y ∈ l₁, ∀ (z), f (f z x) y = f (f z y) x) - (init) : foldl f init l₁ = foldl f init l₂ := by - induction p using recOnSwap' generalizing init with - | nil => simp - | cons x _p IH => - simp only [foldl] - apply IH; intros; apply comm <;> exact .tail _ ‹_› - | swap' x y _p IH => - simp only [foldl] - rw [comm x (.tail _ <| .head _) y (.head _)] - apply IH; intros; apply comm <;> exact .tail _ (.tail _ ‹_›) - | trans p₁ _p₂ IH₁ IH₂ => - refine (IH₁ comm init).trans (IH₂ ?_ _) - intros; apply comm <;> apply p₁.symm.subset <;> assumption - -theorem Perm.rec_heq {β : List α → Sort _} {f : ∀ a l, β l → β (a :: l)} {b : β []} {l l' : List α} - (hl : l ~ l') (f_congr : ∀ {a l l' b b'}, l ~ l' → HEq b b' → HEq (f a l b) (f a l' b')) - (f_swap : ∀ {a a' l b}, HEq (f a (a' :: l) (f a' l b)) (f a' (a :: l) (f a l b))) : - HEq (@List.rec α β b f l) (@List.rec α β b f l') := by - induction hl with - | nil => rfl - | cons a h ih => exact f_congr h ih - | swap a a' l => exact f_swap - | trans _h₁ _h₂ ih₁ ih₂ => exact ih₁.trans ih₂ - -/-- Lemma used to destruct perms element by element. -/ -theorem perm_inv_core {a : α} {l₁ l₂ r₁ r₂ : List α} : - l₁ ++ a :: r₁ ~ l₂ ++ a :: r₂ → l₁ ++ r₁ ~ l₂ ++ r₂ := by - -- Necessary generalization for `induction` - suffices ∀ s₁ s₂ (_ : s₁ ~ s₂) {l₁ l₂ r₁ r₂}, - l₁ ++ a :: r₁ = s₁ → l₂ ++ a :: r₂ = s₂ → l₁ ++ r₁ ~ l₂ ++ r₂ from (this _ _ · rfl rfl) - intro s₁ s₂ p - induction p using Perm.recOnSwap' with intro l₁ l₂ r₁ r₂ e₁ e₂ - | nil => - simp at e₁ - | cons x p IH => - cases l₁ <;> cases l₂ <;> - dsimp at e₁ e₂ <;> injections <;> subst_vars - · exact p - · exact p.trans perm_middle - · exact perm_middle.symm.trans p - · exact (IH rfl rfl).cons _ - | swap' x y p IH => - obtain _ | ⟨y, _ | ⟨z, l₁⟩⟩ := l₁ - <;> obtain _ | ⟨u, _ | ⟨v, l₂⟩⟩ := l₂ - <;> dsimp at e₁ e₂ <;> injections <;> subst_vars - <;> try exact p.cons _ - · exact (p.trans perm_middle).cons u - · exact ((p.trans perm_middle).cons _).trans (swap _ _ _) - · exact (perm_middle.symm.trans p).cons y - · exact (swap _ _ _).trans ((perm_middle.symm.trans p).cons u) - · exact (IH rfl rfl).swap' _ _ - | trans p₁ p₂ IH₁ IH₂ => - subst e₁ e₂ - obtain ⟨l₂, r₂, rfl⟩ := append_of_mem (a := a) (p₁.subset (by simp)) - exact (IH₁ rfl rfl).trans (IH₂ rfl rfl) - -theorem Perm.cons_inv {a : α} {l₁ l₂ : List α} : a :: l₁ ~ a :: l₂ → l₁ ~ l₂ := - perm_inv_core (l₁ := []) (l₂ := []) - -@[simp] theorem perm_cons (a : α) {l₁ l₂ : List α} : a :: l₁ ~ a :: l₂ ↔ l₁ ~ l₂ := - ⟨.cons_inv, .cons a⟩ - -theorem perm_append_left_iff {l₁ l₂ : List α} : ∀ l, l ++ l₁ ~ l ++ l₂ ↔ l₁ ~ l₂ - | [] => .rfl - | a :: l => (perm_cons a).trans (perm_append_left_iff l) - -theorem perm_append_right_iff {l₁ l₂ : List α} (l) : l₁ ++ l ~ l₂ ++ l ↔ l₁ ~ l₂ := by - refine ⟨fun p => ?_, .append_right _⟩ - exact (perm_append_left_iff _).1 <| perm_append_comm.trans <| p.trans perm_append_comm - theorem subperm_cons (a : α) {l₁ l₂ : List α} : a :: l₁ <+~ a :: l₂ ↔ l₁ <+~ l₂ := by refine ⟨fun ⟨l, p, s⟩ => ?_, fun ⟨l, p, s⟩ => ⟨a :: l, p.cons a, s.cons₂ _⟩⟩ match s with @@ -430,7 +144,7 @@ theorem perm_ext_iff_of_nodup {l₁ l₂ : List α} (d₁ : Nodup l₁) (d₂ : theorem Nodup.perm_iff_eq_of_sublist {l₁ l₂ l : List α} (d : Nodup l) (s₁ : l₁ <+ l) (s₂ : l₂ <+ l) : l₁ ~ l₂ ↔ l₁ = l₂ := by refine ⟨fun h => ?_, fun h => by rw [h]⟩ - induction s₂ generalizing l₁ with simp [Nodup] at d + induction s₂ generalizing l₁ with simp [Nodup, List.forall_mem_ne] at d | slnil => exact h.eq_nil | cons a s₂ IH => match s₁ with @@ -449,14 +163,6 @@ section DecidableEq variable [DecidableEq α] -theorem Perm.erase (a : α) {l₁ l₂ : List α} (p : l₁ ~ l₂) : l₁.erase a ~ l₂.erase a := - if h₁ : a ∈ l₁ then - have h₂ : a ∈ l₂ := p.subset h₁ - .cons_inv <| (perm_cons_erase h₁).symm.trans <| p.trans (perm_cons_erase h₂) - else by - have h₂ : a ∉ l₂ := mt p.mem_iff.2 h₁ - rw [erase_of_not_mem h₁, erase_of_not_mem h₂]; exact p - theorem subperm_cons_erase (a : α) (l : List α) : l <+~ a :: l.erase a := if h : a ∈ l then (perm_cons_erase h).subperm @@ -509,7 +215,7 @@ theorem erase_cons_subperm_cons_erase (a b : α) (l : List α) : if h : a = b then rw [h, erase_cons_head]; apply subperm_cons_erase else - have : ¬(a == b) = true := by simp only [beq_false_of_ne h, not_false_eq_true] + have : ¬(a == b) = true := by simp only [beq_false_of_ne h, not_false_eq_true, reduceCtorEq] rw [erase_cons_tail this] theorem subperm_cons_diff {a : α} {l₁ l₂ : List α} : (a :: l₁).diff l₂ <+~ a :: l₁.diff l₂ := by @@ -522,28 +228,6 @@ theorem subperm_cons_diff {a : α} {l₁ l₂ : List α} : (a :: l₁).diff l₂ theorem subset_cons_diff {a : α} {l₁ l₂ : List α} : (a :: l₁).diff l₂ ⊆ a :: l₁.diff l₂ := subperm_cons_diff.subset -theorem cons_perm_iff_perm_erase {a : α} {l₁ l₂ : List α} : - a :: l₁ ~ l₂ ↔ a ∈ l₂ ∧ l₁ ~ l₂.erase a := by - refine ⟨fun h => ?_, fun ⟨m, h⟩ => (h.cons a).trans (perm_cons_erase m).symm⟩ - have : a ∈ l₂ := h.subset (mem_cons_self a l₁) - exact ⟨this, (h.trans <| perm_cons_erase this).cons_inv⟩ - -theorem perm_iff_count {l₁ l₂ : List α} : l₁ ~ l₂ ↔ ∀ a, count a l₁ = count a l₂ := by - refine ⟨Perm.count_eq, fun H => ?_⟩ - induction l₁ generalizing l₂ with - | nil => - match l₂ with - | nil => rfl - | cons b l₂ => - specialize H b - simp at H - | cons a l₁ IH => - have : a ∈ l₂ := count_pos_iff_mem.mp (by rw [← H]; simp) - refine ((IH fun b => ?_).cons a).trans (perm_cons_erase this).symm - specialize H b - rw [(perm_cons_erase this).count_eq] at H - by_cases h : b = a <;> simpa [h] using H - /-- The list version of `add_tsub_cancel_of_le` for multisets. -/ theorem subperm_append_diff_self_of_count_le {l₁ l₂ : List α} (h : ∀ x ∈ l₁, count x l₁ ≤ count x l₂) : l₁ ++ l₂.diff l₁ ~ l₂ := by @@ -586,28 +270,6 @@ theorem Subperm.cons_left {l₁ l₂ : List α} (h : l₁ <+~ l₂) (x : α) (hx refine h y ?_ simpa [hy'] using hy -theorem isPerm_iff : ∀ {l₁ l₂ : List α}, l₁.isPerm l₂ ↔ l₁ ~ l₂ - | [], [] => by simp [isPerm, isEmpty] - | [], _ :: _ => by simp [isPerm, isEmpty, Perm.nil_eq] - | a :: l₁, l₂ => by simp [isPerm, isPerm_iff, cons_perm_iff_perm_erase] - -instance decidablePerm (l₁ l₂ : List α) : Decidable (l₁ ~ l₂) := decidable_of_iff _ isPerm_iff - -protected theorem Perm.insert (a : α) {l₁ l₂ : List α} (p : l₁ ~ l₂) : - l₁.insert a ~ l₂.insert a := by - if h : a ∈ l₁ then - simp [h, p.subset h, p] - else - have := p.cons a - simpa [h, mt p.mem_iff.2 h] using this - -theorem perm_insert_swap (x y : α) (l : List α) : - List.insert x (List.insert y l) ~ List.insert y (List.insert x l) := by - by_cases xl : x ∈ l <;> by_cases yl : y ∈ l <;> simp [xl, yl] - if xy : x = y then simp [xy] else - simp [List.insert, xl, yl, xy, Ne.symm xy] - constructor - theorem perm_insertNth {α} (x : α) (l : List α) {n} (h : n ≤ l.length) : insertNth n x l ~ x :: l := by induction l generalizing n with @@ -647,59 +309,11 @@ theorem Perm.inter {l₁ l₂ t₁ t₂ : List α} (p₁ : l₁ ~ l₂) (p₂ : end DecidableEq -theorem Perm.pairwise_iff {R : α → α → Prop} (S : ∀ {x y}, R x y → R y x) : - ∀ {l₁ l₂ : List α} (_p : l₁ ~ l₂), Pairwise R l₁ ↔ Pairwise R l₂ := - suffices ∀ {l₁ l₂}, l₁ ~ l₂ → Pairwise R l₁ → Pairwise R l₂ - from fun p => ⟨this p, this p.symm⟩ - fun {l₁ l₂} p d => by - induction d generalizing l₂ with - | nil => rw [← p.nil_eq]; constructor - | cons h _ IH => - have : _ ∈ l₂ := p.subset (mem_cons_self _ _) - obtain ⟨s₂, t₂, rfl⟩ := append_of_mem this - have p' := (p.trans perm_middle).cons_inv - refine (pairwise_middle S).2 (pairwise_cons.2 ⟨fun b m => ?_, IH p'⟩) - exact h _ (p'.symm.subset m) - -theorem Pairwise.perm {R : α → α → Prop} {l l' : List α} (hR : l.Pairwise R) (hl : l ~ l') - (hsymm : ∀ {x y}, R x y → R y x) : l'.Pairwise R := (hl.pairwise_iff hsymm).mp hR - -theorem Perm.pairwise {R : α → α → Prop} {l l' : List α} (hl : l ~ l') (hR : l.Pairwise R) - (hsymm : ∀ {x y}, R x y → R y x) : l'.Pairwise R := hR.perm hl hsymm - -theorem Perm.nodup_iff {l₁ l₂ : List α} : l₁ ~ l₂ → (Nodup l₁ ↔ Nodup l₂) := - Perm.pairwise_iff <| @Ne.symm α - -theorem Perm.join {l₁ l₂ : List (List α)} (h : l₁ ~ l₂) : l₁.join ~ l₂.join := by - induction h with - | nil => rfl - | cons _ _ ih => simp only [join_cons, perm_append_left_iff, ih] - | swap => simp only [join_cons, ← append_assoc, perm_append_right_iff]; exact perm_append_comm .. - | trans _ _ ih₁ ih₂ => exact trans ih₁ ih₂ - -theorem Perm.bind_right {l₁ l₂ : List α} (f : α → List β) (p : l₁ ~ l₂) : l₁.bind f ~ l₂.bind f := - (p.map _).join - theorem Perm.join_congr : ∀ {l₁ l₂ : List (List α)} (_ : List.Forall₂ (· ~ ·) l₁ l₂), l₁.join ~ l₂.join | _, _, .nil => .rfl | _ :: _, _ :: _, .cons h₁ h₂ => h₁.append (Perm.join_congr h₂) -theorem Perm.eraseP (f : α → Bool) {l₁ l₂ : List α} - (H : Pairwise (fun a b => f a → f b → False) l₁) (p : l₁ ~ l₂) : eraseP f l₁ ~ eraseP f l₂ := by - induction p with - | nil => simp - | cons a p IH => - if h : f a then simp [h, p] - else simp [h]; exact IH (pairwise_cons.1 H).2 - | swap a b l => - by_cases h₁ : f a <;> by_cases h₂ : f b <;> simp [h₁, h₂] - · cases (pairwise_cons.1 H).1 _ (mem_cons.2 (Or.inl rfl)) h₂ h₁ - · apply swap - | trans p₁ _ IH₁ IH₂ => - refine (IH₁ H).trans (IH₂ ((p₁.pairwise_iff ?_).1 H)) - exact fun h h₁ h₂ => h h₂ h₁ - theorem perm_insertP (p : α → Bool) (a l) : insertP p a l ~ a :: l := by induction l with simp [insertP, insertP.loop, cond] | cons _ _ ih => diff --git a/Batteries/Data/Nat/Basic.lean b/Batteries/Data/Nat/Basic.lean index 1d9658f49e..21b4a4dd43 100644 --- a/Batteries/Data/Nat/Basic.lean +++ b/Batteries/Data/Nat/Basic.lean @@ -21,13 +21,6 @@ protected def recAuxOn {motive : Nat → Sort _} (t : Nat) (zero : motive 0) protected def strongRec {motive : Nat → Sort _} (ind : ∀ n, (∀ m, m < n → motive m) → motive n) (t : Nat) : motive t := ind t fun m _ => Nat.strongRec ind m -/-- - Strong recursor for `Nat` --/ -@[elab_as_elim] -protected def strongRecOn (t : Nat) {motive : Nat → Sort _} - (ind : ∀ n, (∀ m, m < n → motive m) → motive n) : motive t := Nat.strongRec ind t - /-- Strong recursor via a `Nat`-valued measure -/ diff --git a/Batteries/Data/Nat/Lemmas.lean b/Batteries/Data/Nat/Lemmas.lean index 1c3edddeaf..3492037135 100644 --- a/Batteries/Data/Nat/Lemmas.lean +++ b/Batteries/Data/Nat/Lemmas.lean @@ -47,7 +47,7 @@ theorem strongRec_eq {motive : Nat → Sort _} (ind : ∀ n, (∀ m, m < n → m theorem strongRecOn_eq {motive : Nat → Sort _} (ind : ∀ n, (∀ m, m < n → motive m) → motive n) (t : Nat) : Nat.strongRecOn t ind = ind t fun m _ => Nat.strongRecOn m ind := - Nat.strongRec_eq .. + WellFounded.fix_eq WellFoundedRelation.wf ind t @[simp] theorem recDiagAux_zero_left {motive : Nat → Nat → Sort _} (zero_left : ∀ n, motive 0 n) (zero_right : ∀ m, motive m 0) diff --git a/Batteries/Data/RBMap/Basic.lean b/Batteries/Data/RBMap/Basic.lean index d1aff2dc3b..8eb202591c 100644 --- a/Batteries/Data/RBMap/Basic.lean +++ b/Batteries/Data/RBMap/Basic.lean @@ -194,7 +194,8 @@ instance {t : RBNode α} [DecidablePred p] : Decidable (t.Any p) := /-- True if `x` is an element of `t` "exactly", i.e. up to equality, not the `cmp` relation. -/ def EMem (x : α) (t : RBNode α) : Prop := t.Any (x = ·) -instance : Membership α (RBNode α) := ⟨EMem⟩ +instance : Membership α (RBNode α) where + mem t x := EMem x t /-- True if the specified `cut` matches at least one element of of `t`. -/ def MemP (cut : α → Ordering) (t : RBNode α) : Prop := t.Any (cut · = .eq) @@ -768,7 +769,8 @@ def MemP (cut : α → Ordering) (t : RBSet α cmp) : Prop := t.1.MemP cut /-- True if `x` is equivalent to an element of `t`. -/ def Mem (x : α) (t : RBSet α cmp) : Prop := MemP (cmp x) t -instance : Membership α (RBSet α cmp) := ⟨Mem⟩ +instance : Membership α (RBSet α cmp) where + mem t x := Mem x t -- These instances are put in a special namespace because they are usually not what users want -- when deciding membership in a RBSet, since this does a naive linear search through the tree. diff --git a/Batteries/Data/RBMap/Lemmas.lean b/Batteries/Data/RBMap/Lemmas.lean index 43d31af8f7..df9ec1d585 100644 --- a/Batteries/Data/RBMap/Lemmas.lean +++ b/Batteries/Data/RBMap/Lemmas.lean @@ -25,9 +25,9 @@ attribute [simp] fold foldl foldr Any forM foldlM Ordered @[simp] theorem max?_reverse (t : RBNode α) : t.reverse.max? = t.min? := by rw [← min?_reverse, reverse_reverse] -@[simp] theorem mem_nil {x} : ¬x ∈ (.nil : RBNode α) := by simp [(·∈·), EMem] +@[simp] theorem mem_nil {x} : ¬x ∈ (.nil : RBNode α) := by simp [Membership.mem, EMem] @[simp] theorem mem_node {y c a x b} : - y ∈ (.node c a x b : RBNode α) ↔ y = x ∨ y ∈ a ∨ y ∈ b := by simp [(·∈·), EMem] + y ∈ (.node c a x b : RBNode α) ↔ y = x ∨ y ∈ a ∨ y ∈ b := by simp [Membership.mem, EMem] theorem All_def {t : RBNode α} : t.All p ↔ ∀ x ∈ t, p x := by induction t <;> simp [or_imp, forall_and, *] @@ -1185,7 +1185,9 @@ theorem contains_iff_findEntry? {t : RBMap α β cmp} : theorem contains_iff_find? {t : RBMap α β cmp} : t.contains x ↔ ∃ v, t.find? x = some v := by - simp [contains_iff_findEntry?, find?, and_comm, exists_comm] + simp only [contains_iff_findEntry?, Prod.exists, find?, Option.map_eq_some', and_comm, + exists_eq_left] + rw [exists_comm] theorem size_eq (t : RBMap α β cmp) : t.size = t.toList.length := RBNode.size_eq diff --git a/Batteries/Data/String/Lemmas.lean b/Batteries/Data/String/Lemmas.lean index a8d2bfa7b4..7789f9dcd2 100644 --- a/Batteries/Data/String/Lemmas.lean +++ b/Batteries/Data/String/Lemmas.lean @@ -57,7 +57,7 @@ private theorem ne_self_add_add_utf8Size : i ≠ i + (n + Char.utf8Size c) := @[simp] theorem utf8Len_append (cs₁ cs₂) : utf8Len (cs₁ ++ cs₂) = utf8Len cs₁ + utf8Len cs₂ := by induction cs₁ <;> simp [*, Nat.add_right_comm] -@[simp] theorem utf8Len_reverseAux (cs₁ cs₂) : +theorem utf8Len_reverseAux (cs₁ cs₂) : utf8Len (cs₁.reverseAux cs₂) = utf8Len cs₁ + utf8Len cs₂ := by induction cs₁ generalizing cs₂ <;> simp_all [← Nat.add_assoc, Nat.add_right_comm] diff --git a/Batteries/Data/Sum/Lemmas.lean b/Batteries/Data/Sum/Lemmas.lean index 9d55bd1919..ffc24dcd76 100644 --- a/Batteries/Data/Sum/Lemmas.lean +++ b/Batteries/Data/Sum/Lemmas.lean @@ -45,10 +45,10 @@ section get | inr _, _ => rfl @[simp] theorem getLeft?_eq_none_iff {x : α ⊕ β} : x.getLeft? = none ↔ x.isRight := by - cases x <;> simp only [getLeft?, isRight, eq_self_iff_true] + cases x <;> simp only [getLeft?, isRight, eq_self_iff_true, reduceCtorEq] @[simp] theorem getRight?_eq_none_iff {x : α ⊕ β} : x.getRight? = none ↔ x.isLeft := by - cases x <;> simp only [getRight?, isLeft, eq_self_iff_true] + cases x <;> simp only [getRight?, isLeft, eq_self_iff_true, reduceCtorEq] theorem eq_left_getLeft_of_isLeft : ∀ {x : α ⊕ β} (h : x.isLeft), x = inl (x.getLeft h) | inl _, _ => rfl @@ -63,10 +63,10 @@ theorem eq_right_getRight_of_isRight : ∀ {x : α ⊕ β} (h : x.isRight), x = cases x <;> simp at h ⊢ @[simp] theorem getLeft?_eq_some_iff : x.getLeft? = some a ↔ x = inl a := by - cases x <;> simp only [getLeft?, Option.some.injEq, inl.injEq] + cases x <;> simp only [getLeft?, Option.some.injEq, inl.injEq, reduceCtorEq] @[simp] theorem getRight?_eq_some_iff : x.getRight? = some b ↔ x = inr b := by - cases x <;> simp only [getRight?, Option.some.injEq, inr.injEq] + cases x <;> simp only [getRight?, Option.some.injEq, inr.injEq, reduceCtorEq] @[simp] theorem bnot_isLeft (x : α ⊕ β) : !x.isLeft = x.isRight := by cases x <;> rfl diff --git a/Batteries/Data/Thunk.lean b/Batteries/Data/Thunk.lean deleted file mode 100644 index 88c92c3784..0000000000 --- a/Batteries/Data/Thunk.lean +++ /dev/null @@ -1,10 +0,0 @@ -/- -Copyright (c) 2024 François G. Dorais. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE -Authors: François G. Dorais, et al. --/ - -namespace Thunk - -@[ext] protected theorem ext : {a b : Thunk α} → a.get = b.get → a = b - | {..}, {..}, heq => congrArg _ <| funext fun _ => heq diff --git a/Batteries/Data/UInt.lean b/Batteries/Data/UInt.lean index 44ebfd9b80..a7c2a4cbaa 100644 --- a/Batteries/Data/UInt.lean +++ b/Batteries/Data/UInt.lean @@ -26,8 +26,6 @@ theorem UInt8.toNat_lt (x : UInt8) : x.toNat < 2 ^ 8 := x.val.isLt @[simp] theorem UInt8.toUInt64_toNat (x : UInt8) : x.toUInt64.toNat = x.toNat := rfl -theorem UInt8.toNat_zero : (0 : UInt8).toNat = 0 := rfl - theorem UInt8.toNat_add (x y : UInt8) : (x + y).toNat = (x.toNat + y.toNat) % UInt8.size := rfl theorem UInt8.toNat_sub (x y : UInt8) : @@ -35,12 +33,6 @@ theorem UInt8.toNat_sub (x y : UInt8) : theorem UInt8.toNat_mul (x y : UInt8) : (x * y).toNat = (x.toNat * y.toNat) % UInt8.size := rfl -theorem UInt8.toNat_div (x y : UInt8) : (x / y).toNat = x.toNat / y.toNat := rfl - -theorem UInt8.toNat_mod (x y : UInt8) : (x % y).toNat = x.toNat % y.toNat := rfl - -theorem UInt8.toNat_modn (x : UInt8) (n) : (x.modn n).toNat = x.toNat % n := rfl - theorem UInt8.le_antisymm_iff {x y : UInt8} : x = y ↔ x ≤ y ∧ y ≤ x := UInt8.ext_iff.trans Nat.le_antisymm_iff @@ -71,8 +63,6 @@ theorem UInt16.toNat_lt (x : UInt16) : x.toNat < 2 ^ 16 := x.val.isLt @[simp] theorem UInt16.toUInt64_toNat (x : UInt16) : x.toUInt64.toNat = x.toNat := rfl -theorem UInt16.toNat_zero : (0 : UInt16).toNat = 0 := rfl - theorem UInt16.toNat_add (x y : UInt16) : (x + y).toNat = (x.toNat + y.toNat) % UInt16.size := rfl theorem UInt16.toNat_sub (x y : UInt16) : @@ -80,12 +70,6 @@ theorem UInt16.toNat_sub (x y : UInt16) : theorem UInt16.toNat_mul (x y : UInt16) : (x * y).toNat = (x.toNat * y.toNat) % UInt16.size := rfl -theorem UInt16.toNat_div (x y : UInt16) : (x / y).toNat = x.toNat / y.toNat := rfl - -theorem UInt16.toNat_mod (x y : UInt16) : (x % y).toNat = x.toNat % y.toNat := rfl - -theorem UInt16.toNat_modn (x : UInt16) (n) : (x.modn n).toNat = x.toNat % n := rfl - theorem UInt16.le_antisymm_iff {x y : UInt16} : x = y ↔ x ≤ y ∧ y ≤ x := UInt16.ext_iff.trans Nat.le_antisymm_iff @@ -116,8 +100,6 @@ theorem UInt32.toNat_lt (x : UInt32) : x.toNat < 2 ^ 32 := x.val.isLt @[simp] theorem UInt32.toUInt64_toNat (x : UInt32) : x.toUInt64.toNat = x.toNat := rfl -theorem UInt32.toNat_zero : (0 : UInt32).toNat = 0 := rfl - theorem UInt32.toNat_add (x y : UInt32) : (x + y).toNat = (x.toNat + y.toNat) % UInt32.size := rfl theorem UInt32.toNat_sub (x y : UInt32) : @@ -125,12 +107,6 @@ theorem UInt32.toNat_sub (x y : UInt32) : theorem UInt32.toNat_mul (x y : UInt32) : (x * y).toNat = (x.toNat * y.toNat) % UInt32.size := rfl -theorem UInt32.toNat_div (x y : UInt32) : (x / y).toNat = x.toNat / y.toNat := rfl - -theorem UInt32.toNat_mod (x y : UInt32) : (x % y).toNat = x.toNat % y.toNat := rfl - -theorem UInt32.toNat_modn (x : UInt32) (n) : (x.modn n).toNat = x.toNat % n := rfl - theorem UInt32.le_antisymm_iff {x y : UInt32} : x = y ↔ x ≤ y ∧ y ≤ x := UInt32.ext_iff.trans Nat.le_antisymm_iff @@ -161,8 +137,6 @@ theorem UInt64.toNat_lt (x : UInt64) : x.toNat < 2 ^ 64 := x.val.isLt @[simp] theorem UInt64.toUInt32_toNat (x : UInt64) : x.toUInt32.toNat = x.toNat % 2 ^ 32 := rfl -theorem UInt64.toNat_zero : (0 : UInt64).toNat = 0 := rfl - theorem UInt64.toNat_add (x y : UInt64) : (x + y).toNat = (x.toNat + y.toNat) % UInt64.size := rfl theorem UInt64.toNat_sub (x y : UInt64) : @@ -170,12 +144,6 @@ theorem UInt64.toNat_sub (x y : UInt64) : theorem UInt64.toNat_mul (x y : UInt64) : (x * y).toNat = (x.toNat * y.toNat) % UInt64.size := rfl -theorem UInt64.toNat_div (x y : UInt64) : (x / y).toNat = x.toNat / y.toNat := rfl - -theorem UInt64.toNat_mod (x y : UInt64) : (x % y).toNat = x.toNat % y.toNat := rfl - -theorem UInt64.toNat_modn (x : UInt64) (n) : (x.modn n).toNat = x.toNat % n := rfl - theorem UInt64.le_antisymm_iff {x y : UInt64} : x = y ↔ x ≤ y ∧ y ≤ x := UInt64.ext_iff.trans Nat.le_antisymm_iff @@ -220,8 +188,6 @@ theorem USize.toNat_lt (x : USize) : x.toNat < 2 ^ System.Platform.numBits := by @[simp] theorem UInt32.toUSize_toNat (x : UInt32) : x.toUSize.toNat = x.toNat := rfl -theorem USize.toNat_zero : (0 : USize).toNat = 0 := rfl - theorem USize.toNat_add (x y : USize) : (x + y).toNat = (x.toNat + y.toNat) % USize.size := rfl theorem USize.toNat_sub (x y : USize) : @@ -229,12 +195,6 @@ theorem USize.toNat_sub (x y : USize) : theorem USize.toNat_mul (x y : USize) : (x * y).toNat = (x.toNat * y.toNat) % USize.size := rfl -theorem USize.toNat_div (x y : USize) : (x / y).toNat = x.toNat / y.toNat := rfl - -theorem USize.toNat_mod (x y : USize) : (x % y).toNat = x.toNat % y.toNat := rfl - -theorem USize.toNat_modn (x : USize) (n) : (x.modn n).toNat = x.toNat % n := rfl - theorem USize.le_antisymm_iff {x y : USize} : x = y ↔ x ≤ y ∧ y ≤ x := USize.ext_iff.trans Nat.le_antisymm_iff diff --git a/Batteries/Data/UnionFind/Basic.lean b/Batteries/Data/UnionFind/Basic.lean index e0b4174112..549c9e1cde 100644 --- a/Batteries/Data/UnionFind/Basic.lean +++ b/Batteries/Data/UnionFind/Basic.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro -/ import Batteries.Tactic.Lint.Misc import Batteries.Tactic.SeqFocus +import Batteries.Data.Array.Lemmas namespace Batteries @@ -301,19 +302,22 @@ theorem findAux_s {self : UnionFind} {x : Fin self.size} : apply dif_pos exact parent'_lt .. +set_option linter.deprecated false in theorem rankD_findAux {self : UnionFind} {x : Fin self.size} : rankD (findAux self x).s i = self.rank i := by if h : i < self.size then rw [findAux_s]; split <;> [rfl; skip] have := Nat.sub_lt_sub_left (self.lt_rankMax x) (self.rank'_lt _ ‹_›) have := lt_of_parentD (by rwa [parentD_eq]) - rw [rankD_eq' (by simp [FindAux.size_eq, h]), Array.get_modify (by rwa [FindAux.size_eq])] + rw [rankD_eq' (by simp [FindAux.size_eq, h])] + rw [Array.get_modify (by rwa [FindAux.size_eq])] split <;> simp [← rankD_eq, rankD_findAux (x := ⟨_, self.parent'_lt x⟩), -Array.get_eq_getElem] else simp only [rankD, Array.data_length, Array.get_eq_getElem, rank] rw [dif_neg (by rwa [FindAux.size_eq]), dif_neg h] termination_by self.rankMax - self.rank x +set_option linter.deprecated false in theorem parentD_findAux {self : UnionFind} {x : Fin self.size} : parentD (findAux self x).s i = if i = x then self.rootD x else parentD (self.findAux ⟨_, self.parent'_lt x⟩).s i := by diff --git a/Batteries/Lean/AttributeExtra.lean b/Batteries/Lean/AttributeExtra.lean index aeac0267c5..99eb7da612 100644 --- a/Batteries/Lean/AttributeExtra.lean +++ b/Batteries/Lean/AttributeExtra.lean @@ -4,10 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Batteries.Lean.TagAttribute +import Std.Data.HashMap.Basic + open Lean namespace Lean +open Std + /-- `TagAttributeExtra` works around a limitation of `TagAttribute`, which is that definitions must be tagged in the same file that declares the definition. @@ -73,7 +77,7 @@ structure ParametricAttributeExtra (α : Type) where /-- The underlying `ParametricAttribute`. -/ attr : ParametricAttribute α /-- A list of pre-tagged declarations with their values. -/ - base : HashMap Name α + base : Std.HashMap Name α deriving Inhabited /-- @@ -94,7 +98,7 @@ or `none` if `decl` is not tagged. -/ def getParam? [Inhabited α] (attr : ParametricAttributeExtra α) (env : Environment) (decl : Name) : Option α := - attr.attr.getParam? env decl <|> attr.base.find? decl + attr.attr.getParam? env decl <|> attr.base[decl]? /-- Applies attribute `attr` to declaration `decl`, given a value for the parameter. -/ def setParam (attr : ParametricAttributeExtra α) diff --git a/Batteries/Lean/HashMap.lean b/Batteries/Lean/HashMap.lean index a8e963cf94..0cb3b0f87d 100644 --- a/Batteries/Lean/HashMap.lean +++ b/Batteries/Lean/HashMap.lean @@ -4,22 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jannis Limperg -/ -import Lean.Data.HashMap - -namespace Lean.HashMap +import Std.Data.HashMap.Basic +namespace Std.HashMap variable [BEq α] [Hashable α] -instance : ForIn m (HashMap α β) (α × β) where - forIn m init f := do - let mut acc := init - for buckets in m.val.buckets.val do - for d in buckets do - match ← f d acc with - | .done b => return b - | .yield b => acc := b - return acc - /-- `O(|other|)` amortized. Merge two `HashMap`s. The values of keys which appear in both maps are combined using the monadic function `f`. @@ -28,7 +17,7 @@ The values of keys which appear in both maps are combined using the monadic func def mergeWithM {m α β} [BEq α] [Hashable α] [Monad m] (f : α → β → β → m β) (self other : HashMap α β) : m (HashMap α β) := other.foldM (init := self) fun map k v₂ => - match map.find? k with + match map[k]? with | none => return map.insert k v₂ | some v₁ => return map.insert k (← f k v₁ v₂) @@ -41,6 +30,6 @@ def mergeWith (f : α → β → β → β) (self other : HashMap α β) : HashM -- Implementing this function directly, rather than via `mergeWithM`, gives -- us less constrained universes. other.fold (init := self) fun map k v₂ => - match map.find? k with + match map[k]? with | none => map.insert k v₂ | some v₁ => map.insert k <| f k v₁ v₂ diff --git a/Batteries/Lean/HashSet.lean b/Batteries/Lean/HashSet.lean index 0dedb7cd4f..7882358c40 100644 --- a/Batteries/Lean/HashSet.lean +++ b/Batteries/Lean/HashSet.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jannis Limperg -/ -import Lean.Data.HashSet +import Std.Data.HashSet -namespace Lean.HashSet +namespace Std.HashSet variable [BEq α] [Hashable α] @@ -66,10 +66,3 @@ def insert' (s : HashSet α) (a : α) : HashSet α × Bool := @[inline] protected def ofArray [BEq α] [Hashable α] (as : Array α) : HashSet α := HashSet.empty.insertMany as - -/-- -`O(n)`. Obtain a `HashSet` from a list. --/ -@[inline] -protected def ofList [BEq α] [Hashable α] (as : List α) : HashSet α := - HashSet.empty.insertMany as diff --git a/Batteries/Lean/Meta/Inaccessible.lean b/Batteries/Lean/Meta/Inaccessible.lean index e63c4ec4f7..55f203b252 100644 --- a/Batteries/Lean/Meta/Inaccessible.lean +++ b/Batteries/Lean/Meta/Inaccessible.lean @@ -5,7 +5,7 @@ Authors: Jannis Limperg -/ import Lean.Meta.Basic -open Lean Lean.Meta +open Lean Lean.Meta Std /-- Obtain the inaccessible fvars from the given local context. An fvar is @@ -15,7 +15,7 @@ later fvar with the same user name. def Lean.LocalContext.inaccessibleFVars (lctx : LocalContext) : Array LocalDecl := let (result, _) := - lctx.foldr (β := Array LocalDecl × HashSet Name) + lctx.foldr (β := Array LocalDecl × Std.HashSet Name) (init := (Array.mkEmpty lctx.numIndices, {})) fun ldecl (result, seen) => if ldecl.isImplementationDetail then diff --git a/Batteries/Linter/UnnecessarySeqFocus.lean b/Batteries/Linter/UnnecessarySeqFocus.lean index 566da1d490..dca3aac1d7 100644 --- a/Batteries/Linter/UnnecessarySeqFocus.lean +++ b/Batteries/Linter/UnnecessarySeqFocus.lean @@ -8,7 +8,7 @@ import Lean.Linter.Util import Batteries.Lean.AttributeExtra namespace Batteries.Linter -open Lean Elab Command Linter +open Lean Elab Command Linter Std /-- Enables the 'unnecessary `<;>`' linter. This will warn whenever the `<;>` tactic combinator @@ -83,7 +83,7 @@ structure Entry where used : Bool /-- The monad for collecting used tactic syntaxes. -/ -abbrev M (ω) := StateRefT (HashMap String.Range Entry) (ST ω) +abbrev M (ω) := StateRefT (Std.HashMap String.Range Entry) (ST ω) /-- True if this is a `<;>` node in either `tactic` or `conv` classes. -/ @[inline] def isSeqFocus (k : SyntaxNodeKind) : Bool := @@ -120,7 +120,7 @@ partial def markUsedTactics : InfoTree → M ω Unit | .node i c => do if let .ofTacticInfo i := i then if let some r := i.stx.getRange? true then - if let some entry := (← get).find? r then + if let some entry := (← get)[r]? then if i.stx.getKind == ``Parser.Tactic.«tactic_<;>_» then let isBad := do unless i.goalsBefore.length == 1 || !multigoalAttr.hasTag env i.stx[0].getKind do diff --git a/Batteries/Linter/UnreachableTactic.lean b/Batteries/Linter/UnreachableTactic.lean index a4bedbe10c..1e55702674 100644 --- a/Batteries/Linter/UnreachableTactic.lean +++ b/Batteries/Linter/UnreachableTactic.lean @@ -8,7 +8,7 @@ import Lean.Linter.Util import Batteries.Tactic.Unreachable namespace Batteries.Linter -open Lean Elab Command Linter +open Lean Elab Command Linter Std /-- Enables the 'unreachable tactic' linter. This will warn on any tactics that are never executed. @@ -29,14 +29,14 @@ namespace UnreachableTactic def getLinterUnreachableTactic (o : Options) : Bool := getLinterValue linter.unreachableTactic o /-- The monad for collecting used tactic syntaxes. -/ -abbrev M := StateRefT (HashMap String.Range Syntax) IO +abbrev M := StateRefT (Std.HashMap String.Range Syntax) IO /-- A list of blacklisted syntax kinds, which are expected to have subterms that contain unevaluated tactics. -/ initialize ignoreTacticKindsRef : IO.Ref NameHashSet ← - IO.mkRef <| HashSet.empty + IO.mkRef <| Std.HashSet.empty |>.insert ``Parser.Term.binderTactic |>.insert ``Lean.Parser.Term.dynamicQuot |>.insert ``Lean.Parser.Tactic.quotSeq diff --git a/Batteries/Tactic/Lint/Frontend.lean b/Batteries/Tactic/Lint/Frontend.lean index 841c696af8..27ff3bf3b2 100644 --- a/Batteries/Tactic/Lint/Frontend.lean +++ b/Batteries/Tactic/Lint/Frontend.lean @@ -126,7 +126,7 @@ def sortResults (results : HashMap Name α) : CoreM <| Array (Name × α) := do for (n, _) in results.toArray do if let some range ← findDeclarationRanges? n then key := key.insert n <| range.range.pos.line - pure $ results.toArray.qsort fun (a, _) (b, _) => key.findD a 0 < key.findD b 0 + pure $ results.toArray.qsort fun (a, _) (b, _) => key.getD a 0 < key.getD b 0 /-- Formats a linter warning as `#check` command with comment. -/ def printWarning (declName : Name) (warning : MessageData) (useErrorFormat : Bool := false) @@ -158,7 +158,7 @@ def groupedByFilename (results : HashMap Name MessageData) (useErrorFormat : Boo let mod ← findModuleOf? declName let mod := mod.getD (← getEnv).mainModule grouped.insert mod <$> - match grouped.find? mod with + match grouped[mod]? with | some (fp, msgs) => pure (fp, msgs.insert declName msg) | none => do let fp ← if useErrorFormat then @@ -217,7 +217,7 @@ def getDeclsInPackage (pkg : Name) : CoreM (Array Name) := do let mut decls ← getDeclsInCurrModule let modules := env.header.moduleNames.map (pkg.isPrefixOf ·) return env.constants.map₁.fold (init := decls) fun decls declName _ => - if modules[env.const2ModIdx[declName].get! (α := Nat)]! then + if modules[env.const2ModIdx[declName]?.get! (α := Nat)]! then decls.push declName else decls diff --git a/Batteries/Tactic/Lint/Misc.lean b/Batteries/Tactic/Lint/Misc.lean index 3f0e69f648..a0aa8bc884 100644 --- a/Batteries/Tactic/Lint/Misc.lean +++ b/Batteries/Tactic/Lint/Misc.lean @@ -12,7 +12,7 @@ import Lean.Util.Recognizers import Lean.DocString import Batteries.Tactic.Lint.Basic -open Lean Meta +open Lean Meta Std namespace Std.Tactic.Lint @@ -143,7 +143,7 @@ In pseudo-mathematical form, this returns `{{p : parameter | p ∈ u} | (u : lev FIXME: We use `Array Name` instead of `HashSet Name`, since `HashSet` does not have an equality instance. It will ignore `nm₀.proof_i` declarations. -/ -private def univParamsGrouped (e : Expr) (nm₀ : Name) : Lean.HashSet (Array Name) := +private def univParamsGrouped (e : Expr) (nm₀ : Name) : HashSet (Array Name) := runST fun σ => do let res ← ST.mkRef (σ := σ) {} e.forEach fun diff --git a/lean-toolchain b/lean-toolchain index 5a9c76dc98..98556ba065 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.11.0 +leanprover/lean4:v4.12.0-rc1 diff --git a/scripts/check_imports.lean b/scripts/check_imports.lean index 94e778369f..55bf0e8342 100644 --- a/scripts/check_imports.lean +++ b/scripts/check_imports.lean @@ -70,7 +70,7 @@ def writeImportModule (path : FilePath) (imports : Array Name) : IO Unit := do /-- Check for imports and return true if warnings issued. -/ def checkMissingImports (modName : Name) (modData : ModuleData) (reqImports : Array Name) : LogIO Bool := do - let names : HashSet Name := HashSet.ofArray (modData.imports.map (·.module)) + let names : Std.HashSet Name := Std.HashSet.ofArray (modData.imports.map (·.module)) let mut warned := false for req in reqImports do if !names.contains req then diff --git a/test/where.lean b/test/where.lean index 6fd098ce53..a88f1ed8e5 100644 --- a/test/where.lean +++ b/test/where.lean @@ -2,6 +2,7 @@ import Batteries.Tactic.Where -- Return to pristine state set_option linter.missingDocs false +set_option internal.cmdlineSnapshots false /-- info: -- In root namespace with initial scope -/ #guard_msgs in #where From 8feac540abb781cb1349688c816dc02fae66b49c Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 3 Sep 2024 16:39:49 +1000 Subject: [PATCH 066/177] chore: upstream/undeprecate String.toAsciiByteArray (#943) --- Batteries/Data/String/Basic.lean | 23 +++++++++++++++++++++++ 1 file changed, 23 insertions(+) diff --git a/Batteries/Data/String/Basic.lean b/Batteries/Data/String/Basic.lean index d06df6eefa..6ff6ac98bf 100644 --- a/Batteries/Data/String/Basic.lean +++ b/Batteries/Data/String/Basic.lean @@ -104,3 +104,26 @@ def stripSuffix (s : String) (suff : Substring) : String := /-- Count the occurrences of a character in a string. -/ def count (s : String) (c : Char) : Nat := s.foldl (fun n d => if d = c then n + 1 else n) 0 + +/-- +Convert a string of assumed-ASCII characters into a byte array. +(If any characters are non-ASCII they will be reduced modulo 256.) + +Note: if you just need the underlying `ByteArray` of a non-ASCII string, +use `String.toUTF8`. +-/ +def toAsciiByteArray (s : String) : ByteArray := + let rec + /-- + Internal implementation of `toAsciiByteArray`. + `loop p out = out ++ toAsciiByteArray ({ s with startPos := p } : Substring)` + -/ + loop (p : Pos) (out : ByteArray) : ByteArray := + if h : s.atEnd p then out else + let c := s.get p + have : utf8ByteSize s - (next s p).byteIdx < utf8ByteSize s - p.byteIdx := + Nat.sub_lt_sub_left (Nat.lt_of_not_le <| mt decide_eq_true h) + (Nat.lt_add_of_pos_right (Char.utf8Size_pos _)) + loop (s.next p) (out.push c.toUInt8) + termination_by utf8ByteSize s - p.byteIdx + loop 0 ByteArray.empty From 5fc5df1b8089c40095ac348d7a56d71984c6443d Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Mon, 9 Sep 2024 01:03:43 -0400 Subject: [PATCH 067/177] feat: `#where` supports weak options (#932) --- Batteries/Tactic/Where.lean | 10 +++++++--- 1 file changed, 7 insertions(+), 3 deletions(-) diff --git a/Batteries/Tactic/Where.lean b/Batteries/Tactic/Where.lean index 7a07d2ba6b..3ef51a46c9 100644 --- a/Batteries/Tactic/Where.lean +++ b/Batteries/Tactic/Where.lean @@ -40,10 +40,14 @@ private def describeOpenDecls (ds : List OpenDecl) : MessageData := Id.run do private def describeOptions (opts : Options) : CommandElabM (Option MessageData) := do let mut lines := #[] + let decls ← getOptionDecls for (name, val) in opts do - let dval ← getOptionDefaultValue name - if val != dval then - lines := lines.push m!"set_option {name} {val}" + match decls.find? name with + | some decl => + if val != decl.defValue then + lines := lines.push m!"set_option {name} {val}" + | none => + lines := lines.push m!"-- set_option {name} {val} -- unknown" if lines.isEmpty then return none else From 39fef54668aba22be71f474ff281ed640af4dea0 Mon Sep 17 00:00:00 2001 From: Jon Eugster Date: Mon, 9 Sep 2024 07:04:01 +0200 Subject: [PATCH 068/177] chore: use emoji variant of unicode characters (#936) --- .github/workflows/nightly_detect_failure.yml | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/.github/workflows/nightly_detect_failure.yml b/.github/workflows/nightly_detect_failure.yml index 9a01257e21..98aaf457f0 100644 --- a/.github/workflows/nightly_detect_failure.yml +++ b/.github/workflows/nightly_detect_failure.yml @@ -24,7 +24,7 @@ jobs: type: 'stream' topic: 'Batteries status updates' content: | - ❌ The latest CI for Batteries' [`nightly-testing`](https://github.com/leanprover-community/batteries/tree/nightly-testing) branch has [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.event.workflow_run.id }}). + ❌️ The latest CI for Batteries' [`nightly-testing`](https://github.com/leanprover-community/batteries/tree/nightly-testing) branch has [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.event.workflow_run.id }}). # Whenever `nightly-testing` passes CI, # push it to `nightly-testing-YYYY-MM-DD` so we have a known good version of Batteries on that nightly release. @@ -77,13 +77,13 @@ jobs: } response = client.get_messages(request) messages = response['messages'] - if not messages or messages[0]['content'] != "✅ The latest CI for Batteries' [`nightly-testing`](https://github.com/leanprover-community/batteries/tree/nightly-testing) branch has succeeded!": + if not messages or messages[0]['content'] != "✅️ The latest CI for Batteries' [`nightly-testing`](https://github.com/leanprover-community/batteries/tree/nightly-testing) branch has succeeded!": # Post the success message request = { 'type': 'stream', 'to': 'nightly-testing', 'topic': 'Batteries status updates', - 'content': "✅ The latest CI for Batteries' [`nightly-testing`](https://github.com/leanprover-community/batteries/tree/nightly-testing) branch has succeeded!" + 'content': "✅️ The latest CI for Batteries' [`nightly-testing`](https://github.com/leanprover-community/batteries/tree/nightly-testing) branch has succeeded!" } result = client.send_message(request) print(result) From 2b13f6c5c80adfd95becf0f0b03b9a2d7ca52135 Mon Sep 17 00:00:00 2001 From: Shrys Date: Mon, 9 Sep 2024 07:04:26 +0200 Subject: [PATCH 069/177] chore: update .gitpod.yml to pre-build Batteries (#924) --- .gitpod.yml | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/.gitpod.yml b/.gitpod.yml index 5170403ac3..f9614f0d9b 100644 --- a/.gitpod.yml +++ b/.gitpod.yml @@ -4,3 +4,8 @@ image: vscode: extensions: - leanprover.lean4 + +tasks: + - init: | + elan self update + lake build From 869f2ad89b1d25377e779192d19dba5ede26fea4 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Mon, 9 Sep 2024 07:05:57 +0200 Subject: [PATCH 070/177] chore: reduce usage of refine' (#916) --- Batteries/Data/Rat/Lemmas.lean | 8 ++++---- Batteries/Data/String/Lemmas.lean | 4 ++-- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Batteries/Data/Rat/Lemmas.lean b/Batteries/Data/Rat/Lemmas.lean index e70914f6e0..77b911ec94 100644 --- a/Batteries/Data/Rat/Lemmas.lean +++ b/Batteries/Data/Rat/Lemmas.lean @@ -63,10 +63,10 @@ theorem normalize_eq_iff (z₁ : d₁ ≠ 0) (z₂ : d₂ ≠ 0) : normalize n₁ d₁ z₁ = normalize n₂ d₂ z₂ ↔ n₁ * d₂ = n₂ * d₁ := by constructor <;> intro h · simp only [normalize_eq, mk'.injEq] at h - have' hn₁ := Int.ofNat_dvd_left.2 <| Nat.gcd_dvd_left n₁.natAbs d₁ - have' hn₂ := Int.ofNat_dvd_left.2 <| Nat.gcd_dvd_left n₂.natAbs d₂ - have' hd₁ := Int.ofNat_dvd.2 <| Nat.gcd_dvd_right n₁.natAbs d₁ - have' hd₂ := Int.ofNat_dvd.2 <| Nat.gcd_dvd_right n₂.natAbs d₂ + have hn₁ := Int.ofNat_dvd_left.2 <| Nat.gcd_dvd_left n₁.natAbs d₁ + have hn₂ := Int.ofNat_dvd_left.2 <| Nat.gcd_dvd_left n₂.natAbs d₂ + have hd₁ := Int.ofNat_dvd.2 <| Nat.gcd_dvd_right n₁.natAbs d₁ + have hd₂ := Int.ofNat_dvd.2 <| Nat.gcd_dvd_right n₂.natAbs d₂ rw [← Int.ediv_mul_cancel (Int.dvd_trans hd₂ (Int.dvd_mul_left ..)), Int.mul_ediv_assoc _ hd₂, ← Int.ofNat_ediv, ← h.2, Int.ofNat_ediv, ← Int.mul_ediv_assoc _ hd₁, Int.mul_ediv_assoc' _ hn₁, diff --git a/Batteries/Data/String/Lemmas.lean b/Batteries/Data/String/Lemmas.lean index 7789f9dcd2..ddd43d32a4 100644 --- a/Batteries/Data/String/Lemmas.lean +++ b/Batteries/Data/String/Lemmas.lean @@ -922,14 +922,14 @@ theorem takeWhile (p : Char → Bool) : ∀ {s}, ValidFor l m r s → ValidFor l (m.takeWhile p) (m.dropWhile p ++ r) (s.takeWhile p) | _, ⟨⟩ => by simp only [Substring.takeWhile, takeWhileAux_of_valid] - refine' .of_eq .. <;> simp + apply ValidFor.of_eq <;> simp rw [← List.append_assoc, List.takeWhile_append_dropWhile] theorem dropWhile (p : Char → Bool) : ∀ {s}, ValidFor l m r s → ValidFor (l ++ m.takeWhile p) (m.dropWhile p) r (s.dropWhile p) | _, ⟨⟩ => by simp only [Substring.dropWhile, takeWhileAux_of_valid] - refine' .of_eq .. <;> simp + apply ValidFor.of_eq <;> simp rw [Nat.add_assoc, ← utf8Len_append (m.takeWhile p), List.takeWhile_append_dropWhile] -- TODO: takeRightWhile From afe9c5c9dac3d0acd44c56efdab758b0b29aa03c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 9 Sep 2024 01:14:38 -0400 Subject: [PATCH 071/177] feat: add `map` and `mapM` for scalar array types (#902) --- Batteries.lean | 1 + Batteries/Data/ByteArray.lean | 35 +++++++++++++++++++++++++++++ Batteries/Data/FloatArray.lean | 40 ++++++++++++++++++++++++++++++++++ 3 files changed, 76 insertions(+) create mode 100644 Batteries/Data/FloatArray.lean diff --git a/Batteries.lean b/Batteries.lean index 38e7617093..37d88fa3fb 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -23,6 +23,7 @@ import Batteries.Data.ByteSubarray import Batteries.Data.Char import Batteries.Data.DList import Batteries.Data.Fin +import Batteries.Data.FloatArray import Batteries.Data.HashMap import Batteries.Data.Int import Batteries.Data.LazyList diff --git a/Batteries/Data/ByteArray.lean b/Batteries/Data/ByteArray.lean index e051e01fd2..ae4544fd21 100644 --- a/Batteries/Data/ByteArray.lean +++ b/Batteries/Data/ByteArray.lean @@ -146,3 +146,38 @@ where (ofFnAux.go f i acc).data = Array.ofFn.go f i acc.data := by rw [ofFnAux.go, Array.ofFn.go]; split; rw [data_ofFnAux f (i+1), data_push]; rfl termination_by n - i + +/-! ### map/mapM -/ + +/-- +Unsafe optimized implementation of `mapM`. + +This function is unsafe because it relies on the implementation limit that the size of an array is +always less than `USize.size`. +-/ +@[inline] +unsafe def mapMUnsafe [Monad m] (a : ByteArray) (f : UInt8 → m UInt8) : m ByteArray := + loop a 0 a.usize +where + /-- Inner loop for `mapMUnsafe`. -/ + @[specialize] + loop (a : ByteArray) (k s : USize) := do + if k < a.usize then + let x := a.uget k lcProof + let y ← f x + let a := a.uset k y lcProof + loop a (k+1) s + else pure a + +/-- `mapM f a` applies the monadic function `f` to each element of the array. -/ +@[implemented_by mapMUnsafe] +def mapM [Monad m] (a : ByteArray) (f : UInt8 → m UInt8) : m ByteArray := do + let mut r := a + for i in [0:r.size] do + r := r.set! i (← f r[i]!) + return r + +/-- `map f a` applies the function `f` to each element of the array. -/ +@[inline] +def map (a : ByteArray) (f : UInt8 → UInt8) : ByteArray := + mapM (m:=Id) a f diff --git a/Batteries/Data/FloatArray.lean b/Batteries/Data/FloatArray.lean new file mode 100644 index 0000000000..a0ef34073c --- /dev/null +++ b/Batteries/Data/FloatArray.lean @@ -0,0 +1,40 @@ +/- +Copyright (c) 2024 François G. Dorais. All rights reserved. +Released under Apache 2. license as described in the file LICENSE. +Authors: François G. Dorais +-/ + +namespace FloatArray + +/-- +Unsafe optimized implementation of `mapM`. + +This function is unsafe because it relies on the implementation limit that the size of an array is +always less than `USize.size`. +-/ +@[inline] +unsafe def mapMUnsafe [Monad m] (a : FloatArray) (f : Float → m Float) : m FloatArray := + loop a 0 a.usize +where + /-- Inner loop for `mapMUnsafe`. -/ + @[specialize] + loop (a : FloatArray) (k s : USize) := do + if k < s then + let x := a.uget k lcProof + let y ← f x + let a := a.uset k y lcProof + loop a (k+1) s + else pure a + +/-- `mapM f a` applies the monadic function `f` to each element of the array. -/ +@[implemented_by mapMUnsafe] +def mapM [Monad m] (a : FloatArray) (f : Float → m Float) : m FloatArray := do + let mut r := a + for i in [0:r.size] do + r := r.set! i (← f r[i]!) + return r + +/-- `map f a` applies the function `f` to each element of the array. -/ +@[inline] +def map (a : FloatArray) (f : Float → Float) : FloatArray := + mapM (m:=Id) a f From 4bcbb3f1244f9b22cb26c889b65cfd01fd8f7c25 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 9 Sep 2024 01:15:39 -0400 Subject: [PATCH 072/177] feat: lemmas for `Array.insertAt` (#895) --- Batteries/Data/Array/Lemmas.lean | 96 ++++++++++++++++++++++++++++++++ 1 file changed, 96 insertions(+) diff --git a/Batteries/Data/Array/Lemmas.lean b/Batteries/Data/Array/Lemmas.lean index dd4fe12424..c943533670 100644 --- a/Batteries/Data/Array/Lemmas.lean +++ b/Batteries/Data/Array/Lemmas.lean @@ -149,3 +149,99 @@ theorem mem_singleton : a ∈ #[b] ↔ a = b := by simp alias append_empty := append_nil alias empty_append := nil_append + +/-! ### insertAt -/ + +private theorem size_insertAt_loop (as : Array α) (i : Fin (as.size+1)) (j : Fin bs.size) : + (insertAt.loop as i bs j).size = bs.size := by + unfold insertAt.loop + split + · rw [size_insertAt_loop, size_swap] + · rfl + +theorem size_insertAt (as : Array α) (i : Fin (as.size+1)) (v : α) : + (as.insertAt i v).size = as.size + 1 := by + rw [insertAt, size_insertAt_loop, size_push] + +private theorem get_insertAt_loop_lt (as : Array α) (i : Fin (as.size+1)) (j : Fin bs.size) + (k) (hk : k < (insertAt.loop as i bs j).size) (h : k < i) : + (insertAt.loop as i bs j)[k] = bs[k]'(size_insertAt_loop .. ▸ hk) := by + unfold insertAt.loop + split + · have h1 : k ≠ j - 1 := by omega + have h2 : k ≠ j := by omega + rw [get_insertAt_loop_lt, get_swap, if_neg h1, if_neg h2] + exact h + · rfl + +private theorem get_insertAt_loop_gt (as : Array α) (i : Fin (as.size+1)) (j : Fin bs.size) + (k) (hk : k < (insertAt.loop as i bs j).size) (hgt : j < k) : + (insertAt.loop as i bs j)[k] = bs[k]'(size_insertAt_loop .. ▸ hk) := by + unfold insertAt.loop + split + · have h1 : k ≠ j - 1 := by omega + have h2 : k ≠ j := by omega + rw [get_insertAt_loop_gt, get_swap, if_neg h1, if_neg h2] + exact Nat.lt_of_le_of_lt (Nat.pred_le _) hgt + · rfl + +private theorem get_insertAt_loop_eq (as : Array α) (i : Fin (as.size+1)) (j : Fin bs.size) + (k) (hk : k < (insertAt.loop as i bs j).size) (heq : i = k) (h : i.val ≤ j.val) : + (insertAt.loop as i bs j)[k] = bs[j] := by + unfold insertAt.loop + split + · next h => + rw [get_insertAt_loop_eq, Fin.getElem_fin, get_swap, if_pos rfl] + exact Nat.lt_of_le_of_lt (Nat.pred_le _) j.is_lt + exact heq + exact Nat.le_pred_of_lt h + · congr; omega + +private theorem get_insertAt_loop_gt_le (as : Array α) (i : Fin (as.size+1)) (j : Fin bs.size) + (k) (hk : k < (insertAt.loop as i bs j).size) (hgt : i < k) (hle : k ≤ j) : + (insertAt.loop as i bs j)[k] = bs[k-1] := by + unfold insertAt.loop + split + · next h => + if h0 : k = j then + cases h0 + have h1 : j.val ≠ j - 1 := by omega + rw [get_insertAt_loop_gt, get_swap, if_neg h1, if_pos rfl]; rfl + · exact j.is_lt + · exact Nat.pred_lt_of_lt hgt + else + have h1 : k - 1 ≠ j - 1 := by omega + have h2 : k - 1 ≠ j := by omega + rw [get_insertAt_loop_gt_le, get_swap, if_neg h1, if_neg h2] + exact hgt + apply Nat.le_of_lt_add_one + rw [Nat.sub_one_add_one] + exact Nat.lt_of_le_of_ne hle h0 + exact Nat.not_eq_zero_of_lt h + · next h => + absurd h + exact Nat.lt_of_lt_of_le hgt hle + +theorem getElem_insertAt_lt (as : Array α) (i : Fin (as.size+1)) (v : α) + (k) (hlt : k < i.val) {hk : k < (as.insertAt i v).size} {hk' : k < as.size} : + (as.insertAt i v)[k] = as[k] := by + simp only [insertAt] + rw [get_insertAt_loop_lt, get_push, dif_pos hk'] + exact hlt + +theorem getElem_insertAt_gt (as : Array α) (i : Fin (as.size+1)) (v : α) + (k) (hgt : k > i.val) {hk : k < (as.insertAt i v).size} {hk' : k - 1 < as.size} : + (as.insertAt i v)[k] = as[k - 1] := by + simp only [insertAt] + rw [get_insertAt_loop_gt_le, get_push, dif_pos hk'] + exact hgt + rw [size_insertAt] at hk + exact Nat.le_of_lt_succ hk + +theorem getElem_insertAt_eq (as : Array α) (i : Fin (as.size+1)) (v : α) + (k) (heq : i.val = k) {hk : k < (as.insertAt i v).size} : + (as.insertAt i v)[k] = v := by + simp only [insertAt] + rw [get_insertAt_loop_eq, Fin.getElem_fin, get_push_eq] + exact heq + exact Nat.le_of_lt_succ i.is_lt From affe669960e24d47760e3b71782dbbed70f93409 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Mon, 9 Sep 2024 01:16:20 -0400 Subject: [PATCH 073/177] chore: remove upstreamed code actions (#888) --- Batteries/CodeAction/Basic.lean | 124 +-------------------------- Batteries/CodeAction/Deprecated.lean | 2 +- 2 files changed, 2 insertions(+), 124 deletions(-) diff --git a/Batteries/CodeAction/Basic.lean b/Batteries/CodeAction/Basic.lean index 1a7488bd87..17cca0c7d6 100644 --- a/Batteries/CodeAction/Basic.lean +++ b/Batteries/CodeAction/Basic.lean @@ -20,129 +20,7 @@ itself.) -/ namespace Std.CodeAction -open Lean Elab Term Server RequestM - -/-- -The return value of `findTactic?`. -This is the syntax for which code actions will be triggered. --/ -inductive FindTacticResult - /-- The nearest enclosing tactic is a tactic, with the given syntax stack. -/ - | tactic : Syntax.Stack → FindTacticResult - /-- The cursor is between tactics, and the nearest enclosing range is a tactic sequence. - Code actions will insert tactics at index `insertIdx` into the syntax - (which is a nullNode of `tactic;*` inside a `tacticSeqBracketed` or `tacticSeq1Indented`). -/ - | tacticSeq : (preferred : Bool) → (insertIdx : Nat) → Syntax.Stack → FindTacticResult - -/-- -Find the syntax on which to trigger tactic code actions. -This is a pure syntax pass, without regard to elaboration information. - -* `preferred : String.Pos → Bool`: used to select "preferred `tacticSeq`s" based on the cursor - column, when the cursor selection would otherwise be ambiguous. For example, in: - ``` - · foo - · bar - baz - | - ``` - where the cursor is at the `|`, we select the `tacticSeq` starting with `foo`, while if the - cursor was indented to align with `baz` then we would select the `bar; baz` sequence instead. - -* `range`: the cursor selection. We do not do much with range selections; if a range selection - covers more than one tactic then we abort. - -* `root`: the root syntax to process - -The return value is either a selected tactic, or a selected point in a tactic sequence. --/ -partial def findTactic? (preferred : String.Pos → Bool) (range : String.Range) - (root : Syntax) : Option FindTacticResult := do _ ← visit root; ← go [] root -where - /-- Returns `none` if we should not visit this syntax at all, and `some false` if we only - want to visit it in "extended" mode (where we include trailing characters). -/ - visit (stx : Syntax) (prev? : Option String.Pos := none) : Option Bool := do - let left ← stx.getPos? true - guard <| prev?.getD left ≤ range.start - let .original (endPos := right) (trailing := trailing) .. := stx.getTailInfo | none - guard <| right.byteIdx + trailing.bsize ≥ range.stop.byteIdx - return left ≤ range.start && right ≥ range.stop - - /-- Merges the results of two `FindTacticResult`s. This just prefers the second (inner) one, - unless the inner tactic is a dispreferred tactic sequence and the outer one is preferred. - This is used to implement whitespace-sensitive selection of tactic sequences. -/ - merge : (r₁ : Option FindTacticResult) → (r₂ : FindTacticResult) → FindTacticResult - | some r₁@(.tacticSeq (preferred := true) ..), .tacticSeq (preferred := false) .. => r₁ - | _, r₂ => r₂ - - /-- Main recursion for `findTactic?`. This takes a `stack` context and a root syntax `stx`, - and returns the best `FindTacticResult` it can find. It returns `none` (abort) if two or more - results are found, and `some none` (none yet) if no results are found. -/ - go (stack : Syntax.Stack) (stx : Syntax) (prev? : Option String.Pos := none) : - Option (Option FindTacticResult) := do - if stx.getKind == ``Parser.Tactic.tacticSeq then - -- TODO: this implementation is a bit too strict about the beginning of tacticSeqs. - -- We would like to be able to parse - -- · | - -- foo - -- (where `|` is the cursor position) as an insertion into the sequence containing foo - -- at index 0, but we currently use the start of the tacticSeq, which is the foo token, - -- as the earliest possible location that will be associated to the sequence. - let bracket := stx[0].getKind == ``Parser.Tactic.tacticSeqBracketed - let argIdx := if bracket then 1 else 0 - let (stack, stx) := ((stx[0], argIdx) :: (stx, 0) :: stack, stx[0][argIdx]) - let mainRes := stx[0].getPos?.map fun pos => - let i := Id.run do - for i in [0:stx.getNumArgs] do - if let some pos' := stx[2*i].getPos? then - if range.stop < pos' then - return i - (stx.getNumArgs + 1) / 2 - .tacticSeq (bracket || preferred pos) i ((stx, 0) :: stack) - let mut childRes := none - for i in [0:stx.getNumArgs:2] do - if let some inner := visit stx[i] then - let stack := (stx, i) :: stack - if let some child := (← go stack stx[i]) <|> - (if inner then some (.tactic ((stx[i], 0) :: stack)) else none) - then - if childRes.isSome then failure - childRes := merge mainRes child - return childRes <|> mainRes - else - let mut childRes := none - let mut prev? := prev? - for i in [0:stx.getNumArgs] do - if let some _ := visit stx[i] prev? then - if let some child ← go ((stx, i) :: stack) stx[i] prev? then - if childRes.isSome then failure - childRes := child - prev? := stx[i].getTailPos? true <|> prev? - return childRes - -/-- -Returns the info tree corresponding to a syntax, using `kind` and `range` for identification. -(This is not foolproof, but it is a fairly accurate proxy for `Syntax` equality and a lot cheaper -than deep comparison.) --/ -partial def findInfoTree? (kind : SyntaxNodeKind) (tgtRange : String.Range) - (ctx? : Option ContextInfo) (t : InfoTree) - (f : ContextInfo → Info → Bool) (canonicalOnly := false) : - Option (ContextInfo × InfoTree) := - match t with - | .context ctx t => findInfoTree? kind tgtRange (ctx.mergeIntoOuter? ctx?) t f canonicalOnly - | node@(.node i ts) => do - if let some ctx := ctx? then - if let some range := i.stx.getRange? canonicalOnly then - -- FIXME: info tree needs to be organized better so that this works - -- guard <| range.includes tgtRange - if i.stx.getKind == kind && range == tgtRange && f ctx i then - return (ctx, node) - for t in ts do - if let some res := findInfoTree? kind tgtRange (i.updateContext? ctx?) t f canonicalOnly then - return res - none - | _ => none +open Lean Elab Server RequestM CodeAction /-- A code action which calls `@[tactic_code_action]` code actions. -/ @[code_action_provider] def tacticCodeActionProvider : CodeActionProvider := fun params snap => do diff --git a/Batteries/CodeAction/Deprecated.lean b/Batteries/CodeAction/Deprecated.lean index ce98bcc5c6..33d1d75243 100644 --- a/Batteries/CodeAction/Deprecated.lean +++ b/Batteries/CodeAction/Deprecated.lean @@ -15,7 +15,7 @@ whenever the deprecation lint also fires, allowing the user to replace the usage constant. -/ namespace Std -open Lean Elab Server Lsp RequestM +open Lean Elab Server Lsp RequestM CodeAction /-- An environment extension for identifying `@[deprecated]` definitions which can be auto-fixed -/ initialize machineApplicableDeprecated : TagDeclarationExtension ← mkTagDeclarationExtension From e0017bd4c7adacc3de25a0d7036d20968cb9e1c7 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 9 Sep 2024 15:17:06 +1000 Subject: [PATCH 074/177] chore: fix some List.modifyNth lemma names (#831) --- Batteries/Data/List/Lemmas.lean | 12 ++++++++---- 1 file changed, 8 insertions(+), 4 deletions(-) diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index bb7255f4d0..a68d820a52 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -72,11 +72,13 @@ theorem get?_modifyNth (f : α → α) (n) (l : List α) (m) : (modifyNth f n l).get? m = (fun a => if n = m then f a else a) <$> l.get? m := by simp [getElem?_modifyNth] -theorem modifyNthTail_length (f : List α → List α) (H : ∀ l, length (f l) = length l) : +theorem length_modifyNthTail (f : List α → List α) (H : ∀ l, length (f l) = length l) : ∀ n l, length (modifyNthTail f n l) = length l | 0, _ => H _ | _+1, [] => rfl - | _+1, _ :: _ => congrArg (·+1) (modifyNthTail_length _ H _ _) + | _+1, _ :: _ => congrArg (·+1) (length_modifyNthTail _ H _ _) + +@[deprecated (since := "2024-06-07")] alias modifyNthTail_length := length_modifyNthTail theorem modifyNthTail_add (f : List α → List α) (n) (l₁ l₂ : List α) : modifyNthTail f (l₁.length + n) (l₁ ++ l₂) = l₁ ++ modifyNthTail f n l₂ := by @@ -88,8 +90,10 @@ theorem exists_of_modifyNthTail (f : List α → List α) {n} {l : List α} (h : ⟨_, _, (take_append_drop n l).symm, length_take_of_le h⟩ ⟨_, _, eq, hl, hl ▸ eq ▸ modifyNthTail_add (n := 0) ..⟩ -@[simp] theorem modify_get?_length (f : α → α) : ∀ n l, length (modifyNth f n l) = length l := - modifyNthTail_length _ fun l => by cases l <;> rfl +@[simp] theorem length_modifyNth (f : α → α) : ∀ n l, length (modifyNth f n l) = length l := + length_modifyNthTail _ fun l => by cases l <;> rfl + +@[deprecated (since := "2024-06-07")] alias modify_get?_length := length_modifyNth @[simp] theorem getElem?_modifyNth_eq (f : α → α) (n) (l : List α) : (modifyNth f n l)[n]? = f <$> l[n]? := by From 1a28ab04c34939729fca9aecd67e32bf0c0aa9b8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 9 Sep 2024 01:17:52 -0400 Subject: [PATCH 075/177] feat: size lemma for `Array.set!` (#807) --- Batteries/Data/Array/Lemmas.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Batteries/Data/Array/Lemmas.lean b/Batteries/Data/Array/Lemmas.lean index c943533670..02293aa349 100644 --- a/Batteries/Data/Array/Lemmas.lean +++ b/Batteries/Data/Array/Lemmas.lean @@ -131,6 +131,11 @@ theorem size_shrink (a : Array α) (n) : (a.shrink n).size = min a.size n := by simp [shrink, size_shrink_loop] omega +/-! ### set -/ + +theorem size_set! (a : Array α) (i v) : (a.set! i v).size = a.size := by + rw [set!_is_setD, size_setD] + /-! ### map -/ theorem mapM_empty [Monad m] (f : α → m β) : mapM f #[] = pure #[] := by From 1af15aa183d46ea5284289e18f75a8d018a3f439 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 9 Sep 2024 15:20:30 +1000 Subject: [PATCH 076/177] chore: fix imports --- Batteries/CodeAction/Basic.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Batteries/CodeAction/Basic.lean b/Batteries/CodeAction/Basic.lean index 17cca0c7d6..378f6aa00d 100644 --- a/Batteries/CodeAction/Basic.lean +++ b/Batteries/CodeAction/Basic.lean @@ -6,6 +6,7 @@ Authors: Mario Carneiro import Lean.Elab.BuiltinTerm import Lean.Elab.BuiltinNotation import Lean.Server.InfoUtils +import Lean.Server.CodeActions.Provider import Batteries.CodeAction.Attr /-! From d11566f4c8ca80dbe87630b606c608274c7380fd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 9 Sep 2024 01:26:59 -0400 Subject: [PATCH 077/177] feat: add `Array.Pairwise` (#897) --- Batteries/Data/Array.lean | 1 + Batteries/Data/Array/Pairwise.lean | 61 ++++++++++++++++++++++++++++++ 2 files changed, 62 insertions(+) create mode 100644 Batteries/Data/Array/Pairwise.lean diff --git a/Batteries/Data/Array.lean b/Batteries/Data/Array.lean index 3a9bb1fef1..4d41e75755 100644 --- a/Batteries/Data/Array.lean +++ b/Batteries/Data/Array.lean @@ -4,3 +4,4 @@ import Batteries.Data.Array.Lemmas import Batteries.Data.Array.Match import Batteries.Data.Array.Merge import Batteries.Data.Array.Monadic +import Batteries.Data.Array.Pairwise diff --git a/Batteries/Data/Array/Pairwise.lean b/Batteries/Data/Array/Pairwise.lean new file mode 100644 index 0000000000..84dff1f680 --- /dev/null +++ b/Batteries/Data/Array/Pairwise.lean @@ -0,0 +1,61 @@ +/- +Copyright (c) 2024 François G. Dorais. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: François G. Dorais +-/ +import Batteries.Data.Array.Lemmas +import Batteries.Data.List.Pairwise + +namespace Array + +/-- +`Pairwise R as` means that all the elements of the array `as` are `R`-related to all elements with +larger indices. + +`Pairwise R #[1, 2, 3] ↔ R 1 2 ∧ R 1 3 ∧ R 2 3` + +For example `as.Pairwise (· ≠ ·)` asserts that `as` has no duplicates, `as.Pairwise (· < ·)` asserts +that `as` is strictly sorted and `as.Pairwise (· ≤ ·)` asserts that `as` is weakly sorted. +-/ +def Pairwise (R : α → α → Prop) (as : Array α) : Prop := as.data.Pairwise R + +theorem pairwise_iff_get {as : Array α} : as.Pairwise R ↔ + ∀ (i j : Fin as.size), i < j → R (as.get i) (as.get j) := by + unfold Pairwise; simp [List.pairwise_iff_get, getElem_fin_eq_data_get]; rfl + +theorem pairwise_iff_getElem {as : Array α} : as.Pairwise R ↔ + ∀ (i j : Nat) (_ : i < as.size) (_ : j < as.size), i < j → R as[i] as[j] := by + unfold Pairwise; simp [List.pairwise_iff_getElem, data_length]; rfl + +instance (R : α → α → Prop) [DecidableRel R] (as) : Decidable (Pairwise R as) := + have : (∀ (j : Fin as.size) (i : Fin j.val), R as[i.val] (as[j.val])) ↔ Pairwise R as := by + rw [pairwise_iff_getElem] + constructor + · intro h i j _ hj hlt; exact h ⟨j, hj⟩ ⟨i, hlt⟩ + · intro h ⟨j, hj⟩ ⟨i, hlt⟩; exact h i j (Nat.lt_trans hlt hj) hj hlt + decidable_of_iff _ this + +theorem pairwise_empty : #[].Pairwise R := by + unfold Pairwise; exact List.Pairwise.nil + +theorem pairwise_singleton (R : α → α → Prop) (a) : #[a].Pairwise R := by + unfold Pairwise; exact List.pairwise_singleton .. + +theorem pairwise_pair : #[a, b].Pairwise R ↔ R a b := by + unfold Pairwise; exact List.pairwise_pair + +theorem pairwise_append {as bs : Array α} : + (as ++ bs).Pairwise R ↔ as.Pairwise R ∧ bs.Pairwise R ∧ (∀ x ∈ as, ∀ y ∈ bs, R x y) := by + unfold Pairwise; simp [← mem_data, append_data, ← List.pairwise_append] + +theorem pairwise_push {as : Array α} : + (as.push a).Pairwise R ↔ as.Pairwise R ∧ (∀ x ∈ as, R x a) := by + unfold Pairwise + simp [← mem_data, push_data, List.pairwise_append, List.pairwise_singleton, List.mem_singleton] + +theorem pairwise_extract {as : Array α} (h : as.Pairwise R) (start stop) : + (as.extract start stop).Pairwise R := by + simp only [pairwise_iff_getElem, get_extract, size_extract] at h ⊢ + intro _ _ _ _ hlt + apply h + exact Nat.add_lt_add_left hlt start From 46fed98b5cac2b1ea64e363b420c382ed1af0d85 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 10 Sep 2024 16:50:25 +1000 Subject: [PATCH 078/177] chore: remove >6 month deprecations (#945) --- Batteries.lean | 4 -- Batteries/Data/BitVec.lean | 1 - Batteries/Data/BitVec/Lemmas.lean | 16 ------ Batteries/Data/Bool.lean | 19 ------- Batteries/Data/Int.lean | 3 -- Batteries/Data/Int/DivMod.lean | 84 ------------------------------- Batteries/Data/Int/Lemmas.lean | 6 --- Batteries/Data/Int/Order.lean | 10 ---- Batteries/Data/List/Basic.lean | 17 ------- Batteries/Data/List/Count.lean | 8 --- Batteries/Data/List/Pairwise.lean | 12 ----- Batteries/Data/Nat/Lemmas.lean | 21 -------- Batteries/Data/Option.lean | 1 - Batteries/Data/Option/Lemmas.lean | 13 ----- scripts/check_imports.lean | 8 +-- 15 files changed, 4 insertions(+), 219 deletions(-) delete mode 100644 Batteries/Data/BitVec.lean delete mode 100644 Batteries/Data/BitVec/Lemmas.lean delete mode 100644 Batteries/Data/Bool.lean delete mode 100644 Batteries/Data/Int.lean delete mode 100644 Batteries/Data/Int/DivMod.lean delete mode 100644 Batteries/Data/Int/Lemmas.lean delete mode 100644 Batteries/Data/Int/Order.lean delete mode 100644 Batteries/Data/Option.lean delete mode 100644 Batteries/Data/Option/Lemmas.lean diff --git a/Batteries.lean b/Batteries.lean index 37d88fa3fb..3bd1f4ef79 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -16,8 +16,6 @@ import Batteries.Data.Array import Batteries.Data.AssocList import Batteries.Data.BinaryHeap import Batteries.Data.BinomialHeap -import Batteries.Data.BitVec -import Batteries.Data.Bool import Batteries.Data.ByteArray import Batteries.Data.ByteSubarray import Batteries.Data.Char @@ -25,12 +23,10 @@ import Batteries.Data.DList import Batteries.Data.Fin import Batteries.Data.FloatArray import Batteries.Data.HashMap -import Batteries.Data.Int import Batteries.Data.LazyList import Batteries.Data.List import Batteries.Data.MLList import Batteries.Data.Nat -import Batteries.Data.Option import Batteries.Data.PairingHeap import Batteries.Data.RBMap import Batteries.Data.Range diff --git a/Batteries/Data/BitVec.lean b/Batteries/Data/BitVec.lean deleted file mode 100644 index ff4358f07d..0000000000 --- a/Batteries/Data/BitVec.lean +++ /dev/null @@ -1 +0,0 @@ -import Batteries.Data.BitVec.Lemmas diff --git a/Batteries/Data/BitVec/Lemmas.lean b/Batteries/Data/BitVec/Lemmas.lean deleted file mode 100644 index 2bcd3ce4aa..0000000000 --- a/Batteries/Data/BitVec/Lemmas.lean +++ /dev/null @@ -1,16 +0,0 @@ -/- -Copyright (c) 2023 Lean FRO, LLC. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Joe Hendrix --/ -import Batteries.Tactic.Alias - -namespace BitVec - -@[deprecated (since := "2024-02-07")] alias zero_is_unique := eq_nil - -/-! ### sub/neg -/ - -@[deprecated (since := "2024-02-07")] alias sub_toNat := toNat_sub - -@[deprecated (since := "2024-02-07")] alias neg_toNat := toNat_neg diff --git a/Batteries/Data/Bool.lean b/Batteries/Data/Bool.lean deleted file mode 100644 index b6c0c9b4e7..0000000000 --- a/Batteries/Data/Bool.lean +++ /dev/null @@ -1,19 +0,0 @@ -/- -Copyright (c) 2023 F. G. Dorais. No rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: F. G. Dorais --/ - -import Batteries.Tactic.Alias - -namespace Bool - -/-! ### injectivity lemmas -/ - -@[deprecated (since := "2023-10-27")] alias not_inj' := not_inj_iff - -@[deprecated (since := "2023-10-27")] alias and_or_inj_right' := and_or_inj_right_iff - -@[deprecated (since := "2023-10-27")] alias and_or_inj_left' := and_or_inj_left_iff - -end Bool diff --git a/Batteries/Data/Int.lean b/Batteries/Data/Int.lean deleted file mode 100644 index 29b6c2e4a9..0000000000 --- a/Batteries/Data/Int.lean +++ /dev/null @@ -1,3 +0,0 @@ -import Batteries.Data.Int.DivMod -import Batteries.Data.Int.Lemmas -import Batteries.Data.Int.Order diff --git a/Batteries/Data/Int/DivMod.lean b/Batteries/Data/Int/DivMod.lean deleted file mode 100644 index 7558c98a3d..0000000000 --- a/Batteries/Data/Int/DivMod.lean +++ /dev/null @@ -1,84 +0,0 @@ -/- -Copyright (c) 2016 Jeremy Avigad. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jeremy Avigad, Mario Carneiro --/ -import Batteries.Data.Int.Order - -/-! -# Lemmas about integer division --/ - - -open Nat - -namespace Int - -/-! -### The following lemmas have been commented out here for a while, and need restoration. --/ - -/- -theorem eq_mul_ediv_of_mul_eq_mul_of_dvd_left {a b c d : Int} - (hb : b ≠ 0) (hbc : b ∣ c) (h : b * a = c * d) : a = c / b * d := by - rcases hbc with ⟨k, hk⟩ - subst hk - rw [Int.mul_ediv_cancel_left _ hb] - rw [Int.mul_assoc] at h - apply mul_left_cancel₀ hb h - -/-- If an integer with larger absolute value divides an integer, it is -zero. -/ -theorem eq_zero_of_dvd_ofNatAbs_lt_natAbs {a b : Int} (w : a ∣ b) (h : natAbs b < natAbs a) : - b = 0 := by - rw [← natAbs_dvd, ← dvd_natAbs, ofNat_dvd] at w - rw [← natAbs_eq_zero] - exact eq_zero_of_dvd_of_lt w h - -theorem eq_zero_of_dvd_of_nonneg_of_lt {a b : Int} (w₁ : 0 ≤ a) (w₂ : a < b) (h : b ∣ a) : a = 0 := - eq_zero_of_dvd_ofNatAbs_lt_natAbs h (natAbs_lt_natAbs_of_nonneg_of_lt w₁ w₂) - -/-- If two integers are congruent to a sufficiently large modulus, -they are equal. -/ -theorem eq_of_mod_eq_ofNatAbs_sub_lt_natAbs {a b c : Int} - (h1 : a % b = c) (h2 : natAbs (a - c) < natAbs b) : a = c := - Int.eq_of_sub_eq_zero (eq_zero_of_dvd_ofNatAbs_lt_natAbs (dvd_sub_of_emod_eq h1) h2) - -theorem ofNat_add_negSucc_of_lt {m n : Nat} (h : m < n.succ) : ofNat m + -[n+1] = -[n+1 - m] := by - change subNatNat _ _ = _ - have h' : n.succ - m = (n - m).succ - apply succ_sub - apply le_of_lt_succ h - simp [*, subNatNat] - -theorem ofNat_add_negSucc_of_ge {m n : Nat} (h : n.succ ≤ m) : - ofNat m + -[n+1] = ofNat (m - n.succ) := by - change subNatNat _ _ = _ - have h' : n.succ - m = 0 - apply tsub_eq_zero_iff_le.mpr h - simp [*, subNatNat] - -@[simp] -theorem neg_add_neg (m n : Nat) : -[m+1] + -[n+1] = -[Nat+1.succ (m + n)] := - rfl - -theorem natAbs_le_of_dvd_ne_zero {s t : Int} (hst : s ∣ t) (ht : t ≠ 0) : natAbs s ≤ natAbs t := - not_lt.mp (mt (eq_zero_of_dvd_ofNatAbs_lt_natAbs hst) ht) - -theorem natAbs_eq_of_dvd_dvd {s t : Int} (hst : s ∣ t) (hts : t ∣ s) : natAbs s = natAbs t := - Nat.dvd_antisymm (natAbs_dvd_iff_dvd.mpr hst) (natAbs_dvd_iff_dvd.mpr hts) - -theorem div_dvd_of_dvd {s t : Int} (hst : s ∣ t) : t / s ∣ t := by - rcases eq_or_ne s 0 with (rfl | hs) - · simpa using hst - rcases hst with ⟨c, hc⟩ - simp [hc, Int.mul_div_cancel_left _ hs] - -theorem dvd_div_of_mul_dvd {a b c : Int} (h : a * b ∣ c) : b ∣ c / a := by - rcases eq_or_ne a 0 with (rfl | ha) - · simp only [Int.div_zero, dvd_zero] - - rcases h with ⟨d, rfl⟩ - refine' ⟨d, _⟩ - rw [mul_assoc, Int.mul_div_cancel_left _ ha] --/ diff --git a/Batteries/Data/Int/Lemmas.lean b/Batteries/Data/Int/Lemmas.lean deleted file mode 100644 index 3ce6c74c04..0000000000 --- a/Batteries/Data/Int/Lemmas.lean +++ /dev/null @@ -1,6 +0,0 @@ --- This is a backwards compatibility shim, --- after `Batteries.Data.Int.Lemmas` was split into smaller files. --- Hopefully it can later be removed. - -import Batteries.Data.Int.Order -import Batteries.Data.Int.DivMod diff --git a/Batteries/Data/Int/Order.lean b/Batteries/Data/Int/Order.lean deleted file mode 100644 index daab75c612..0000000000 --- a/Batteries/Data/Int/Order.lean +++ /dev/null @@ -1,10 +0,0 @@ -/- -Copyright (c) 2016 Jeremy Avigad. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jeremy Avigad, Deniz Aydin, Floris van Doorn, Mario Carneiro --/ -import Batteries.Tactic.Alias - -namespace Int - -@[deprecated (since := "2024-01-24")] alias ofNat_natAbs_eq_of_nonneg := natAbs_of_nonneg diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index a44f430c8d..963920017c 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -706,23 +706,6 @@ Defined as `pwFilter (≠)`. eraseDup [1, 0, 2, 2, 1] = [0, 2, 1] -/ @[inline] def eraseDup [BEq α] : List α → List α := pwFilter (· != ·) -/-- -`ilast' x xs` returns the last element of `xs` if `xs` is non-empty; it returns `x` otherwise. -Use `List.getLastD` instead. --/ -@[simp, deprecated getLastD (since := "2024-01-09")] def ilast' {α} : α → List α → α - | a, [] => a - | _, b :: l => ilast' b l - -/-- -`last' xs` returns the last element of `xs` if `xs` is non-empty; it returns `none` otherwise. -Use `List.getLast?` instead --/ -@[simp, deprecated getLast? (since := "2024-01-09")] def last' {α} : List α → Option α - | [] => none - | [a] => some a - | _ :: l => last' l - /-- `rotate l n` rotates the elements of `l` to the left by `n` ``` diff --git a/Batteries/Data/List/Count.lean b/Batteries/Data/List/Count.lean index 036a76b4b8..151169f74b 100644 --- a/Batteries/Data/List/Count.lean +++ b/Batteries/Data/List/Count.lean @@ -29,11 +29,3 @@ variable [DecidableEq α] theorem count_singleton' (a b : α) : count a [b] = if b = a then 1 else 0 := by simp [count_cons] theorem count_concat (a : α) (l : List α) : count a (concat l a) = succ (count a l) := by simp - -@[deprecated filter_eq (since := "2023-12-14")] -theorem filter_eq' (l : List α) (a : α) : l.filter (a = ·) = replicate (count a l) a := by - simpa only [eq_comm] using filter_eq l a - -@[deprecated filter_beq (since := "2023-12-14")] -theorem filter_beq' (l : List α) (a : α) : l.filter (a == ·) = replicate (count a l) a := by - simpa only [eq_comm (b := a)] using filter_eq l a diff --git a/Batteries/Data/List/Pairwise.lean b/Batteries/Data/List/Pairwise.lean index 9df947bf4c..de13a5e292 100644 --- a/Batteries/Data/List/Pairwise.lean +++ b/Batteries/Data/List/Pairwise.lean @@ -29,18 +29,6 @@ namespace List /-! ### Pairwise -/ -@[deprecated pairwise_iff_forall_sublist (since := "2023-09-18")] -theorem pairwise_of_reflexive_on_dupl_of_forall_ne [DecidableEq α] {l : List α} {r : α → α → Prop} - (hr : ∀ a, 1 < count a l → r a a) (h : ∀ a ∈ l, ∀ b ∈ l, a ≠ b → r a b) : l.Pairwise r := by - apply pairwise_iff_forall_sublist.mpr - intro a b hab - if heq : a = b then - cases heq; apply hr - rwa [show [a,a] = replicate 2 a from rfl, ← le_count_iff_replicate_sublist] at hab - else - apply h <;> try (apply hab.subset; simp) - exact heq - theorem pairwise_iff_get : Pairwise R l ↔ ∀ (i j) (_hij : i < j), R (get l i) (get l j) := by rw [pairwise_iff_getElem] constructor <;> intro h diff --git a/Batteries/Data/Nat/Lemmas.lean b/Batteries/Data/Nat/Lemmas.lean index 3492037135..7fa93705f7 100644 --- a/Batteries/Data/Nat/Lemmas.lean +++ b/Batteries/Data/Nat/Lemmas.lean @@ -149,24 +149,6 @@ protected def sum_trichotomy (a b : Nat) : a < b ⊕' a = b ⊕' b < a := | .eq => .inr (.inl (Nat.compare_eq_eq.1 h)) | .gt => .inr (.inr (Nat.compare_eq_gt.1 h)) -/-! ## add -/ - -@[deprecated (since := "2023-11-25")] alias succ_add_eq_succ_add := Nat.succ_add_eq_add_succ - -/-! ## sub -/ - -@[deprecated (since := "2023-11-25")] -protected alias le_of_le_of_sub_le_sub_right := Nat.le_of_sub_le_sub_right - -@[deprecated (since := "2023-11-25")] -protected alias le_of_le_of_sub_le_sub_left := Nat.le_of_sub_le_sub_left - -/-! ### mul -/ - -@[deprecated (since := "2024-01-11")] protected alias mul_lt_mul := Nat.mul_lt_mul_of_lt_of_le' - -@[deprecated (since := "2024-01-11")] protected alias mul_lt_mul' := Nat.mul_lt_mul_of_le_of_lt - /-! ### div/mod -/ -- TODO mod_core_congr, mod_def @@ -179,6 +161,3 @@ protected alias le_of_le_of_sub_le_sub_left := Nat.le_of_sub_le_sub_left @[simp] theorem sum_append : Nat.sum (l₁ ++ l₂) = Nat.sum l₁ + Nat.sum l₂ := by induction l₁ <;> simp [*, Nat.add_assoc] - -@[deprecated (since := "2024-03-05")] protected alias lt_connex := Nat.lt_or_gt_of_ne -@[deprecated (since := "2024-02-09")] alias pow_two_pos := Nat.two_pow_pos diff --git a/Batteries/Data/Option.lean b/Batteries/Data/Option.lean deleted file mode 100644 index e99c38ee11..0000000000 --- a/Batteries/Data/Option.lean +++ /dev/null @@ -1 +0,0 @@ -import Batteries.Data.Option.Lemmas diff --git a/Batteries/Data/Option/Lemmas.lean b/Batteries/Data/Option/Lemmas.lean deleted file mode 100644 index 05a38b25a7..0000000000 --- a/Batteries/Data/Option/Lemmas.lean +++ /dev/null @@ -1,13 +0,0 @@ -/- -Copyright (c) 2017 Mario Carneiro. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro --/ -import Batteries.Tactic.Alias - -namespace Option - -@[deprecated (since := "2024-03-05")] alias to_list_some := toList_some -@[deprecated (since := "2024-03-05")] alias to_list_none := toList_none - -end Option diff --git a/scripts/check_imports.lean b/scripts/check_imports.lean index 55bf0e8342..bc651dcc12 100644 --- a/scripts/check_imports.lean +++ b/scripts/check_imports.lean @@ -34,7 +34,7 @@ def warn (fixable : Bool) (msg : String) : LogIO Unit := do -- | Predicate indicates if warnings are present and if they fixable. def getWarningInfo : LogIO (Bool × Bool) := get -def createModuleHashmap (env : Environment) : HashMap Name ModuleData := Id.run do +def createModuleHashmap (env : Environment) : Std.HashMap Name ModuleData := Id.run do let mut nameMap := {} for i in [0:env.header.moduleNames.size] do let nm := env.header.moduleNames[i]! @@ -80,11 +80,11 @@ def checkMissingImports (modName : Name) (modData : ModuleData) (reqImports : Ar /-- Check directory entry in `Batteries/Data/` -/ def checkBatteriesDataDir - (modMap : HashMap Name ModuleData) + (modMap : Std.HashMap Name ModuleData) (entry : IO.FS.DirEntry) (autofix : Bool := false) : LogIO Unit := do let moduleName := `Batteries.Data ++ .mkSimple entry.fileName let requiredImports ← addModulesIn (recurse := true) #[] (root := moduleName) entry.path - let .some module := modMap.find? moduleName + let .some module := modMap[moduleName]? | warn true s!"Could not find {moduleName}; Not imported into Batteries." let path := modulePath moduleName -- We refuse to generate imported modules whose path doesn't exist. @@ -134,7 +134,7 @@ def checkBatteriesDataImports : MetaM Unit := do if ← entry.path.isDir then checkBatteriesDataDir (autofix := autofix) modMap entry let batteriesImports ← expectedBatteriesImports - let .some batteriesMod := modMap.find? `Batteries + let .some batteriesMod := modMap[`Batteries]? | warn false "Missing Batteries module!; Run `lake build`." let warned ← checkMissingImports `Batteries batteriesMod batteriesImports if autofix && warned then From 2ce0037d487217469a1efeb9ea8196fe15ab9c46 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 18 Sep 2024 16:44:55 +1000 Subject: [PATCH 079/177] chore: deprecate HashSet.insert' (#949) --- Batteries/Lean/HashSet.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Batteries/Lean/HashSet.lean b/Batteries/Lean/HashSet.lean index 7882358c40..4b6df398b3 100644 --- a/Batteries/Lean/HashSet.lean +++ b/Batteries/Lean/HashSet.lean @@ -54,7 +54,7 @@ instance : BEq (HashSet α) where `O(1)` amortized. Similar to `insert`, but also returns a Boolean flag indicating whether an existing entry has been replaced with `a => b`. -/ -@[inline] +@[inline, deprecated containsThenInsert (since := "2024-09-17")] def insert' (s : HashSet α) (a : α) : HashSet α × Bool := let oldSize := s.size let s := s.insert a From 35d1cd731ad832c9f1d860c4d8ec1c7c3ab96823 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 23 Sep 2024 21:16:42 +1000 Subject: [PATCH 080/177] feat: lemmas about Vector (#952) --- Batteries/Data/Vector/Basic.lean | 2 +- Batteries/Data/Vector/Lemmas.lean | 25 +++++++++++++++++++++++++ 2 files changed, 26 insertions(+), 1 deletion(-) diff --git a/Batteries/Data/Vector/Basic.lean b/Batteries/Data/Vector/Basic.lean index 6fc8166b6d..d8b11a5e75 100644 --- a/Batteries/Data/Vector/Basic.lean +++ b/Batteries/Data/Vector/Basic.lean @@ -80,7 +80,7 @@ def get (v : Vector α n) (i : Fin n) : α := v.toArray.get <| i.cast v.size_eq. /-- Vector lookup function that takes an index `i` of type `USize` -/ def uget (v : Vector α n) (i : USize) (h : i.toNat < n) : α := v.toArray.uget i (v.size_eq.symm ▸ h) -/-- `Vector α n` nstance for the `GetElem` typeclass. -/ +/-- `Vector α n` instance for the `GetElem` typeclass. -/ instance : GetElem (Vector α n) Nat α fun _ i => i < n where getElem := fun x i h => get x ⟨i, h⟩ diff --git a/Batteries/Data/Vector/Lemmas.lean b/Batteries/Data/Vector/Lemmas.lean index 02acccb8ce..572ec9a79d 100644 --- a/Batteries/Data/Vector/Lemmas.lean +++ b/Batteries/Data/Vector/Lemmas.lean @@ -8,6 +8,7 @@ import Batteries.Data.Vector.Basic import Batteries.Data.List.Basic import Batteries.Data.List.Lemmas import Batteries.Data.Array.Lemmas +import Batteries.Tactic.Lint.Simp /-! ## Vectors @@ -45,3 +46,27 @@ protected theorem ext {a b : Vector α n} (h : (i : Nat) → (_ : i < n) → a[i · intro i hi _ rw [a.size_eq] at hi exact h i hi + +@[simp] theorem getElem_mk {data : Array α} {size : data.size = n} {i : Nat} (h : i < n) : + (Vector.mk data size)[i] = data[i] := rfl + +@[simp] theorem push_mk {data : Array α} {size : data.size = n} {x : α} : + (Vector.mk data size).push x = + Vector.mk (data.push x) (by simp [size, Nat.succ_eq_add_one]) := rfl + +@[simp] theorem pop_mk {data : Array α} {size : data.size = n} : + (Vector.mk data size).pop = Vector.mk data.pop (by simp [size]) := rfl + +@[simp] theorem getElem_push_last {v : Vector α n} {x : α} : (v.push x)[n] = x := by + rcases v with ⟨data, rfl⟩ + simp + +-- The `simpNF` linter incorrectly claims that this lemma can not be applied by `simp`. +@[simp, nolint simpNF] theorem getElem_push_lt {v : Vector α n} {x : α} {i : Nat} (h : i < n) : + (v.push x)[i] = v[i] := by + rcases v with ⟨data, rfl⟩ + simp [Array.get_push_lt, h] + +@[simp] theorem getElem_pop {v : Vector α n} {i : Nat} (h : i < n - 1) : (v.pop)[i] = v[i] := by + rcases v with ⟨data, rfl⟩ + simp From c3817c4a0eb086efe5dbad137d72e340d3324bfd Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Tue, 24 Sep 2024 10:05:43 -0400 Subject: [PATCH 081/177] chore: fix more Std -> Batteries (#887) Co-authored-by: Matthew Robert Ballard <100034030+mattrobball@users.noreply.github.com> Co-authored-by: Matthew Ballard --- Batteries/CodeAction/Attr.lean | 2 +- Batteries/CodeAction/Basic.lean | 2 +- Batteries/CodeAction/Deprecated.lean | 4 +++- Batteries/CodeAction/Misc.lean | 2 +- Batteries/Tactic/Alias.lean | 2 +- Batteries/Tactic/Lint/Basic.lean | 2 +- Batteries/Tactic/Lint/Frontend.lean | 22 +++++++++++----------- Batteries/Tactic/Lint/Misc.lean | 4 ++-- Batteries/Tactic/Lint/Simp.lean | 4 +--- Batteries/Tactic/Lint/TypeClass.lean | 2 +- scripts/runLinter.lean | 2 +- test/lintTC.lean | 2 +- test/lintsimp.lean | 2 +- 13 files changed, 26 insertions(+), 26 deletions(-) diff --git a/Batteries/CodeAction/Attr.lean b/Batteries/CodeAction/Attr.lean index 79802d6932..308c4c7ad3 100644 --- a/Batteries/CodeAction/Attr.lean +++ b/Batteries/CodeAction/Attr.lean @@ -14,7 +14,7 @@ import Lean.Server.CodeActions.Basic * Attribute `@[tactic_code_action]` collects code actions which will be called on each occurrence of a tactic. -/ -namespace Std.CodeAction +namespace Batteries.CodeAction open Lean Elab Server Lsp RequestM Snapshots diff --git a/Batteries/CodeAction/Basic.lean b/Batteries/CodeAction/Basic.lean index 378f6aa00d..120de7eb05 100644 --- a/Batteries/CodeAction/Basic.lean +++ b/Batteries/CodeAction/Basic.lean @@ -19,7 +19,7 @@ on each occurrence of a hole (`_`, `?_` or `sorry`). attempt to use this code action provider when browsing the `Batteries.CodeAction.Hole.Attr` file itself.) -/ -namespace Std.CodeAction +namespace Batteries.CodeAction open Lean Elab Server RequestM CodeAction diff --git a/Batteries/CodeAction/Deprecated.lean b/Batteries/CodeAction/Deprecated.lean index 33d1d75243..30745d161d 100644 --- a/Batteries/CodeAction/Deprecated.lean +++ b/Batteries/CodeAction/Deprecated.lean @@ -14,7 +14,9 @@ This is an opt-in mechanism for making machine-applicable `@[deprecated]` defini whenever the deprecation lint also fires, allowing the user to replace the usage of the deprecated constant. -/ -namespace Std + +namespace Batteries + open Lean Elab Server Lsp RequestM CodeAction /-- An environment extension for identifying `@[deprecated]` definitions which can be auto-fixed -/ diff --git a/Batteries/CodeAction/Misc.lean b/Batteries/CodeAction/Misc.lean index 0bea77ff55..4ad0a8099d 100644 --- a/Batteries/CodeAction/Misc.lean +++ b/Batteries/CodeAction/Misc.lean @@ -15,7 +15,7 @@ import Lean.Server.CodeActions.Provider This declares some basic tactic code actions, using the `@[tactic_code_action]` API. -/ -namespace Std.CodeAction +namespace Batteries.CodeAction open Lean Meta Elab Server RequestM CodeAction diff --git a/Batteries/Tactic/Alias.lean b/Batteries/Tactic/Alias.lean index 07b91dc76e..471dfd883a 100644 --- a/Batteries/Tactic/Alias.lean +++ b/Batteries/Tactic/Alias.lean @@ -19,7 +19,7 @@ an iff theorem. namespace Batteries.Tactic.Alias -open Lean Elab Parser.Command Std +open Lean Elab Parser.Command /-- An alias can be in one of three forms -/ inductive AliasInfo where diff --git a/Batteries/Tactic/Lint/Basic.lean b/Batteries/Tactic/Lint/Basic.lean index 948bab5456..a0fceb7fa2 100644 --- a/Batteries/Tactic/Lint/Basic.lean +++ b/Batteries/Tactic/Lint/Basic.lean @@ -9,7 +9,7 @@ import Lean.Elab.Exception open Lean Meta -namespace Std.Tactic.Lint +namespace Batteries.Tactic.Lint /-! # Basic linter types and attributes diff --git a/Batteries/Tactic/Lint/Frontend.lean b/Batteries/Tactic/Lint/Frontend.lean index 27ff3bf3b2..cb0369db5f 100644 --- a/Batteries/Tactic/Lint/Frontend.lean +++ b/Batteries/Tactic/Lint/Frontend.lean @@ -39,7 +39,7 @@ You can append `only name1 name2 ...` to any command to run a subset of linters, You can add custom linters by defining a term of type `Linter` with the `@[env_linter]` attribute. -A linter defined with the name `Std.Tactic.Lint.myNewCheck` can be run with `#lint myNewCheck` +A linter defined with the name `Batteries.Tactic.Lint.myNewCheck` can be run with `#lint myNewCheck` or `#lint only myNewCheck`. If you add the attribute `@[env_linter disabled]` to `linter.myNewCheck` it will be registered, but not run by default. @@ -52,8 +52,8 @@ omits it from only the specified linter checks. sanity check, lint, cleanup, command, tactic -/ -namespace Std.Tactic.Lint -open Lean Std +namespace Batteries.Tactic.Lint +open Lean /-- Verbosity for the linter output. -/ inductive LintVerbosity @@ -97,7 +97,7 @@ Runs all the specified linters on all the specified declarations in parallel, producing a list of results. -/ def lintCore (decls : Array Name) (linters : Array NamedLinter) : - CoreM (Array (NamedLinter × HashMap Name MessageData)) := do + CoreM (Array (NamedLinter × Std.HashMap Name MessageData)) := do let env ← getEnv let options ← getOptions -- TODO: sanitize options? @@ -114,15 +114,15 @@ def lintCore (decls : Array Name) (linters : Array NamedLinter) : | Except.error err => pure m!"LINTER FAILED:\n{err.toMessageData}" tasks.mapM fun (linter, decls) => do - let mut msgs : HashMap Name MessageData := {} + let mut msgs : Std.HashMap Name MessageData := {} for (declName, msg?) in decls do if let some msg := msg?.get then msgs := msgs.insert declName msg pure (linter, msgs) /-- Sorts a map with declaration keys as names by line number. -/ -def sortResults (results : HashMap Name α) : CoreM <| Array (Name × α) := do - let mut key : HashMap Name Nat := {} +def sortResults (results : Std.HashMap Name α) : CoreM <| Array (Name × α) := do + let mut key : Std.HashMap Name Nat := {} for (n, _) in results.toArray do if let some range ← findDeclarationRanges? n then key := key.insert n <| range.range.pos.line @@ -140,7 +140,7 @@ def printWarning (declName : Name) (warning : MessageData) (useErrorFormat : Boo addMessageContextPartial m!"#check {← mkConstWithLevelParams declName} /- {warning} -/" /-- Formats a map of linter warnings using `print_warning`, sorted by line number. -/ -def printWarnings (results : HashMap Name MessageData) (filePath : System.FilePath := default) +def printWarnings (results : Std.HashMap Name MessageData) (filePath : System.FilePath := default) (useErrorFormat : Bool := false) : CoreM MessageData := do (MessageData.joinSep ·.toList Format.line) <$> (← sortResults results).mapM fun (declName, warning) => @@ -150,10 +150,10 @@ def printWarnings (results : HashMap Name MessageData) (filePath : System.FilePa Formats a map of linter warnings grouped by filename with `-- filename` comments. The first `drop_fn_chars` characters are stripped from the filename. -/ -def groupedByFilename (results : HashMap Name MessageData) (useErrorFormat : Bool := false) : +def groupedByFilename (results : Std.HashMap Name MessageData) (useErrorFormat : Bool := false) : CoreM MessageData := do let sp ← if useErrorFormat then initSrcSearchPath ["."] else pure {} - let grouped : HashMap Name (System.FilePath × HashMap Name MessageData) ← + let grouped : Std.HashMap Name (System.FilePath × Std.HashMap Name MessageData) ← results.foldM (init := {}) fun grouped declName msg => do let mod ← findModuleOf? declName let mod := mod.getD (← getEnv).mainModule @@ -174,7 +174,7 @@ def groupedByFilename (results : HashMap Name MessageData) (useErrorFormat : Boo Formats the linter results as Lean code with comments and `#check` commands. -/ def formatLinterResults - (results : Array (NamedLinter × HashMap Name MessageData)) + (results : Array (NamedLinter × Std.HashMap Name MessageData)) (decls : Array Name) (groupByFilename : Bool) (whereDesc : String) (runSlowLinters : Bool) diff --git a/Batteries/Tactic/Lint/Misc.lean b/Batteries/Tactic/Lint/Misc.lean index a0aa8bc884..075c139aea 100644 --- a/Batteries/Tactic/Lint/Misc.lean +++ b/Batteries/Tactic/Lint/Misc.lean @@ -14,7 +14,7 @@ import Batteries.Tactic.Lint.Basic open Lean Meta Std -namespace Std.Tactic.Lint +namespace Batteries.Tactic.Lint /-! # Various linters @@ -143,7 +143,7 @@ In pseudo-mathematical form, this returns `{{p : parameter | p ∈ u} | (u : lev FIXME: We use `Array Name` instead of `HashSet Name`, since `HashSet` does not have an equality instance. It will ignore `nm₀.proof_i` declarations. -/ -private def univParamsGrouped (e : Expr) (nm₀ : Name) : HashSet (Array Name) := +private def univParamsGrouped (e : Expr) (nm₀ : Name) : Std.HashSet (Array Name) := runST fun σ => do let res ← ST.mkRef (σ := σ) {} e.forEach fun diff --git a/Batteries/Tactic/Lint/Simp.lean b/Batteries/Tactic/Lint/Simp.lean index 9196e4ab74..c6c0ad8828 100644 --- a/Batteries/Tactic/Lint/Simp.lean +++ b/Batteries/Tactic/Lint/Simp.lean @@ -9,7 +9,7 @@ import Batteries.Tactic.OpenPrivate import Batteries.Util.LibraryNote open Lean Meta -namespace Std.Tactic.Lint +namespace Batteries.Tactic.Lint /-! # Linter for simplification lemmas @@ -86,8 +86,6 @@ where | Trie.node vs children => children.foldl (init := arr ++ vs) fun arr (_, child) => trieElements arr child -open Std - /-- Add message `msg` to any errors thrown inside `k`. -/ def decorateError (msg : MessageData) (k : MetaM α) : MetaM α := do try k catch e => throw (.error e.getRef m!"{msg}\n{e.toMessageData}") diff --git a/Batteries/Tactic/Lint/TypeClass.lean b/Batteries/Tactic/Lint/TypeClass.lean index e6cf632f87..a025777d18 100644 --- a/Batteries/Tactic/Lint/TypeClass.lean +++ b/Batteries/Tactic/Lint/TypeClass.lean @@ -6,7 +6,7 @@ Authors: Gabriel Ebner import Lean.Meta.Instances import Batteries.Tactic.Lint.Basic -namespace Std.Tactic.Lint +namespace Batteries.Tactic.Lint open Lean Meta /-- diff --git a/scripts/runLinter.lean b/scripts/runLinter.lean index 3fab25f834..7b46106099 100644 --- a/scripts/runLinter.lean +++ b/scripts/runLinter.lean @@ -2,7 +2,7 @@ import Lean.Util.SearchPath import Batteries.Tactic.Lint import Batteries.Data.Array.Basic -open Lean Core Elab Command Std.Tactic.Lint +open Lean Core Elab Command Batteries.Tactic.Lint open System (FilePath) /-- The list of `nolints` pulled from the `nolints.json` file -/ diff --git a/test/lintTC.lean b/test/lintTC.lean index 1f0ade6987..2a5f6f2640 100644 --- a/test/lintTC.lean +++ b/test/lintTC.lean @@ -1,7 +1,7 @@ import Batteries.Tactic.Lint.TypeClass import Lean.Elab.Command -open Std.Tactic.Lint +open Batteries.Tactic.Lint namespace A diff --git a/test/lintsimp.lean b/test/lintsimp.lean index 0950578665..ad2a589247 100644 --- a/test/lintsimp.lean +++ b/test/lintsimp.lean @@ -1,6 +1,6 @@ import Batteries.Tactic.Lint -open Std.Tactic.Lint +open Batteries.Tactic.Lint set_option linter.missingDocs false def f : Nat := 0 From 6d5e1c81277e960372c94f19172440e39b3c5980 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 25 Sep 2024 08:39:16 +1000 Subject: [PATCH 082/177] chore: cleanup a breaking non-terminal simp on HashMap (#954) --- Batteries/Data/HashMap/WF.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/Batteries/Data/HashMap/WF.lean b/Batteries/Data/HashMap/WF.lean index 9f69994638..258723a89b 100644 --- a/Batteries/Data/HashMap/WF.lean +++ b/Batteries/Data/HashMap/WF.lean @@ -94,12 +94,14 @@ where .sum (source.data.map (·.toList.length)) + target.size := by unfold expand.go; split · next H => - refine (go (i+1) _ _ fun j hj => ?a).trans ?b <;> simp + refine (go (i+1) _ _ fun j hj => ?a).trans ?b · case a => + simp only [Array.data_length, Array.data_set] simp [List.getD_eq_getElem?_getD, List.getElem?_set, Option.map_eq_map]; split · cases source.data[j]? <;> rfl · next H => exact hs _ (Nat.lt_of_le_of_ne (Nat.le_of_lt_succ hj) (Ne.symm H)) · case b => + simp only [Array.data_length, Array.data_set, Array.get_eq_getElem, AssocList.foldl_eq] refine have ⟨l₁, l₂, h₁, _, eq⟩ := List.exists_of_set H; eq ▸ ?_ rw [h₁] simp only [Buckets.size_eq, List.map_append, List.map_cons, AssocList.toList, From 1745fbd2f61394a65fddb8323cd6bb23600926d5 Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 26 Sep 2024 00:32:37 +0200 Subject: [PATCH 083/177] fix: add trailing line-break in nolints.json (#955) --- scripts/runLinter.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/runLinter.lean b/scripts/runLinter.lean index 7b46106099..89976f0e47 100644 --- a/scripts/runLinter.lean +++ b/scripts/runLinter.lean @@ -15,7 +15,7 @@ def readJsonFile (α) [FromJson α] (path : System.FilePath) : IO α := do /-- Serialize the given value `a : α` to the file as JSON. -/ def writeJsonFile [ToJson α] (path : System.FilePath) (a : α) : IO Unit := - IO.FS.writeFile path <| toJson a |>.pretty + IO.FS.writeFile path <| toJson a |>.pretty.push '\n' /-- Usage: `runLinter [--update] [Batteries.Data.Nat.Basic]` From c57ab80c8dd20b345b29c81c446c78a6b3677d20 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 26 Sep 2024 00:34:21 +0200 Subject: [PATCH 084/177] refactor: avoid relying on rfl's behavior on ground terms (#832) --- Batteries/Data/Fin/Lemmas.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Batteries/Data/Fin/Lemmas.lean b/Batteries/Data/Fin/Lemmas.lean index bc1eb0d2c9..5ddde152de 100644 --- a/Batteries/Data/Fin/Lemmas.lean +++ b/Batteries/Data/Fin/Lemmas.lean @@ -40,14 +40,14 @@ theorem list_succ (n) : list (n+1) = 0 :: (list n).map Fin.succ := by theorem list_succ_last (n) : list (n+1) = (list n).map castSucc ++ [last n] := by rw [list_succ] induction n with - | zero => rfl + | zero => simp [last] | succ n ih => rw [list_succ, List.map_cons castSucc, ih] simp [Function.comp_def, succ_castSucc] theorem list_reverse (n) : (list n).reverse = (list n).map rev := by induction n with - | zero => rfl + | zero => simp [last] | succ n ih => conv => lhs; rw [list_succ_last] conv => rhs; rw [list_succ] From 98f2215707ae293a5612217dc29c05b515269512 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 26 Sep 2024 19:44:23 +1000 Subject: [PATCH 085/177] feat: compare List.ofFn and Array.ofFn (#956) Co-authored-by: F. G. Dorais Co-authored-by: Matthew Robert Ballard <100034030+mattrobball@users.noreply.github.com> --- Batteries/Data/Array.lean | 1 + Batteries/Data/Array/OfFn.lean | 23 ++++++++++++++++ Batteries/Data/List.lean | 1 + Batteries/Data/List/Basic.lean | 23 +--------------- Batteries/Data/List/Lemmas.lean | 18 ++++++++++++ Batteries/Data/List/OfFn.lean | 49 +++++++++++++++++++++++++++++++++ 6 files changed, 93 insertions(+), 22 deletions(-) create mode 100644 Batteries/Data/Array/OfFn.lean create mode 100644 Batteries/Data/List/OfFn.lean diff --git a/Batteries/Data/Array.lean b/Batteries/Data/Array.lean index 4d41e75755..1f500b700a 100644 --- a/Batteries/Data/Array.lean +++ b/Batteries/Data/Array.lean @@ -4,4 +4,5 @@ import Batteries.Data.Array.Lemmas import Batteries.Data.Array.Match import Batteries.Data.Array.Merge import Batteries.Data.Array.Monadic +import Batteries.Data.Array.OfFn import Batteries.Data.Array.Pairwise diff --git a/Batteries/Data/Array/OfFn.lean b/Batteries/Data/Array/OfFn.lean new file mode 100644 index 0000000000..a5d3bc78bc --- /dev/null +++ b/Batteries/Data/Array/OfFn.lean @@ -0,0 +1,23 @@ +/- +Copyright (c) 2024 Lean FRO. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kim Morrison +-/ +import Batteries.Data.List.OfFn + +open List + +namespace Array + +/-! ### ofFn -/ + +@[simp] +theorem data_ofFn (f : Fin n → α) : (ofFn f).data = List.ofFn f := by + ext1 + simp only [getElem?_eq, data_length, size_ofFn, length_ofFn, getElem_ofFn] + split + · rw [← getElem_eq_data_getElem] + simp + · rfl + +end Array diff --git a/Batteries/Data/List.lean b/Batteries/Data/List.lean index 998160b904..f93f90a6c2 100644 --- a/Batteries/Data/List.lean +++ b/Batteries/Data/List.lean @@ -3,5 +3,6 @@ import Batteries.Data.List.Count import Batteries.Data.List.EraseIdx import Batteries.Data.List.Init.Lemmas import Batteries.Data.List.Lemmas +import Batteries.Data.List.OfFn import Batteries.Data.List.Pairwise import Batteries.Data.List.Perm diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index 963920017c..c2d76ed0e6 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -601,28 +601,7 @@ def sigmaTR {σ : α → Type _} (l₁ : List α) (l₂ : ∀ a, List (σ a)) : ofFn f = [f 0, f 1, ... , f (n - 1)] ``` -/ -def ofFn {n} (f : Fin n → α) : List α := go n 0 rfl where - /-- Auxiliary for `List.ofFn`. `ofFn.go f i j _ = [f j, ..., f (n - 1)]`. -/ - -- This used to be defined via `Array.ofFn` but mathlib relies on reducing it, - -- so we use a structurally recursive definition here. - go : (i j : Nat) → (h : i + j = n) → List α - | 0, _, _ => [] - | i+1, j, h => f ⟨j, by omega⟩ :: go i (j+1) (Nat.add_right_comm .. ▸ h :) - -/-- Tail-recursive version of `ofFn`. -/ -@[inline] def ofFnTR {n} (f : Fin n → α) : List α := go n (Nat.le_refl _) [] where - /-- Auxiliary for `List.ofFnTR`. `ofFnTR.go f i _ acc = f 0 :: ... :: f (i - 1) :: acc`. -/ - go : (i : Nat) → (h : i ≤ n) → List α → List α - | 0, _, acc => acc - | i+1, h, acc => go i (Nat.le_of_lt h) (f ⟨i, h⟩ :: acc) - -@[csimp] theorem ofFn_eq_ofFnTR : @ofFn = @ofFnTR := by - funext α n f; simp [ofFnTR] - let rec go (i j h h') : ofFnTR.go f j h' (ofFn.go f i j h) = ofFn f := by - unfold ofFnTR.go; split - · subst h; rfl - · next l j h' => exact go (i+1) j ((Nat.succ_add ..).trans h) (Nat.le_of_lt h') - exact (go 0 n (Nat.zero_add _) (Nat.le_refl _)).symm +def ofFn {n} (f : Fin n → α) : List α := Fin.foldr n (f · :: ·) [] /-- `ofFnNthVal f i` returns `some (f i)` if `i < n` and `none` otherwise. -/ def ofFnNthVal {n} (f : Fin n → α) (i : Nat) : Option α := diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index a68d820a52..9cc0ca1d95 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -17,6 +17,24 @@ open Nat @[simp] theorem mem_toArray {a : α} {l : List α} : a ∈ l.toArray ↔ a ∈ l := by simp [Array.mem_def] +/-! ### toArray-/ + +@[simp] theorem size_toArrayAux (l : List α) (r : Array α) : + (l.toArrayAux r).size = r.size + l.length := by + induction l generalizing r with + | nil => simp [toArrayAux] + | cons a l ih => + simp [ih, List.toArrayAux] + omega + +@[simp] theorem getElem_mk {xs : List α} {i : Nat} (h : i < xs.length) : + (Array.mk xs)[i] = xs[i] := rfl + +@[simp] theorem getElem_toArray (l : List α) (i : Nat) (h : i < l.toArray.size) : + l.toArray[i] = l[i]'(by simpa using h) := by + rw [Array.getElem_eq_data_getElem] + simp + /-! ### next? -/ @[simp] theorem next?_nil : @next? α [] = none := rfl diff --git a/Batteries/Data/List/OfFn.lean b/Batteries/Data/List/OfFn.lean new file mode 100644 index 0000000000..93214cc99b --- /dev/null +++ b/Batteries/Data/List/OfFn.lean @@ -0,0 +1,49 @@ +/- +Copyright (c) 2024 Lean FRO. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kim Morrison +-/ +import Batteries.Data.List.Lemmas +import Batteries.Data.Fin.Lemmas + +/-! +# Theorems about `List.ofFn` +-/ + +namespace List + +@[simp] +theorem length_ofFn (f : Fin n → α) : (ofFn f).length = n := by + simp only [ofFn] + induction n with + | zero => simp + | succ n ih => simp [Fin.foldr_succ, ih] + +@[simp] +protected theorem getElem_ofFn (f : Fin n → α) (i : Nat) (h : i < (ofFn f).length) : + (ofFn f)[i] = f ⟨i, by simp_all⟩ := by + simp only [ofFn] + induction n generalizing i with + | zero => simp at h + | succ n ih => + match i with + | 0 => simp [Fin.foldr_succ] + | i+1 => + simp only [Fin.foldr_succ] + apply ih + simp_all + +@[simp] +protected theorem getElem?_ofFn (f : Fin n → α) (i) : (ofFn f)[i]? = ofFnNthVal f i := + if h : i < (ofFn f).length + then by + rw [getElem?_eq_getElem h, List.getElem_ofFn] + · simp only [length_ofFn] at h; simp [ofFnNthVal, h] + else by + rw [ofFnNthVal, dif_neg] <;> + simpa using h + +@[simp] theorem toArray_ofFn (f : Fin n → α) : (ofFn f).toArray = Array.ofFn f := by + ext <;> simp + +end List From 40d378f10d013d4ec275bb5d364cce672aa87113 Mon Sep 17 00:00:00 2001 From: Seppel3210 <34406239+Seppel3210@users.noreply.github.com> Date: Thu, 26 Sep 2024 16:15:33 +0200 Subject: [PATCH 086/177] feat: fill in proof of Array.data_erase (#690) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: François G. Dorais --- Batteries/Data/Array/Lemmas.lean | 57 +++++++++++++++++++++++++++++++- Batteries/Data/List/Lemmas.lean | 56 +++++++++++++++++++++++++++++++ 2 files changed, 112 insertions(+), 1 deletion(-) diff --git a/Batteries/Data/Array/Lemmas.lean b/Batteries/Data/Array/Lemmas.lean index 02293aa349..869c35b37e 100644 --- a/Batteries/Data/Array/Lemmas.lean +++ b/Batteries/Data/Array/Lemmas.lean @@ -115,9 +115,64 @@ theorem mem_join : ∀ {L : Array (Array α)}, a ∈ L.join ↔ ∃ l, l ∈ L · rintro ⟨s, h₁, h₂⟩ refine ⟨s.data, ⟨⟨s, h₁, rfl⟩, h₂⟩⟩ +/-! ### indexOf? -/ + +theorem indexOf?_data [BEq α] {a : α} {l : Array α} : + l.data.indexOf? a = (l.indexOf? a).map Fin.val := by + simpa using aux l 0 +where + aux (l : Array α) (i : Nat) : + ((l.data.drop i).indexOf? a).map (·+i) = (indexOfAux l a i).map Fin.val := by + rw [indexOfAux] + if h : i < l.size then + rw [List.drop_eq_getElem_cons h, ←getElem_eq_data_getElem, List.indexOf?_cons] + if h' : l[i] == a then + simp [h, h'] + else + simp [h, h', ←aux l (i+1), Function.comp_def, ←Nat.add_assoc, Nat.add_right_comm] + else + have h' : l.size ≤ i := Nat.le_of_not_lt h + simp [h, List.drop_of_length_le h', List.indexOf?] + termination_by l.size - i + /-! ### erase -/ -@[simp] proof_wanted data_erase [BEq α] {l : Array α} {a : α} : (l.erase a).data = l.data.erase a +theorem eraseIdx_data_swap {l : Array α} (i : Nat) (lt : i + 1 < size l) : + (l.swap ⟨i+1, lt⟩ ⟨i, Nat.lt_of_succ_lt lt⟩).data.eraseIdx (i+1) = l.data.eraseIdx i := by + let ⟨xs⟩ := l + induction i generalizing xs <;> let x₀::x₁::xs := xs + case zero => simp [swap, get] + case succ i ih _ => + have lt' := Nat.lt_of_succ_lt_succ lt + have : (swap ⟨x₀::x₁::xs⟩ ⟨i.succ + 1, lt⟩ ⟨i.succ, Nat.lt_of_succ_lt lt⟩).data + = x₀::(swap ⟨x₁::xs⟩ ⟨i + 1, lt'⟩ ⟨i, Nat.lt_of_succ_lt lt'⟩).data := by + simp [swap_def, getElem_eq_data_getElem] + simp [this, ih] + +@[simp] theorem data_feraseIdx {l : Array α} (i : Fin l.size) : + (l.feraseIdx i).data = l.data.eraseIdx i := by + induction l, i using feraseIdx.induct with + | @case1 a i lt a' i' ih => + rw [feraseIdx] + simp [lt, ih, a', eraseIdx_data_swap i lt] + | case2 a i lt => + have : i + 1 ≥ a.size := Nat.ge_of_not_lt lt + have last : i + 1 = a.size := Nat.le_antisymm i.is_lt this + simp [feraseIdx, lt, List.dropLast_eq_eraseIdx last] + +@[simp] theorem data_erase [BEq α] (l : Array α) (a : α) : (l.erase a).data = l.data.erase a := by + match h : indexOf? l a with + | none => + simp only [erase, h] + apply Eq.symm + apply List.erase_of_forall_bne + rw [←List.indexOf?_eq_none_iff, indexOf?_data, h, Option.map_none'] + | some i => + simp only [erase, h] + rw [data_feraseIdx, ←List.eraseIdx_indexOf_eq_erase] + congr + rw [List.indexOf_eq_indexOf?, indexOf?_data] + simp [h] /-! ### shrink -/ diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 9cc0ca1d95..d211cb04f8 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -40,6 +40,22 @@ open Nat @[simp] theorem next?_nil : @next? α [] = none := rfl @[simp] theorem next?_cons (a l) : @next? α (a :: l) = some (a, l) := rfl +/-! ### dropLast -/ + +theorem dropLast_eq_eraseIdx {xs : List α} {i : Nat} (last_idx : i + 1 = xs.length) : + xs.dropLast = List.eraseIdx xs i := by + induction i generalizing xs with + | zero => + let [x] := xs + rfl + | succ n ih => + let x::xs := xs + simp at last_idx + rw [dropLast, eraseIdx] + congr + exact ih last_idx + exact fun _ => nomatch xs + /-! ### get? -/ @[deprecated getElem_eq_iff (since := "2024-06-12")] @@ -228,6 +244,25 @@ theorem get?_set_of_lt' (a : α) {m n} (l : List α) (h : m < length l) : @[deprecated (since := "2024-04-22")] alias sublist.erase := Sublist.erase +theorem erase_of_forall_bne [BEq α] (a : α) (xs : List α) (h : ∀ (x : α), x ∈ xs → ¬x == a) : + xs.erase a = xs := by + rw [erase_eq_eraseP', eraseP_of_forall_not h] + +-- TODO a version of the above theorem with LawfulBEq and ∉ + +/-! ### findIdx? -/ + +theorem findIdx_eq_findIdx? (p : α → Bool) (l : List α) : + l.findIdx p = (match l.findIdx? p with | some i => i | none => l.length) := by + induction l with + | nil => rfl + | cons x xs ih => + rw [findIdx_cons, findIdx?_cons] + if h : p x then + simp [h] + else + cases h' : findIdx? p xs <;> simp [h, h', ih] + /-! ### replaceF -/ theorem replaceF_nil : [].replaceF p = [] := rfl @@ -571,6 +606,14 @@ theorem indexesOf_cons [BEq α] : (x :: xs : List α).indexesOf y = bif x == y then 0 :: (xs.indexesOf y).map (· + 1) else (xs.indexesOf y).map (· + 1) := by simp [indexesOf, findIdxs_cons] +@[simp] theorem eraseIdx_indexOf_eq_erase [BEq α] (a : α) (l : List α) : + l.eraseIdx (l.indexOf a) = l.erase a := by + induction l with + | nil => rfl + | cons x xs ih => + rw [List.erase, indexOf_cons] + cases x == a <;> simp [ih] + theorem indexOf_mem_indexesOf [BEq α] [LawfulBEq α] {xs : List α} (m : x ∈ xs) : xs.indexOf x ∈ xs.indexesOf x := by induction xs with @@ -585,6 +628,19 @@ theorem indexOf_mem_indexesOf [BEq α] [LawfulBEq α] {xs : List α} (m : x ∈ specialize ih m simpa +@[simp] theorem indexOf?_nil [BEq α] : ([] : List α).indexOf? x = none := rfl +theorem indexOf?_cons [BEq α] : + (x :: xs : List α).indexOf? y = if x == y then some 0 else (xs.indexOf? y).map Nat.succ := by + simp [indexOf?] + +theorem indexOf?_eq_none_iff [BEq α] {a : α} {l : List α} : + l.indexOf? a = none ↔ ∀ x ∈ l, ¬x == a := by + simp [indexOf?, findIdx?_eq_none_iff] + +theorem indexOf_eq_indexOf? [BEq α] (a : α) (l : List α) : + l.indexOf a = (match l.indexOf? a with | some i => i | none => l.length) := by + simp [indexOf, indexOf?, findIdx_eq_findIdx?] + /-! ### insertP -/ theorem insertP_loop (a : α) (l r : List α) : From 27c99fb6e7a1088b1746fe5a63ef1be077b9dd0b Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 27 Sep 2024 11:15:30 +1000 Subject: [PATCH 087/177] feat: automatically test Mathlib against Batteries PRs (#958) --- .github/workflows/test_mathlib.yml | 75 ++++++++++++++++++++++++++++++ 1 file changed, 75 insertions(+) create mode 100644 .github/workflows/test_mathlib.yml diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml new file mode 100644 index 0000000000..97bbd9a792 --- /dev/null +++ b/.github/workflows/test_mathlib.yml @@ -0,0 +1,75 @@ +# Test Mathlib against a Batteries PR + +name: Test Mathlib + +on: + workflow_run: # https://docs.github.com/en/actions/using-workflows/events-that-trigger-workflows#workflow_run + workflows: [ci] + types: [completed] + +jobs: + on-success: + runs-on: ubuntu-latest + if: github.event.workflow_run.conclusion == 'success' && github.event.workflow_run.event == 'pull_request' && github.repository == 'leanprover-community/batteries' + steps: + - name: Retrieve information about the original workflow + uses: potiuk/get-workflow-origin@v1_1 # https://github.com/marketplace/actions/get-workflow-origin + # This action is deprecated and archived, but it seems hard to find a better solution for getting the PR number + # see https://github.com/orgs/community/discussions/25220 for some discussion + id: workflow-info + with: + token: ${{ secrets.GITHUB_TOKEN }} + sourceRunId: ${{ github.event.workflow_run.id }} + + # We automatically create a Mathlib branch using this Batteries branch. + # Mathlib CI will be responsible for reporting back success or failure + # to the PR comments asynchronously. + + # Checkout the mathlib4 repository with all branches + - name: Checkout mathlib4 repository + if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.ready.outputs.mathlib_ready == 'true' + uses: actions/checkout@v4 + with: + repository: leanprover-community/mathlib4 + token: ${{ secrets.MATHLIB4_BOT }} + ref: master + fetch-depth: 0 # This ensures we check out all tags and branches. + + - name: install elan + run: | + set -o pipefail + curl -sSfL https://github.com/leanprover/elan/releases/download/v3.0.0/elan-x86_64-unknown-linux-gnu.tar.gz | tar xz + ./elan-init -y --default-toolchain none + echo "$HOME/.elan/bin" >> "${GITHUB_PATH}" + + - name: Check if tag exists + if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.ready.outputs.mathlib_ready == 'true' + id: check_mathlib_tag + run: | + git config user.name "leanprover-community-mathlib4-bot" + git config user.email "leanprover-community-mathlib4-bot@users.noreply.github.com" + + BASE=master + echo "Using base tag: $BASE" + + EXISTS="$(git ls-remote --heads origin batteries-pr-testing-${{ steps.workflow-info.outputs.pullRequestNumber }} | wc -l)" + echo "Branch exists: $EXISTS" + if [ "$EXISTS" = "0" ]; then + echo "Branch does not exist, creating it." + git switch -c batteries-pr-testing-${{ steps.workflow-info.outputs.pullRequestNumber }} "$BASE" + sed -i 's,require "leanprover-community" / "batteries" @ git ".\+",require "leanprover-community" / "batteries" @ git "refs/pull/${{ steps.workflow-info.outputs.pullRequestNumber }}/head",' lakefile.lean + lake update batteries + git add lakefile.lean lake-manifest.json + git commit -m "Update Batteries branch for testing https://github.com/leanprover-community/batteries/pull/${{ steps.workflow-info.outputs.pullRequestNumber }}" + else + echo "Branch already exists, merging $BASE and bumping Batteries." + git switch batteries-pr-testing-${{ steps.workflow-info.outputs.pullRequestNumber }} + git merge "$BASE" --strategy-option ours --no-commit --allow-unrelated-histories + lake update batteries + git commit --allow-empty -m "Trigger CI for https://github.com/leanprover-community/batteries/pull/${{ steps.workflow-info.outputs.pullRequestNumber }}" + fi + + - name: Push changes + if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.ready.outputs.mathlib_ready == 'true' + run: | + git push origin batteries-pr-testing-${{ steps.workflow-info.outputs.pullRequestNumber }} From 3ad208b73483726912b6694b5b72d5396ceb6967 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 27 Sep 2024 14:08:25 +1000 Subject: [PATCH 088/177] chore: fix test_mathlib.yml (#959) --- .github/workflows/test_mathlib.yml | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index 97bbd9a792..7c6690f21d 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -27,7 +27,7 @@ jobs: # Checkout the mathlib4 repository with all branches - name: Checkout mathlib4 repository - if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.ready.outputs.mathlib_ready == 'true' + if: steps.workflow-info.outputs.pullRequestNumber != '' uses: actions/checkout@v4 with: repository: leanprover-community/mathlib4 @@ -43,7 +43,7 @@ jobs: echo "$HOME/.elan/bin" >> "${GITHUB_PATH}" - name: Check if tag exists - if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.ready.outputs.mathlib_ready == 'true' + if: steps.workflow-info.outputs.pullRequestNumber != '' id: check_mathlib_tag run: | git config user.name "leanprover-community-mathlib4-bot" @@ -70,6 +70,6 @@ jobs: fi - name: Push changes - if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.ready.outputs.mathlib_ready == 'true' + if: steps.workflow-info.outputs.pullRequestNumber != '' run: | git push origin batteries-pr-testing-${{ steps.workflow-info.outputs.pullRequestNumber }} From 949f8827e1147a5579fcf4b6468719703c847f4d Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 27 Sep 2024 14:28:17 +1000 Subject: [PATCH 089/177] chore: fix test_mathlib.yml again (#961) --- .github/workflows/test_mathlib.yml | 22 +++++++++++++++++++++- 1 file changed, 21 insertions(+), 1 deletion(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index 7c6690f21d..f41593d601 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -42,6 +42,23 @@ jobs: ./elan-init -y --default-toolchain none echo "$HOME/.elan/bin" >> "${GITHUB_PATH}" + - name: Retrieve PR information + if: steps.workflow-info.outputs.pullRequestNumber != '' + id: pr-info + uses: actions/github-script@v6 + with: + script: | + const prNumber = ${{ steps.workflow-info.outputs.pullRequestNumber }}; + const pr = await github.pulls.get({ + owner: 'leanprover-community', + repo: 'batteries', + pull_number: prNumber + }); + return { + headRepo: pr.data.head.repo.full_name, + headBranch: pr.data.head.ref + }; + - name: Check if tag exists if: steps.workflow-info.outputs.pullRequestNumber != '' id: check_mathlib_tag @@ -57,7 +74,10 @@ jobs: if [ "$EXISTS" = "0" ]; then echo "Branch does not exist, creating it." git switch -c batteries-pr-testing-${{ steps.workflow-info.outputs.pullRequestNumber }} "$BASE" - sed -i 's,require "leanprover-community" / "batteries" @ git ".\+",require "leanprover-community" / "batteries" @ git "refs/pull/${{ steps.workflow-info.outputs.pullRequestNumber }}/head",' lakefile.lean + + # Use the fork and branch name to modify the lakefile.lean + sed -i 's,require "leanprover-community" / "batteries" @ git ".\+",require "leanprover-community" / "batteries" @ git "${{ steps.pr-info.outputs.headBranch }}" from git "https://github.com/${{ steps.pr-info.outputs.headRepo }}",' lakefile.lean + lake update batteries git add lakefile.lean lake-manifest.json git commit -m "Update Batteries branch for testing https://github.com/leanprover-community/batteries/pull/${{ steps.workflow-info.outputs.pullRequestNumber }}" From d41ca4d21d5174cb3ef7f59ba585269b7d157bc2 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 27 Sep 2024 14:35:16 +1000 Subject: [PATCH 090/177] chore: fix test_mathlib.yml (3) (#962) --- .github/workflows/test_mathlib.yml | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index f41593d601..0e3782c354 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -49,14 +49,14 @@ jobs: with: script: | const prNumber = ${{ steps.workflow-info.outputs.pullRequestNumber }}; - const pr = await github.pulls.get({ - owner: 'leanprover-community', - repo: 'batteries', + const { data: pr } = await github.rest.pulls.get({ + owner: context.repo.owner, + repo: context.repo.repo, pull_number: prNumber }); return { - headRepo: pr.data.head.repo.full_name, - headBranch: pr.data.head.ref + headRepo: pr.head.repo.full_name, + headBranch: pr.head.ref }; - name: Check if tag exists From d3d6925545b56facace36ef39b376d05f45113b7 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 27 Sep 2024 14:52:37 +1000 Subject: [PATCH 091/177] chore: fix test_mathlib.yml (4) (#963) --- .github/workflows/test_mathlib.yml | 46 +++++++++++++++--------------- 1 file changed, 23 insertions(+), 23 deletions(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index 0e3782c354..cd831183a1 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -3,7 +3,7 @@ name: Test Mathlib on: - workflow_run: # https://docs.github.com/en/actions/using-workflows/events-that-trigger-workflows#workflow_run + workflow_run: workflows: [ci] types: [completed] @@ -13,19 +13,12 @@ jobs: if: github.event.workflow_run.conclusion == 'success' && github.event.workflow_run.event == 'pull_request' && github.repository == 'leanprover-community/batteries' steps: - name: Retrieve information about the original workflow - uses: potiuk/get-workflow-origin@v1_1 # https://github.com/marketplace/actions/get-workflow-origin - # This action is deprecated and archived, but it seems hard to find a better solution for getting the PR number - # see https://github.com/orgs/community/discussions/25220 for some discussion + uses: potiuk/get-workflow-origin@v1_1 id: workflow-info with: token: ${{ secrets.GITHUB_TOKEN }} sourceRunId: ${{ github.event.workflow_run.id }} - # We automatically create a Mathlib branch using this Batteries branch. - # Mathlib CI will be responsible for reporting back success or failure - # to the PR comments asynchronously. - - # Checkout the mathlib4 repository with all branches - name: Checkout mathlib4 repository if: steps.workflow-info.outputs.pullRequestNumber != '' uses: actions/checkout@v4 @@ -33,7 +26,7 @@ jobs: repository: leanprover-community/mathlib4 token: ${{ secrets.MATHLIB4_BOT }} ref: master - fetch-depth: 0 # This ensures we check out all tags and branches. + fetch-depth: 0 - name: install elan run: | @@ -46,50 +39,57 @@ jobs: if: steps.workflow-info.outputs.pullRequestNumber != '' id: pr-info uses: actions/github-script@v6 + env: + PR_NUMBER: ${{ steps.workflow-info.outputs.pullRequestNumber }} with: script: | - const prNumber = ${{ steps.workflow-info.outputs.pullRequestNumber }}; + const prNumber = process.env.PR_NUMBER; const { data: pr } = await github.rest.pulls.get({ owner: context.repo.owner, repo: context.repo.repo, pull_number: prNumber }); - return { - headRepo: pr.head.repo.full_name, - headBranch: pr.head.ref - }; + core.exportVariable('HEAD_REPO', pr.head.repo.full_name); + core.exportVariable('HEAD_BRANCH', pr.head.ref); - name: Check if tag exists if: steps.workflow-info.outputs.pullRequestNumber != '' id: check_mathlib_tag + env: + PR_NUMBER: ${{ steps.workflow-info.outputs.pullRequestNumber }} + HEAD_REPO: ${{ env.HEAD_REPO }} + HEAD_BRANCH: ${{ env.HEAD_BRANCH }} run: | git config user.name "leanprover-community-mathlib4-bot" git config user.email "leanprover-community-mathlib4-bot@users.noreply.github.com" + echo "PR info: $HEAD_REPO $HEAD_BRANCH" + BASE=master echo "Using base tag: $BASE" - EXISTS="$(git ls-remote --heads origin batteries-pr-testing-${{ steps.workflow-info.outputs.pullRequestNumber }} | wc -l)" + EXISTS="$(git ls-remote --heads origin batteries-pr-testing-$PR_NUMBER | wc -l)" echo "Branch exists: $EXISTS" if [ "$EXISTS" = "0" ]; then echo "Branch does not exist, creating it." - git switch -c batteries-pr-testing-${{ steps.workflow-info.outputs.pullRequestNumber }} "$BASE" + git switch -c batteries-pr-testing-$PR_NUMBER "$BASE" - # Use the fork and branch name to modify the lakefile.lean - sed -i 's,require "leanprover-community" / "batteries" @ git ".\+",require "leanprover-community" / "batteries" @ git "${{ steps.pr-info.outputs.headBranch }}" from git "https://github.com/${{ steps.pr-info.outputs.headRepo }}",' lakefile.lean + sed -i 's,require "leanprover-community" / "batteries" @ git ".\+",require "leanprover-community" / "batteries" from git "https://github.com/$HEAD_REPO" @ "$HEAD_BRANCH",' lakefile.lean lake update batteries git add lakefile.lean lake-manifest.json - git commit -m "Update Batteries branch for testing https://github.com/leanprover-community/batteries/pull/${{ steps.workflow-info.outputs.pullRequestNumber }}" + git commit -m "Update Batteries branch for testing https://github.com/leanprover-community/batteries/pull/$PR_NUMBER" else echo "Branch already exists, merging $BASE and bumping Batteries." - git switch batteries-pr-testing-${{ steps.workflow-info.outputs.pullRequestNumber }} + git switch batteries-pr-testing-$PR_NUMBER git merge "$BASE" --strategy-option ours --no-commit --allow-unrelated-histories lake update batteries - git commit --allow-empty -m "Trigger CI for https://github.com/leanprover-community/batteries/pull/${{ steps.workflow-info.outputs.pullRequestNumber }}" + git commit --allow-empty -m "Trigger CI for https://github.com/leanprover-community/batteries/pull/$PR_NUMBER" fi - name: Push changes if: steps.workflow-info.outputs.pullRequestNumber != '' + env: + PR_NUMBER: ${{ steps.workflow-info.outputs.pullRequestNumber }} run: | - git push origin batteries-pr-testing-${{ steps.workflow-info.outputs.pullRequestNumber }} + git push origin batteries-pr-testing-$PR_NUMBER From e023a8173c15fa0b065780dd81b675e59ccf3e13 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 27 Sep 2024 15:00:22 +1000 Subject: [PATCH 092/177] chore: fix test_mathlib.yml (5) (#964) --- .github/workflows/test_mathlib.yml | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index cd831183a1..d6ae30146b 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -74,7 +74,8 @@ jobs: echo "Branch does not exist, creating it." git switch -c batteries-pr-testing-$PR_NUMBER "$BASE" - sed -i 's,require "leanprover-community" / "batteries" @ git ".\+",require "leanprover-community" / "batteries" from git "https://github.com/$HEAD_REPO" @ "$HEAD_BRANCH",' lakefile.lean + # Use the fork and branch name to modify the lakefile.lean + sed -i "s,require \"leanprover-community\" / \"batteries\" @ git \".\+\",require \"leanprover-community\" / \"batteries\" from git \"https://github.com/$HEAD_REPO\" @ \"$HEAD_BRANCH\",g" lakefile.lean lake update batteries git add lakefile.lean lake-manifest.json From e0b13c946e9c3805f1eec785c72955e103a9cbaf Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 27 Sep 2024 16:33:31 +1000 Subject: [PATCH 093/177] chore: add lake-manifest.json when updating batteries-pr-testing branches (#965) --- .github/workflows/test_mathlib.yml | 1 + 1 file changed, 1 insertion(+) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index d6ae30146b..15062cb521 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -85,6 +85,7 @@ jobs: git switch batteries-pr-testing-$PR_NUMBER git merge "$BASE" --strategy-option ours --no-commit --allow-unrelated-histories lake update batteries + git add lake-manifest.json git commit --allow-empty -m "Trigger CI for https://github.com/leanprover-community/batteries/pull/$PR_NUMBER" fi From 8311edc805fe01405a1a28de83ec9bfddeb3aefa Mon Sep 17 00:00:00 2001 From: Seppel3210 <34406239+Seppel3210@users.noreply.github.com> Date: Sat, 28 Sep 2024 17:14:03 +0200 Subject: [PATCH 094/177] fix: remove List.erase_of_forall_bne and replace with List.erase_eq_self_iff_forall_bne (#967) --- Batteries/Data/Array/Lemmas.lean | 4 ++-- Batteries/Data/List/Lemmas.lean | 8 +++----- 2 files changed, 5 insertions(+), 7 deletions(-) diff --git a/Batteries/Data/Array/Lemmas.lean b/Batteries/Data/Array/Lemmas.lean index 869c35b37e..d384a334ab 100644 --- a/Batteries/Data/Array/Lemmas.lean +++ b/Batteries/Data/Array/Lemmas.lean @@ -165,8 +165,8 @@ theorem eraseIdx_data_swap {l : Array α} (i : Nat) (lt : i + 1 < size l) : | none => simp only [erase, h] apply Eq.symm - apply List.erase_of_forall_bne - rw [←List.indexOf?_eq_none_iff, indexOf?_data, h, Option.map_none'] + rw [List.erase_eq_self_iff_forall_bne, ←List.indexOf?_eq_none_iff, indexOf?_data, + h, Option.map_none'] | some i => simp only [erase, h] rw [data_feraseIdx, ←List.eraseIdx_indexOf_eq_erase] diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index d211cb04f8..5df9338026 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -244,11 +244,9 @@ theorem get?_set_of_lt' (a : α) {m n} (l : List α) (h : m < length l) : @[deprecated (since := "2024-04-22")] alias sublist.erase := Sublist.erase -theorem erase_of_forall_bne [BEq α] (a : α) (xs : List α) (h : ∀ (x : α), x ∈ xs → ¬x == a) : - xs.erase a = xs := by - rw [erase_eq_eraseP', eraseP_of_forall_not h] - --- TODO a version of the above theorem with LawfulBEq and ∉ +theorem erase_eq_self_iff_forall_bne [BEq α] (a : α) (xs : List α) : + xs.erase a = xs ↔ ∀ (x : α), x ∈ xs → ¬x == a := by + rw [erase_eq_eraseP', eraseP_eq_self_iff] /-! ### findIdx? -/ From 51c38e3828d06c82741a7a65df93611e2ce209e1 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 29 Sep 2024 01:51:10 +1000 Subject: [PATCH 095/177] chore: simplify lakefile (#950) --- .github/workflows/build.yml | 2 +- lakefile.lean | 7 +------ 2 files changed, 2 insertions(+), 7 deletions(-) diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 15d72c6876..61af30e96e 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -20,7 +20,7 @@ jobs: name: build, test, and lint batteries uses: leanprover/lean-action@v1 with: - build-args: '-Kwerror' + build-args: '--wfail' - name: Check that all files are imported run: lake env lean scripts/check_imports.lean diff --git a/lakefile.lean b/lakefile.lean index e55e8b2639..7c287c4c95 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -2,13 +2,8 @@ import Lake open Lake DSL -macro "opt_arg?" x:ident v:term : term => `(if get_config? $x |>.isSome then $v else default) - package batteries where - moreLeanArgs := opt_arg? werror #["-DwarningAsError=true"] - leanOptions := - #[⟨`linter.missingDocs, true⟩] ++ - opt_arg? disable_new_compiler #[⟨`compiler.enableNew, false⟩] + leanOptions := #[⟨`linter.missingDocs, true⟩] @[default_target] lean_lib Batteries From deb7e080f0d5e48de4477e1675fc5b743295ac14 Mon Sep 17 00:00:00 2001 From: FR Date: Sun, 29 Sep 2024 01:45:15 +0800 Subject: [PATCH 096/177] chore: use implicit arguments in iff lemmas (#957) Co-authored-by: F. G. Dorais --- Batteries/Logic.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Batteries/Logic.lean b/Batteries/Logic.lean index 03c82eb5ac..9d7d14a70d 100644 --- a/Batteries/Logic.lean +++ b/Batteries/Logic.lean @@ -79,25 +79,25 @@ theorem cast_eq_iff_heq : cast e a = a' ↔ HEq a a' := ⟨heq_of_cast_eq _, fun h => by cases h; rfl⟩ theorem eqRec_eq_cast {α : Sort _} {a : α} {motive : (a' : α) → a = a' → Sort _} - (x : motive a (rfl : a = a)) {a' : α} (e : a = a') : + (x : motive a rfl) {a' : α} (e : a = a') : @Eq.rec α a motive x a' e = cast (e ▸ rfl) x := by subst e; rfl --Porting note: new theorem. More general version of `eqRec_heq` theorem eqRec_heq_self {α : Sort _} {a : α} {motive : (a' : α) → a = a' → Sort _} - (x : motive a (rfl : a = a)) {a' : α} (e : a = a') : + (x : motive a rfl) {a' : α} (e : a = a') : HEq (@Eq.rec α a motive x a' e) x := by subst e; rfl @[simp] theorem eqRec_heq_iff_heq {α : Sort _} {a : α} {motive : (a' : α) → a = a' → Sort _} - (x : motive a (rfl : a = a)) {a' : α} (e : a = a') {β : Sort _} (y : β) : + {x : motive a rfl} {a' : α} {e : a = a'} {β : Sort _} {y : β} : HEq (@Eq.rec α a motive x a' e) y ↔ HEq x y := by subst e; rfl @[simp] theorem heq_eqRec_iff_heq {α : Sort _} {a : α} {motive : (a' : α) → a = a' → Sort _} - (x : motive a (rfl : a = a)) {a' : α} (e : a = a') {β : Sort _} (y : β) : + {x : motive a rfl} {a' : α} {e : a = a'} {β : Sort _} {y : β} : HEq y (@Eq.rec α a motive x a' e) ↔ HEq y x := by subst e; rfl From 82c92a6f46dff1c0c8a93a132d9f33953e4246ee Mon Sep 17 00:00:00 2001 From: Pim Otte Date: Sat, 28 Sep 2024 20:20:03 +0200 Subject: [PATCH 097/177] feat: getElem_tail lemmas (#905) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: François G. Dorais Co-authored-by: Kim Morrison --- Batteries/Data/List/Lemmas.lean | 11 +++++++++++ lake-manifest.json | 2 +- 2 files changed, 12 insertions(+), 1 deletion(-) diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 5df9338026..ce439cffc5 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -228,6 +228,17 @@ theorem get?_set_of_lt' (a : α) {m n} (l : List α) (h : m < length l) : @[deprecated (since := "2024-05-06")] alias length_removeNth := length_eraseIdx +/-! ### tail -/ + +theorem length_tail_add_one (l : List α) (h : 0 < length l) : (length (tail l)) + 1 = length l := by + simp [Nat.sub_add_cancel h] + +@[simp] theorem getElem?_tail (l : List α) : l.tail[n]? = l[n + 1]? := by cases l <;> simp + +@[simp] theorem getElem_tail (l : List α) (h : n < l.tail.length) : + l.tail[n] = l[n + 1]'(by simp at h; omega) := by + cases l; contradiction; simp + /-! ### eraseP -/ @[simp] theorem extractP_eq_find?_eraseP diff --git a/lake-manifest.json b/lake-manifest.json index 68d5a83837..b44d4b4a5f 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -1,4 +1,4 @@ -{"version": "1.0.0", +{"version": "1.1.0", "packagesDir": ".lake/packages", "packages": [], "name": "batteries", From c94610ae66955eabcdff619ab80ad1f172b18d3c Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 29 Sep 2024 06:50:10 +1000 Subject: [PATCH 098/177] feat: decidable quantifiers for `Vector` (#953) --- Batteries/Data/Vector/Basic.lean | 3 ++ Batteries/Data/Vector/Lemmas.lean | 56 +++++++++++++++++++++++++++++++ test/vector.lean | 25 ++++++++++++++ 3 files changed, 84 insertions(+) create mode 100644 test/vector.lean diff --git a/Batteries/Data/Vector/Basic.lean b/Batteries/Data/Vector/Basic.lean index d8b11a5e75..7726a5b9e5 100644 --- a/Batteries/Data/Vector/Basic.lean +++ b/Batteries/Data/Vector/Basic.lean @@ -102,6 +102,9 @@ if it exists. Else the vector is empty and it returns `none` -/ abbrev back? (v : Vector α n) : Option α := v[n - 1]? +/-- Abbreviation for the last element of a non-empty `Vector`.-/ +abbrev back (v : Vector α (n + 1)) : α := v[n] + /-- `Vector.head` produces the head of a vector -/ abbrev head (v : Vector α (n+1)) := v[0] diff --git a/Batteries/Data/Vector/Lemmas.lean b/Batteries/Data/Vector/Lemmas.lean index 572ec9a79d..0f8caa5c45 100644 --- a/Batteries/Data/Vector/Lemmas.lean +++ b/Batteries/Data/Vector/Lemmas.lean @@ -70,3 +70,59 @@ protected theorem ext {a b : Vector α n} (h : (i : Nat) → (_ : i < n) → a[i @[simp] theorem getElem_pop {v : Vector α n} {i : Nat} (h : i < n - 1) : (v.pop)[i] = v[i] := by rcases v with ⟨data, rfl⟩ simp + +/-- +Variant of `getElem_pop` that will sometimes fire when `getElem_pop` gets stuck because of +defeq issues in the implicit size argument. +-/ +@[simp] theorem getElem_pop' (v : Vector α (n + 1)) (i : Nat) (h : i < n + 1 - 1) : + @getElem (Vector α n) Nat α (fun _ i => i < n) instGetElemNatLt v.pop i h = v[i] := + getElem_pop h + +@[simp] theorem push_pop_back (v : Vector α (n + 1)) : v.pop.push v.back = v := by + ext i + by_cases h : i < n + · simp [h] + · replace h : i = n := by omega + subst h + simp + +/-! ### Decidable quantifiers. -/ + +theorem forall_zero_iff {P : Vector α 0 → Prop} : + (∀ v, P v) ↔ P .empty := by + constructor + · intro h + apply h + · intro h v + obtain (rfl : v = .empty) := (by ext i h; simp at h) + apply h + +theorem forall_cons_iff {P : Vector α (n + 1) → Prop} : + (∀ v : Vector α (n + 1), P v) ↔ (∀ (x : α) (v : Vector α n), P (v.push x)) := by + constructor + · intro h _ _ + apply h + · intro h v + have w : v = v.pop.push v.back := by simp + rw [w] + apply h + +instance instDecidableForallVectorZero (P : Vector α 0 → Prop) : + ∀ [Decidable (P .empty)], Decidable (∀ v, P v) + | .isTrue h => .isTrue fun ⟨v, s⟩ => by + obtain (rfl : v = .empty) := (by ext i h₁ h₂; exact s; cases h₂) + exact h + | .isFalse h => .isFalse (fun w => h (w _)) + +instance instDecidableForallVectorSucc (P : Vector α (n+1) → Prop) + [Decidable (∀ (x : α) (v : Vector α n), P (v.push x))] : Decidable (∀ v, P v) := + decidable_of_iff' (∀ x (v : Vector α n), P (v.push x)) forall_cons_iff + +instance instDecidableExistsVectorZero (P : Vector α 0 → Prop) [Decidable (P .empty)] : + Decidable (∃ v, P v) := + decidable_of_iff (¬ ∀ v, ¬ P v) Classical.not_forall_not + +instance instDecidableExistsVectorSucc (P : Vector α (n+1) → Prop) + [Decidable (∀ (x : α) (v : Vector α n), ¬ P (v.push x))] : Decidable (∃ v, P v) := + decidable_of_iff (¬ ∀ v, ¬ P v) Classical.not_forall_not diff --git a/test/vector.lean b/test/vector.lean new file mode 100644 index 0000000000..5a8b9e8587 --- /dev/null +++ b/test/vector.lean @@ -0,0 +1,25 @@ +import Batteries.Data.Vector + +/-! ### Testing decidable quantifiers for `Vector`. -/ + +open Batteries + +example : ∃ v : Vector Bool 6, v.toList.count true = 3 := by decide + +inductive Gate : Nat → Type +| const : Bool → Gate 0 +| if : ∀ {n}, Gate n → Gate n → Gate (n + 1) + +namespace Gate + +def and : Gate 2 := .if (.if (.const true) (.const false)) (.if (.const false) (.const false)) + +def eval (g : Gate n) (v : Vector Bool n) : Bool := + match g, v with + | .const b, _ => b + | .if g₁ g₂, v => if v.1.back then eval g₁ v.pop else eval g₂ v.pop + +example : ∀ v, and.eval v = (v[0] && v[1]) := by decide +example : ∃ v, and.eval v = false := by decide + +end Gate From bf12ff6041cbab6eba6b54d9467baed807bb2bfd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sat, 28 Sep 2024 19:22:37 -0400 Subject: [PATCH 099/177] fix: disable flaky test (#970) --- test/MLList.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/test/MLList.lean b/test/MLList.lean index ad53377237..bc7546df0a 100644 --- a/test/MLList.lean +++ b/test/MLList.lean @@ -23,6 +23,8 @@ We generate three tasks which sleep for `100`, `50`, and `1` milliseconds respec and then verify that `MLList.ofTaskList` return their results in the order they complete. -/ +/- This test is very flaky, so it's disabled for now. + def sleep (n : UInt32) : BaseIO (Task UInt32) := IO.asTask (do IO.sleep n; return n) |>.map fun t => t.map fun | .ok n => n @@ -35,3 +37,4 @@ def sleeps : MLList BaseIO UInt32 := .squash fun _ => do /-- info: [1, 50, 100] -/ #guard_msgs in #eval sleeps.force +-/ From 4756e0fc48acce0cc808df0ad149de5973240df6 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 1 Oct 2024 16:06:07 +1000 Subject: [PATCH 100/177] chore: update toolchain to v4.12.0 (#973) --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 98556ba065..89985206ac 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.12.0-rc1 +leanprover/lean4:v4.12.0 From 7815c9dbe98c295ad2df777d8af96c16366a3c0a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 3 Oct 2024 16:46:57 +1000 Subject: [PATCH 101/177] chore: cleanup some unused arguments (#974) --- Batteries/Classes/SatisfiesM.lean | 2 +- Batteries/Data/List/Basic.lean | 2 +- Batteries/Data/List/Lemmas.lean | 16 ++++++++-------- Batteries/Data/RBMap/Alter.lean | 4 ++-- Batteries/Logic.lean | 2 +- Batteries/WF.lean | 2 +- 6 files changed, 14 insertions(+), 14 deletions(-) diff --git a/Batteries/Classes/SatisfiesM.lean b/Batteries/Classes/SatisfiesM.lean index c8fc1f0be9..55ded469fc 100644 --- a/Batteries/Classes/SatisfiesM.lean +++ b/Batteries/Classes/SatisfiesM.lean @@ -52,7 +52,7 @@ protected theorem trivial [Applicative m] [LawfulApplicative m] {x : m α} : /-- The `SatisfiesM p x` predicate is monotonic in `p`. -/ theorem imp [Functor m] [LawfulFunctor m] {x : m α} (h : SatisfiesM p x) (H : ∀ {a}, p a → q a) : SatisfiesM q x := - let ⟨x, h⟩ := h; ⟨(fun ⟨a, h⟩ => ⟨_, H h⟩) <$> x, by rw [← h, ← comp_map]; rfl⟩ + let ⟨x, h⟩ := h; ⟨(fun ⟨_, h⟩ => ⟨_, H h⟩) <$> x, by rw [← h, ← comp_map]; rfl⟩ /-- `SatisfiesM` distributes over `<$>`, general version. -/ protected theorem map [Functor m] [LawfulFunctor m] {x : m α} diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index c2d76ed0e6..29061c409e 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -758,7 +758,7 @@ where | x::xs, n+1, acc => go m xs n (acc.push x) theorem dropSlice_zero₂ : ∀ n l, @dropSlice α n 0 l = l - | 0, [] | 0, _::_ | n+1, [] => rfl + | 0, [] | 0, _::_ | _+1, [] => rfl | n+1, x::xs => by simp [dropSlice, dropSlice_zero₂] @[csimp] theorem dropSlice_eq_dropSliceTR : @dropSlice = @dropSliceTR := by diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index ce439cffc5..f03e693d16 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -85,8 +85,8 @@ theorem modifyNthTail_id : ∀ n (l : List α), l.modifyNthTail id n = l theorem eraseIdx_eq_modifyNthTail : ∀ n (l : List α), eraseIdx l n = modifyNthTail tail n l | 0, l => by cases l <;> rfl - | n+1, [] => rfl - | n+1, a :: l => congrArg (cons _) (eraseIdx_eq_modifyNthTail _ _) + | _+1, [] => rfl + | _+1, _ :: _ => congrArg (cons _) (eraseIdx_eq_modifyNthTail _ _) @[deprecated (since := "2024-05-06")] alias removeNth_eq_nth_tail := eraseIdx_eq_modifyNthTail @@ -171,8 +171,8 @@ theorem modifyNth_eq_take_cons_drop (f : α → α) {n l} (h : n < length l) : theorem set_eq_modifyNth (a : α) : ∀ n (l : List α), set l n a = modifyNth (fun _ => a) n l | 0, l => by cases l <;> rfl - | n+1, [] => rfl - | n+1, b :: l => congrArg (cons _) (set_eq_modifyNth _ _ _) + | _+1, [] => rfl + | _+1, _ :: _ => congrArg (cons _) (set_eq_modifyNth _ _ _) theorem set_eq_take_cons_drop (a : α) {n l} (h : n < length l) : set l n a = take n l ++ a :: drop (n + 1) l := by @@ -181,7 +181,7 @@ theorem set_eq_take_cons_drop (a : α) {n l} (h : n < length l) : theorem modifyNth_eq_set_get? (f : α → α) : ∀ n (l : List α), l.modifyNth f n = ((fun a => l.set n (f a)) <$> l.get? n).getD l | 0, l => by cases l <;> rfl - | n+1, [] => rfl + | _+1, [] => rfl | n+1, b :: l => (congrArg (cons _) (modifyNth_eq_set_get? ..)).trans <| by cases h : l[n]? <;> simp [h] @@ -293,10 +293,10 @@ theorem replaceF_of_forall_none {l : List α} (h : ∀ a, a ∈ l → p a = none | nil => rfl | cons _ _ ih => simp [h _ (.head ..), ih (forall_mem_cons.1 h).2] -theorem exists_of_replaceF : ∀ {l : List α} {a a'} (al : a ∈ l) (pa : p a = some a'), +theorem exists_of_replaceF : ∀ {l : List α} {a a'} (_ : a ∈ l) (_ : p a = some a'), ∃ a a' l₁ l₂, (∀ b ∈ l₁, p b = none) ∧ p a = some a' ∧ l = l₁ ++ a :: l₂ ∧ l.replaceF p = l₁ ++ a' :: l₂ - | b :: l, a, a', al, pa => + | b :: l, _, _, al, pa => match pb : p b with | some b' => ⟨b, b', [], l, forall_mem_nil _, pb, by simp [pb]⟩ | none => @@ -489,7 +489,7 @@ theorem Sublist.diff_right : ∀ {l₁ l₂ l₃ : List α}, l₁ <+ l₂ → l theorem Sublist.erase_diff_erase_sublist {a : α} : ∀ {l₁ l₂ : List α}, l₁ <+ l₂ → (l₂.erase a).diff (l₁.erase a) <+ l₂.diff l₁ - | [], l₂, _ => erase_sublist _ _ + | [], _, _ => erase_sublist _ _ | b :: l₁, l₂, h => by if heq : b = a then simp [heq] diff --git a/Batteries/Data/RBMap/Alter.lean b/Batteries/Data/RBMap/Alter.lean index 1148221081..3f6d25c2b3 100644 --- a/Batteries/Data/RBMap/Alter.lean +++ b/Batteries/Data/RBMap/Alter.lean @@ -204,7 +204,7 @@ theorem _root_.Batteries.RBNode.Ordered.zoom {t : RBNode α} theorem Ordered.ins : ∀ {path : Path α} {t : RBNode α}, t.Ordered cmp → path.Ordered cmp → t.All (path.RootOrdered cmp) → (path.ins t).Ordered cmp - | .root, t, ht, _, _ => Ordered.setBlack.2 ht + | .root, _, ht, _, _ => Ordered.setBlack.2 ht | .left red parent x b, a, ha, ⟨hp, xb, xp, bp, hb⟩, H => by unfold ins; have ⟨ax, ap⟩ := All_and.1 H; exact hp.ins ⟨ax, xb, ha, hb⟩ ⟨xp, ap, bp⟩ | .right red a x parent, b, hb, ⟨hp, ax, xp, ap, ha⟩, H => by @@ -222,7 +222,7 @@ theorem Ordered.insertNew {path : Path α} (hp : path.Ordered cmp) (vp : path.Ro theorem Ordered.del : ∀ {path : Path α} {t : RBNode α} {c}, t.Ordered cmp → path.Ordered cmp → t.All (path.RootOrdered cmp) → (path.del t c).Ordered cmp - | .root, t, _, ht, _, _ => Ordered.setBlack.2 ht + | .root, _, _, ht, _, _ => Ordered.setBlack.2 ht | .left _ parent x b, a, red, ha, ⟨hp, xb, xp, bp, hb⟩, H => by unfold del; have ⟨ax, ap⟩ := All_and.1 H; exact hp.del ⟨ax, xb, ha, hb⟩ ⟨xp, ap, bp⟩ | .right _ a x parent, b, red, hb, ⟨hp, ax, xp, ap, ha⟩, H => by diff --git a/Batteries/Logic.lean b/Batteries/Logic.lean index 9d7d14a70d..f8fe5a6e38 100644 --- a/Batteries/Logic.lean +++ b/Batteries/Logic.lean @@ -31,7 +31,7 @@ end Classical theorem heq_iff_eq : HEq a b ↔ a = b := ⟨eq_of_heq, heq_of_eq⟩ @[simp] theorem eq_rec_constant {α : Sort _} {a a' : α} {β : Sort _} (y : β) (h : a = a') : - (@Eq.rec α a (fun α _ => β) y a' h) = y := by cases h; rfl + (@Eq.rec α a (fun _ _ => β) y a' h) = y := by cases h; rfl theorem congrArg₂ (f : α → β → γ) {x x' : α} {y y' : β} (hx : x = x') (hy : y = y') : f x y = f x' y' := by subst hx hy; rfl diff --git a/Batteries/WF.lean b/Batteries/WF.lean index 73428fe3bd..2c36dccebd 100644 --- a/Batteries/WF.lean +++ b/Batteries/WF.lean @@ -50,7 +50,7 @@ instance wfRel {r : α → α → Prop} : WellFoundedRelation { val // Acc r val (intro : (x : α) → (h : ∀ (y : α), r y x → Acc r y) → ((y : α) → (hr : r y x) → motive y (h y hr)) → motive x (intro x h)) {a : α} (t : Acc r a) : motive a t := - intro a (fun x h => t.inv h) (fun y hr => recC intro (t.inv hr)) + intro a (fun _ h => t.inv h) (fun _ hr => recC intro (t.inv hr)) termination_by Subtype.mk a t unseal recC From 34e690ec07f6f6375668adba5a16d0d723226c2c Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 3 Oct 2024 17:06:09 +1000 Subject: [PATCH 102/177] chore: only test Mathlib on PRs to main (#977) --- .github/workflows/test_mathlib.yml | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index 15062cb521..8eca882a4c 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -20,7 +20,7 @@ jobs: sourceRunId: ${{ github.event.workflow_run.id }} - name: Checkout mathlib4 repository - if: steps.workflow-info.outputs.pullRequestNumber != '' + if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.workflow-info.outputs.targetBranch == 'main' uses: actions/checkout@v4 with: repository: leanprover-community/mathlib4 @@ -29,6 +29,7 @@ jobs: fetch-depth: 0 - name: install elan + if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.workflow-info.outputs.targetBranch == 'main' run: | set -o pipefail curl -sSfL https://github.com/leanprover/elan/releases/download/v3.0.0/elan-x86_64-unknown-linux-gnu.tar.gz | tar xz @@ -36,7 +37,7 @@ jobs: echo "$HOME/.elan/bin" >> "${GITHUB_PATH}" - name: Retrieve PR information - if: steps.workflow-info.outputs.pullRequestNumber != '' + if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.workflow-info.outputs.targetBranch == 'main' id: pr-info uses: actions/github-script@v6 env: @@ -53,7 +54,7 @@ jobs: core.exportVariable('HEAD_BRANCH', pr.head.ref); - name: Check if tag exists - if: steps.workflow-info.outputs.pullRequestNumber != '' + if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.workflow-info.outputs.targetBranch == 'main' id: check_mathlib_tag env: PR_NUMBER: ${{ steps.workflow-info.outputs.pullRequestNumber }} @@ -90,7 +91,7 @@ jobs: fi - name: Push changes - if: steps.workflow-info.outputs.pullRequestNumber != '' + if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.workflow-info.outputs.targetBranch == 'main' env: PR_NUMBER: ${{ steps.workflow-info.outputs.pullRequestNumber }} run: | From f274aed7ae8d1addd3e70adaf3183ccc6e1ed43d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Thu, 3 Oct 2024 05:14:29 -0400 Subject: [PATCH 103/177] fix: argument order and make abbrev (#975) --- Batteries/Data/Array/Basic.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/Batteries/Data/Array/Basic.lean b/Batteries/Data/Array/Basic.lean index ad6b8129e6..f25b5e091b 100644 --- a/Batteries/Data/Array/Basic.lean +++ b/Batteries/Data/Array/Basic.lean @@ -148,7 +148,7 @@ should prove the index bound. A proof by `get_elem_tactic` is provided as a default argument for `h`. This will perform the update destructively provided that `a` has a reference count of 1 when called. -/ -def setN (a : Array α) (i : Nat) (h : i < a.size := by get_elem_tactic) (x : α) : Array α := +abbrev setN (a : Array α) (i : Nat) (x : α) (h : i < a.size := by get_elem_tactic) : Array α := a.set ⟨i, h⟩ x /-- @@ -157,7 +157,7 @@ Uses `get_elem_tactic` to supply a proof that the indices are in range. `hi` and `hj` are both given a default argument `by get_elem_tactic`. This will perform the update destructively provided that `a` has a reference count of 1 when called. -/ -def swapN (a : Array α) (i j : Nat) +abbrev swapN (a : Array α) (i j : Nat) (hi : i < a.size := by get_elem_tactic) (hj : j < a.size := by get_elem_tactic) : Array α := Array.swap a ⟨i,hi⟩ ⟨j, hj⟩ @@ -166,8 +166,8 @@ def swapN (a : Array α) (i j : Nat) The old entry is returned alongwith the modified vector. Automatically generates proof of `i < a.size` with `get_elem_tactic` where feasible. -/ -def swapAtN (a : Array α) (i : Nat) (h : i < a.size := by get_elem_tactic) (x : α) : α × Array α := - swapAt a ⟨i,h⟩ x +abbrev swapAtN (a : Array α) (i : Nat) (x : α) (h : i < a.size := by get_elem_tactic) : + α × Array α := swapAt a ⟨i,h⟩ x /-- `eraseIdxN a i h` Removes the element at position `i` from a vector of length `n`. @@ -176,7 +176,7 @@ that the index is valid. This function takes worst case O(n) time because it has to backshift all elements at positions greater than i. -/ -def eraseIdxN (a : Array α) (i : Nat) (h : i < a.size := by get_elem_tactic) : Array α := +abbrev eraseIdxN (a : Array α) (i : Nat) (h : i < a.size := by get_elem_tactic) : Array α := a.feraseIdx ⟨i, h⟩ end Array From fc871f7039ac6d8ab993335bb35aba43286004a0 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 3 Oct 2024 22:11:13 +1000 Subject: [PATCH 104/177] chore: move to v4.13.0-rc1 (#979) Co-authored-by: leanprover-community-mathlib4-bot Co-authored-by: Joachim Breitner Co-authored-by: Kyle Miller Co-authored-by: Matthew Ballard Co-authored-by: Jeremy Tan Jie Rui Co-authored-by: Mario Carneiro --- Batteries.lean | 2 - Batteries/Classes/SatisfiesM.lean | 4 +- Batteries/Data/Array/Basic.lean | 23 ---- Batteries/Data/Array/Lemmas.lean | 127 ++++++++-------------- Batteries/Data/Array/Merge.lean | 4 + Batteries/Data/Array/OfFn.lean | 9 +- Batteries/Data/Array/Pairwise.lean | 13 ++- Batteries/Data/AssocList.lean | 2 +- Batteries/Data/ByteArray.lean | 4 +- Batteries/Data/Fin/Basic.lean | 2 +- Batteries/Data/Fin/Lemmas.lean | 6 +- Batteries/Data/HashMap/Basic.lean | 6 +- Batteries/Data/HashMap/WF.lean | 104 +++++++++--------- Batteries/Data/List/Basic.lean | 37 ++++--- Batteries/Data/List/Lemmas.lean | 73 ++----------- Batteries/Data/List/Perm.lean | 8 +- Batteries/Data/RBMap/Alter.lean | 28 +++-- Batteries/Data/RBMap/WF.lean | 3 +- Batteries/Data/Rat/Basic.lean | 14 +-- Batteries/Data/Rat/Lemmas.lean | 16 +-- Batteries/Data/String/Lemmas.lean | 6 +- Batteries/Data/UnionFind/Basic.lean | 27 +++-- Batteries/Data/Vector/Basic.lean | 2 +- Batteries/Lean/HashSet.lean | 21 ---- Batteries/Lean/Meta/AssertHypotheses.lean | 40 ------- Batteries/Lean/Meta/Clear.lean | 26 ----- lean-toolchain | 2 +- test/lint_lean.lean | 8 +- 28 files changed, 217 insertions(+), 400 deletions(-) delete mode 100644 Batteries/Lean/Meta/AssertHypotheses.lean delete mode 100644 Batteries/Lean/Meta/Clear.lean diff --git a/Batteries.lean b/Batteries.lean index 3bd1f4ef79..f04a7a1f76 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -45,9 +45,7 @@ import Batteries.Lean.HashMap import Batteries.Lean.HashSet import Batteries.Lean.IO.Process import Batteries.Lean.Json -import Batteries.Lean.Meta.AssertHypotheses import Batteries.Lean.Meta.Basic -import Batteries.Lean.Meta.Clear import Batteries.Lean.Meta.DiscrTree import Batteries.Lean.Meta.Expr import Batteries.Lean.Meta.Inaccessible diff --git a/Batteries/Classes/SatisfiesM.lean b/Batteries/Classes/SatisfiesM.lean index 55ded469fc..89f264963a 100644 --- a/Batteries/Classes/SatisfiesM.lean +++ b/Batteries/Classes/SatisfiesM.lean @@ -171,7 +171,7 @@ theorem SatisfiesM_StateRefT_eq [Monad m] : · refine ⟨fun s => (fun ⟨⟨a, h⟩, s'⟩ => ⟨⟨a, s'⟩, h⟩) <$> f s, fun s => ?_⟩ rw [← comp_map, map_eq_pure_bind]; rfl · refine ⟨fun s => (fun ⟨⟨a, s'⟩, h⟩ => ⟨⟨a, h⟩, s'⟩) <$> f s, funext fun s => ?_⟩ - show _ >>= _ = _; simp [map_eq_pure_bind, ← h] + show _ >>= _ = _; simp [← h] @[simp] theorem SatisfiesM_ExceptT_eq [Monad m] [LawfulMonad m] : SatisfiesM (m := ExceptT ρ m) (α := α) p x ↔ SatisfiesM (m := m) (∀ a, · = .ok a → p a) x := by @@ -179,4 +179,4 @@ theorem SatisfiesM_StateRefT_eq [Monad m] : · exists (fun | .ok ⟨a, h⟩ => ⟨.ok a, fun | _, rfl => h⟩ | .error e => ⟨.error e, nofun⟩) <$> f show _ = _ >>= _; rw [← comp_map, map_eq_pure_bind]; congr; funext a; cases a <;> rfl · exists ((fun | ⟨.ok a, h⟩ => .ok ⟨a, h _ rfl⟩ | ⟨.error e, _⟩ => .error e) <$> f : m _) - show _ >>= _ = _; simp [← comp_map, map_eq_pure_bind]; congr; funext ⟨a, h⟩; cases a <;> rfl + show _ >>= _ = _; simp [← comp_map, ← bind_pure_comp]; congr; funext ⟨a, h⟩; cases a <;> rfl diff --git a/Batteries/Data/Array/Basic.lean b/Batteries/Data/Array/Basic.lean index f25b5e091b..7fb0cf8cf8 100644 --- a/Batteries/Data/Array/Basic.lean +++ b/Batteries/Data/Array/Basic.lean @@ -27,13 +27,6 @@ arrays, remove duplicates and then compare them elementwise. def equalSet [BEq α] (xs ys : Array α) : Bool := xs.all (ys.contains ·) && ys.all (xs.contains ·) -set_option linter.unusedVariables.funArgs false in -/-- -Sort an array using `compare` to compare elements. --/ -def qsortOrd [ord : Ord α] (xs : Array α) : Array α := - xs.qsort fun x y => compare x y |>.isLT - set_option linter.unusedVariables.funArgs false in /-- Returns the first minimal element among `d` and elements of the array. @@ -184,22 +177,6 @@ end Array namespace Subarray -/-- -The empty subarray. --/ -protected def empty : Subarray α where - array := #[] - start := 0 - stop := 0 - start_le_stop := Nat.le_refl 0 - stop_le_array_size := Nat.le_refl 0 - -instance : EmptyCollection (Subarray α) := - ⟨Subarray.empty⟩ - -instance : Inhabited (Subarray α) := - ⟨{}⟩ - /-- Check whether a subarray is empty. -/ diff --git a/Batteries/Data/Array/Lemmas.lean b/Batteries/Data/Array/Lemmas.lean index d384a334ab..3f28087212 100644 --- a/Batteries/Data/Array/Lemmas.lean +++ b/Batteries/Data/Array/Lemmas.lean @@ -11,30 +11,33 @@ import Batteries.Util.ProofWanted namespace Array -theorem forIn_eq_forIn_data [Monad m] +theorem forIn_eq_forIn_toList [Monad m] (as : Array α) (b : β) (f : α → β → m (ForInStep β)) : - forIn as b f = forIn as.data b f := by + forIn as b f = forIn as.toList b f := by let rec loop : ∀ {i h b j}, j + i = as.size → - Array.forIn.loop as f i h b = forIn (as.data.drop j) b f + Array.forIn.loop as f i h b = forIn (as.toList.drop j) b f | 0, _, _, _, rfl => by rw [List.drop_length]; rfl | i+1, _, _, j, ij => by simp only [forIn.loop, Nat.add] have j_eq : j = size as - 1 - i := by simp [← ij, ← Nat.add_assoc] have : as.size - 1 - i < as.size := j_eq ▸ ij ▸ Nat.lt_succ_of_le (Nat.le_add_right ..) - have : as[size as - 1 - i] :: as.data.drop (j + 1) = as.data.drop j := by + have : as[size as - 1 - i] :: as.toList.drop (j + 1) = as.toList.drop j := by rw [j_eq]; exact List.getElem_cons_drop _ _ this simp only [← this, List.forIn_cons]; congr; funext x; congr; funext b rw [loop (i := i)]; rw [← ij, Nat.succ_add]; rfl conv => lhs; simp only [forIn, Array.forIn] rw [loop (Nat.zero_add _)]; rfl + +@[deprecated (since := "2024-09-09")] alias forIn_eq_forIn_data := forIn_eq_forIn_toList @[deprecated (since := "2024-08-13")] alias forIn_eq_data_forIn := forIn_eq_forIn_data /-! ### zipWith / zip -/ -theorem data_zipWith (f : α → β → γ) (as : Array α) (bs : Array β) : - (as.zipWith bs f).data = as.data.zipWith f bs.data := by +theorem toList_zipWith (f : α → β → γ) (as : Array α) (bs : Array β) : + (as.zipWith bs f).toList = as.toList.zipWith f bs.toList := by let rec loop : ∀ (i : Nat) cs, i ≤ as.size → i ≤ bs.size → - (zipWithAux f as bs i cs).data = cs.data ++ (as.data.drop i).zipWith f (bs.data.drop i) := by + (zipWithAux f as bs i cs).toList = + cs.toList ++ (as.toList.drop i).zipWith f (bs.toList.drop i) := by intro i cs hia hib unfold zipWithAux by_cases h : i = as.size ∨ i = bs.size @@ -59,27 +62,30 @@ theorem data_zipWith (f : α → β → γ) (as : Array α) (bs : Array β) : have has : i < as.size := Nat.lt_of_le_of_ne hia h.1 have hbs : i < bs.size := Nat.lt_of_le_of_ne hib h.2 simp only [has, hbs, dite_true] - rw [loop (i+1) _ has hbs, Array.push_data] + rw [loop (i+1) _ has hbs, Array.push_toList] have h₁ : [f as[i] bs[i]] = List.zipWith f [as[i]] [bs[i]] := rfl - let i_as : Fin as.data.length := ⟨i, has⟩ - let i_bs : Fin bs.data.length := ⟨i, hbs⟩ + let i_as : Fin as.toList.length := ⟨i, has⟩ + let i_bs : Fin bs.toList.length := ⟨i, hbs⟩ rw [h₁, List.append_assoc] congr - rw [← List.zipWith_append (h := by simp), getElem_eq_data_getElem, getElem_eq_data_getElem] - show List.zipWith f (as.data[i_as] :: List.drop (i_as + 1) as.data) - ((List.get bs.data i_bs) :: List.drop (i_bs + 1) bs.data) = - List.zipWith f (List.drop i as.data) (List.drop i bs.data) - simp only [data_length, Fin.getElem_fin, List.getElem_cons_drop, List.get_eq_getElem] + rw [← List.zipWith_append (h := by simp), getElem_eq_getElem_toList, + getElem_eq_getElem_toList] + show List.zipWith f (as.toList[i_as] :: List.drop (i_as + 1) as.toList) + ((List.get bs.toList i_bs) :: List.drop (i_bs + 1) bs.toList) = + List.zipWith f (List.drop i as.toList) (List.drop i bs.toList) + simp only [length_toList, Fin.getElem_fin, List.getElem_cons_drop, List.get_eq_getElem] simp [zipWith, loop 0 #[] (by simp) (by simp)] +@[deprecated (since := "2024-09-09")] alias data_zipWith := toList_zipWith @[deprecated (since := "2024-08-13")] alias zipWith_eq_zipWith_data := data_zipWith theorem size_zipWith (as : Array α) (bs : Array β) (f : α → β → γ) : (as.zipWith bs f).size = min as.size bs.size := by - rw [size_eq_length_data, data_zipWith, List.length_zipWith] + rw [size_eq_length_toList, toList_zipWith, List.length_zipWith] -theorem data_zip (as : Array α) (bs : Array β) : - (as.zip bs).data = as.data.zip bs.data := - data_zipWith Prod.mk as bs +theorem toList_zip (as : Array α) (bs : Array β) : + (as.zip bs).toList = as.toList.zip bs.toList := + toList_zipWith Prod.mk as bs +@[deprecated (since := "2024-09-09")] alias data_zip := toList_zip @[deprecated (since := "2024-08-13")] alias zip_eq_zip_data := data_zip theorem size_zip (as : Array α) (bs : Array β) : @@ -90,42 +96,43 @@ theorem size_zip (as : Array α) (bs : Array β) : theorem size_filter_le (p : α → Bool) (l : Array α) : (l.filter p).size ≤ l.size := by - simp only [← data_length, filter_data] + simp only [← length_toList, toList_filter] apply List.length_filter_le /-! ### join -/ -@[simp] theorem data_join {l : Array (Array α)} : l.join.data = (l.data.map data).join := by +@[simp] theorem toList_join {l : Array (Array α)} : l.join.toList = (l.toList.map toList).join := by dsimp [join] - simp only [foldl_eq_foldl_data] - generalize l.data = l - have : ∀ a : Array α, (List.foldl ?_ a l).data = a.data ++ ?_ := ?_ + simp only [foldl_eq_foldl_toList] + generalize l.toList = l + have : ∀ a : Array α, (List.foldl ?_ a l).toList = a.toList ++ ?_ := ?_ exact this #[] induction l with | nil => simp - | cons h => induction h.data <;> simp [*] + | cons h => induction h.toList <;> simp [*] +@[deprecated (since := "2024-09-09")] alias data_join := toList_join @[deprecated (since := "2024-08-13")] alias join_data := data_join theorem mem_join : ∀ {L : Array (Array α)}, a ∈ L.join ↔ ∃ l, l ∈ L ∧ a ∈ l := by - simp only [mem_def, data_join, List.mem_join, List.mem_map] + simp only [mem_def, toList_join, List.mem_join, List.mem_map] intro l constructor · rintro ⟨_, ⟨s, m, rfl⟩, h⟩ exact ⟨s, m, h⟩ · rintro ⟨s, h₁, h₂⟩ - refine ⟨s.data, ⟨⟨s, h₁, rfl⟩, h₂⟩⟩ + refine ⟨s.toList, ⟨⟨s, h₁, rfl⟩, h₂⟩⟩ /-! ### indexOf? -/ -theorem indexOf?_data [BEq α] {a : α} {l : Array α} : - l.data.indexOf? a = (l.indexOf? a).map Fin.val := by +theorem indexOf?_toList [BEq α] {a : α} {l : Array α} : + l.toList.indexOf? a = (l.indexOf? a).map Fin.val := by simpa using aux l 0 where aux (l : Array α) (i : Nat) : - ((l.data.drop i).indexOf? a).map (·+i) = (indexOfAux l a i).map Fin.val := by + ((l.toList.drop i).indexOf? a).map (·+i) = (indexOfAux l a i).map Fin.val := by rw [indexOfAux] if h : i < l.size then - rw [List.drop_eq_getElem_cons h, ←getElem_eq_data_getElem, List.indexOf?_cons] + rw [List.drop_eq_getElem_cons h, ←getElem_eq_getElem_toList, List.indexOf?_cons] if h' : l[i] == a then simp [h, h'] else @@ -137,42 +144,8 @@ where /-! ### erase -/ -theorem eraseIdx_data_swap {l : Array α} (i : Nat) (lt : i + 1 < size l) : - (l.swap ⟨i+1, lt⟩ ⟨i, Nat.lt_of_succ_lt lt⟩).data.eraseIdx (i+1) = l.data.eraseIdx i := by - let ⟨xs⟩ := l - induction i generalizing xs <;> let x₀::x₁::xs := xs - case zero => simp [swap, get] - case succ i ih _ => - have lt' := Nat.lt_of_succ_lt_succ lt - have : (swap ⟨x₀::x₁::xs⟩ ⟨i.succ + 1, lt⟩ ⟨i.succ, Nat.lt_of_succ_lt lt⟩).data - = x₀::(swap ⟨x₁::xs⟩ ⟨i + 1, lt'⟩ ⟨i, Nat.lt_of_succ_lt lt'⟩).data := by - simp [swap_def, getElem_eq_data_getElem] - simp [this, ih] - -@[simp] theorem data_feraseIdx {l : Array α} (i : Fin l.size) : - (l.feraseIdx i).data = l.data.eraseIdx i := by - induction l, i using feraseIdx.induct with - | @case1 a i lt a' i' ih => - rw [feraseIdx] - simp [lt, ih, a', eraseIdx_data_swap i lt] - | case2 a i lt => - have : i + 1 ≥ a.size := Nat.ge_of_not_lt lt - have last : i + 1 = a.size := Nat.le_antisymm i.is_lt this - simp [feraseIdx, lt, List.dropLast_eq_eraseIdx last] - -@[simp] theorem data_erase [BEq α] (l : Array α) (a : α) : (l.erase a).data = l.data.erase a := by - match h : indexOf? l a with - | none => - simp only [erase, h] - apply Eq.symm - rw [List.erase_eq_self_iff_forall_bne, ←List.indexOf?_eq_none_iff, indexOf?_data, - h, Option.map_none'] - | some i => - simp only [erase, h] - rw [data_feraseIdx, ←List.eraseIdx_indexOf_eq_erase] - congr - rw [List.indexOf_eq_indexOf?, indexOf?_data] - simp [h] +@[simp] proof_wanted toList_erase [BEq α] {l : Array α} {a : α} : + (l.erase a).toList = l.toList.erase a /-! ### shrink -/ @@ -196,12 +169,10 @@ theorem size_set! (a : Array α) (i v) : (a.set! i v).size = a.size := by theorem mapM_empty [Monad m] (f : α → m β) : mapM f #[] = pure #[] := by rw [mapM, mapM.map]; rfl -@[simp] theorem map_empty (f : α → β) : map f #[] = #[] := mapM_empty .. +theorem map_empty (f : α → β) : map f #[] = #[] := mapM_empty f /-! ### mem -/ -alias not_mem_empty := not_mem_nil - theorem mem_singleton : a ∈ #[b] ↔ a = b := by simp /-! ### append -/ @@ -230,7 +201,7 @@ private theorem get_insertAt_loop_lt (as : Array α) (i : Fin (as.size+1)) (j : split · have h1 : k ≠ j - 1 := by omega have h2 : k ≠ j := by omega - rw [get_insertAt_loop_lt, get_swap, if_neg h1, if_neg h2] + rw [get_insertAt_loop_lt, getElem_swap, if_neg h1, if_neg h2] exact h · rfl @@ -241,7 +212,7 @@ private theorem get_insertAt_loop_gt (as : Array α) (i : Fin (as.size+1)) (j : split · have h1 : k ≠ j - 1 := by omega have h2 : k ≠ j := by omega - rw [get_insertAt_loop_gt, get_swap, if_neg h1, if_neg h2] + rw [get_insertAt_loop_gt, getElem_swap, if_neg h1, if_neg h2] exact Nat.lt_of_le_of_lt (Nat.pred_le _) hgt · rfl @@ -251,8 +222,7 @@ private theorem get_insertAt_loop_eq (as : Array α) (i : Fin (as.size+1)) (j : unfold insertAt.loop split · next h => - rw [get_insertAt_loop_eq, Fin.getElem_fin, get_swap, if_pos rfl] - exact Nat.lt_of_le_of_lt (Nat.pred_le _) j.is_lt + rw [get_insertAt_loop_eq, Fin.getElem_fin, getElem_swap, if_pos rfl] exact heq exact Nat.le_pred_of_lt h · congr; omega @@ -266,18 +236,17 @@ private theorem get_insertAt_loop_gt_le (as : Array α) (i : Fin (as.size+1)) (j if h0 : k = j then cases h0 have h1 : j.val ≠ j - 1 := by omega - rw [get_insertAt_loop_gt, get_swap, if_neg h1, if_pos rfl]; rfl - · exact j.is_lt - · exact Nat.pred_lt_of_lt hgt + rw [get_insertAt_loop_gt, getElem_swap, if_neg h1, if_pos rfl]; rfl + exact Nat.pred_lt_of_lt hgt else have h1 : k - 1 ≠ j - 1 := by omega have h2 : k - 1 ≠ j := by omega - rw [get_insertAt_loop_gt_le, get_swap, if_neg h1, if_neg h2] - exact hgt + rw [get_insertAt_loop_gt_le, getElem_swap, if_neg h1, if_neg h2] apply Nat.le_of_lt_add_one rw [Nat.sub_one_add_one] exact Nat.lt_of_le_of_ne hle h0 exact Nat.not_eq_zero_of_lt h + exact hgt · next h => absurd h exact Nat.lt_of_lt_of_le hgt hle diff --git a/Batteries/Data/Array/Merge.lean b/Batteries/Data/Array/Merge.lean index aa83e0200b..4b806a35d6 100644 --- a/Batteries/Data/Array/Merge.lean +++ b/Batteries/Data/Array/Merge.lean @@ -31,6 +31,8 @@ set_option linter.unusedVariables false in def mergeSortedPreservingDuplicates [ord : Ord α] (xs ys : Array α) : Array α := merge (compare · · |>.isLT) xs ys +-- We name `ord` so it can be provided as a named argument. +set_option linter.unusedVariables.funArgs false in /-- `O(|xs| + |ys|)`. Merge arrays `xs` and `ys`, which must be sorted according to `compare` and must not contain duplicates. Equal elements are merged using `merge`. If `merge` respects the order @@ -85,6 +87,8 @@ where @[deprecated (since := "2024-04-24")] alias mergeUnsortedDeduplicating := mergeUnsortedDedup +-- We name `eq` so it can be provided as a named argument. +set_option linter.unusedVariables.funArgs false in /-- `O(|xs|)`. Replace each run `[x₁, ⋯, xₙ]` of equal elements in `xs` with `f ⋯ (f (f x₁ x₂) x₃) ⋯ xₙ`. diff --git a/Batteries/Data/Array/OfFn.lean b/Batteries/Data/Array/OfFn.lean index a5d3bc78bc..e02be7cba1 100644 --- a/Batteries/Data/Array/OfFn.lean +++ b/Batteries/Data/Array/OfFn.lean @@ -12,12 +12,7 @@ namespace Array /-! ### ofFn -/ @[simp] -theorem data_ofFn (f : Fin n → α) : (ofFn f).data = List.ofFn f := by - ext1 - simp only [getElem?_eq, data_length, size_ofFn, length_ofFn, getElem_ofFn] - split - · rw [← getElem_eq_data_getElem] - simp - · rfl +theorem toList_ofFn (f : Fin n → α) : (ofFn f).toList = List.ofFn f := by + apply ext_getElem <;> simp end Array diff --git a/Batteries/Data/Array/Pairwise.lean b/Batteries/Data/Array/Pairwise.lean index 84dff1f680..9e61e642ed 100644 --- a/Batteries/Data/Array/Pairwise.lean +++ b/Batteries/Data/Array/Pairwise.lean @@ -17,15 +17,15 @@ larger indices. For example `as.Pairwise (· ≠ ·)` asserts that `as` has no duplicates, `as.Pairwise (· < ·)` asserts that `as` is strictly sorted and `as.Pairwise (· ≤ ·)` asserts that `as` is weakly sorted. -/ -def Pairwise (R : α → α → Prop) (as : Array α) : Prop := as.data.Pairwise R +def Pairwise (R : α → α → Prop) (as : Array α) : Prop := as.toList.Pairwise R theorem pairwise_iff_get {as : Array α} : as.Pairwise R ↔ ∀ (i j : Fin as.size), i < j → R (as.get i) (as.get j) := by - unfold Pairwise; simp [List.pairwise_iff_get, getElem_fin_eq_data_get]; rfl + unfold Pairwise; simp [List.pairwise_iff_get, getElem_fin_eq_toList_get] theorem pairwise_iff_getElem {as : Array α} : as.Pairwise R ↔ ∀ (i j : Nat) (_ : i < as.size) (_ : j < as.size), i < j → R as[i] as[j] := by - unfold Pairwise; simp [List.pairwise_iff_getElem, data_length]; rfl + unfold Pairwise; simp [List.pairwise_iff_getElem, length_toList] instance (R : α → α → Prop) [DecidableRel R] (as) : Decidable (Pairwise R as) := have : (∀ (j : Fin as.size) (i : Fin j.val), R as[i.val] (as[j.val])) ↔ Pairwise R as := by @@ -46,16 +46,17 @@ theorem pairwise_pair : #[a, b].Pairwise R ↔ R a b := by theorem pairwise_append {as bs : Array α} : (as ++ bs).Pairwise R ↔ as.Pairwise R ∧ bs.Pairwise R ∧ (∀ x ∈ as, ∀ y ∈ bs, R x y) := by - unfold Pairwise; simp [← mem_data, append_data, ← List.pairwise_append] + unfold Pairwise; simp [← mem_toList, toList_append, ← List.pairwise_append] theorem pairwise_push {as : Array α} : (as.push a).Pairwise R ↔ as.Pairwise R ∧ (∀ x ∈ as, R x a) := by unfold Pairwise - simp [← mem_data, push_data, List.pairwise_append, List.pairwise_singleton, List.mem_singleton] + simp [← mem_toList, push_toList, List.pairwise_append, List.pairwise_singleton, + List.mem_singleton] theorem pairwise_extract {as : Array α} (h : as.Pairwise R) (start stop) : (as.extract start stop).Pairwise R := by - simp only [pairwise_iff_getElem, get_extract, size_extract] at h ⊢ + simp only [pairwise_iff_getElem, getElem_extract, size_extract] at h ⊢ intro _ _ _ _ hlt apply h exact Nat.add_lt_add_left hlt start diff --git a/Batteries/Data/AssocList.lean b/Batteries/Data/AssocList.lean index 1492304efa..1e94b1d2f6 100644 --- a/Batteries/Data/AssocList.lean +++ b/Batteries/Data/AssocList.lean @@ -78,7 +78,7 @@ def toListTR (as : AssocList α β) : List (α × β) := @[csimp] theorem toList_eq_toListTR : @toList = @toListTR := by funext α β as; simp [toListTR] - exact .symm <| (Array.foldl_data_eq_map (toList as) _ id).trans (List.map_id _) + exact .symm <| (Array.foldl_toList_eq_map (toList as) _ id).trans (List.map_id _) /-- `O(n)`. Run monadic function `f` on all elements in the list, from head to tail. -/ @[specialize] def forM [Monad m] (f : α → β → m PUnit) : AssocList α β → m PUnit diff --git a/Batteries/Data/ByteArray.lean b/Batteries/Data/ByteArray.lean index ae4544fd21..c10cf6ba65 100644 --- a/Batteries/Data/ByteArray.lean +++ b/Batteries/Data/ByteArray.lean @@ -85,12 +85,12 @@ theorem size_append (a b : ByteArray) : (a ++ b).size = a.size + b.size := by theorem get_append_left {a b : ByteArray} (hlt : i < a.size) (h : i < (a ++ b).size := size_append .. ▸ Nat.lt_of_lt_of_le hlt (Nat.le_add_right ..)) : (a ++ b)[i] = a[i] := by - simp [getElem_eq_data_getElem]; exact Array.get_append_left hlt + simp [getElem_eq_data_getElem]; exact Array.getElem_append_left hlt theorem get_append_right {a b : ByteArray} (hle : a.size ≤ i) (h : i < (a ++ b).size) (h' : i - a.size < b.size := Nat.sub_lt_left_of_lt_add hle (size_append .. ▸ h)) : (a ++ b)[i] = b[i - a.size] := by - simp [getElem_eq_data_getElem]; exact Array.get_append_right hle + simp [getElem_eq_data_getElem]; exact Array.getElem_append_right hle /-! ### extract -/ diff --git a/Batteries/Data/Fin/Basic.lean b/Batteries/Data/Fin/Basic.lean index 3632d4f254..346f3006e5 100644 --- a/Batteries/Data/Fin/Basic.lean +++ b/Batteries/Data/Fin/Basic.lean @@ -13,4 +13,4 @@ def clamp (n m : Nat) : Fin (m + 1) := ⟨min n m, Nat.lt_succ_of_le (Nat.min_le def enum (n) : Array (Fin n) := Array.ofFn id /-- `list n` is the list of all elements of `Fin n` in order -/ -def list (n) : List (Fin n) := (enum n).data +def list (n) : List (Fin n) := (enum n).toList diff --git a/Batteries/Data/Fin/Lemmas.lean b/Batteries/Data/Fin/Lemmas.lean index 5ddde152de..33534b4549 100644 --- a/Batteries/Data/Fin/Lemmas.lean +++ b/Batteries/Data/Fin/Lemmas.lean @@ -26,7 +26,7 @@ attribute [norm_cast] val_last @[simp] theorem getElem_list (i : Nat) (h : i < (list n).length) : (list n)[i] = cast (length_list n) ⟨i, h⟩ := by - simp only [list]; rw [← Array.getElem_eq_data_getElem, getElem_enum, cast_mk] + simp only [list]; rw [← Array.getElem_eq_getElem_toList, getElem_enum, cast_mk] @[deprecated getElem_list (since := "2024-06-12")] theorem get_list (i : Fin (list n).length) : (list n).get i = i.cast (length_list n) := by @@ -40,14 +40,14 @@ theorem list_succ (n) : list (n+1) = 0 :: (list n).map Fin.succ := by theorem list_succ_last (n) : list (n+1) = (list n).map castSucc ++ [last n] := by rw [list_succ] induction n with - | zero => simp [last] + | zero => simp | succ n ih => rw [list_succ, List.map_cons castSucc, ih] simp [Function.comp_def, succ_castSucc] theorem list_reverse (n) : (list n).reverse = (list n).map rev := by induction n with - | zero => simp [last] + | zero => simp | succ n ih => conv => lhs; rw [list_succ_last] conv => rhs; rw [list_succ] diff --git a/Batteries/Data/HashMap/Basic.lean b/Batteries/Data/HashMap/Basic.lean index 3419ba35ab..98905c3343 100644 --- a/Batteries/Data/HashMap/Basic.lean +++ b/Batteries/Data/HashMap/Basic.lean @@ -37,7 +37,7 @@ def update (data : Buckets α β) (i : USize) The number of elements in the bucket array. Note: this is marked `noncomputable` because it is only intended for specification. -/ -noncomputable def size (data : Buckets α β) : Nat := .sum (data.1.data.map (·.toList.length)) +noncomputable def size (data : Buckets α β) : Nat := .sum (data.1.toList.map (·.toList.length)) @[simp] theorem update_size (self : Buckets α β) (i d h) : (self.update i d h).1.size = self.1.size := Array.size_uset .. @@ -52,7 +52,7 @@ The well-formedness invariant for the bucket array says that every element hashe -/ structure WF [BEq α] [Hashable α] (buckets : Buckets α β) : Prop where /-- The elements of a bucket are all distinct according to the `BEq` relation. -/ - distinct [LawfulHashable α] [PartialEquivBEq α] : ∀ bucket ∈ buckets.1.data, + distinct [LawfulHashable α] [PartialEquivBEq α] : ∀ bucket ∈ buckets.1.toList, bucket.toList.Pairwise fun a b => ¬(a.1 == b.1) /-- Every element in a bucket should hash to its location. -/ hash_self (i : Nat) (h : i < buckets.1.size) : @@ -237,7 +237,7 @@ inductive WF [BEq α] [Hashable α] : Imp α β → Prop where /-- Replacing an element in a well formed hash map yields a well formed hash map. -/ | modify : WF m → WF (modify m a f) -theorem WF.empty [BEq α] [Hashable α] : WF (empty n : Imp α β) := by unfold empty; apply empty' +theorem WF.empty [BEq α] [Hashable α] : WF (empty n : Imp α β) := empty' end Imp diff --git a/Batteries/Data/HashMap/WF.lean b/Batteries/Data/HashMap/WF.lean index 258723a89b..11b68de28d 100644 --- a/Batteries/Data/HashMap/WF.lean +++ b/Batteries/Data/HashMap/WF.lean @@ -15,26 +15,28 @@ attribute [-simp] Bool.not_eq_true namespace Buckets -@[ext] protected theorem ext : ∀ {b₁ b₂ : Buckets α β}, b₁.1.data = b₂.1.data → b₁ = b₂ +@[ext] protected theorem ext : ∀ {b₁ b₂ : Buckets α β}, b₁.1.toList = b₂.1.toList → b₁ = b₂ | ⟨⟨_⟩, _⟩, ⟨⟨_⟩, _⟩, rfl => rfl -theorem update_data (self : Buckets α β) (i d h) : - (self.update i d h).1.data = self.1.data.set i.toNat d := rfl +theorem toList_update (self : Buckets α β) (i d h) : + (self.update i d h).1.toList = self.1.toList.set i.toNat d := rfl + +@[deprecated (since := "2024-09-09")] alias update_data := toList_update theorem exists_of_update (self : Buckets α β) (i d h) : - ∃ l₁ l₂, self.1.data = l₁ ++ self.1[i] :: l₂ ∧ List.length l₁ = i.toNat ∧ - (self.update i d h).1.data = l₁ ++ d :: l₂ := by - simp only [Array.data_length, Array.ugetElem_eq_getElem, Array.getElem_eq_data_getElem] + ∃ l₁ l₂, self.1.toList = l₁ ++ self.1[i] :: l₂ ∧ List.length l₁ = i.toNat ∧ + (self.update i d h).1.toList = l₁ ++ d :: l₂ := by + simp only [Array.length_toList, Array.ugetElem_eq_getElem, Array.getElem_eq_getElem_toList] exact List.exists_of_set h theorem update_update (self : Buckets α β) (i d d' h h') : (self.update i d h).update i d' h' = self.update i d' h := by - simp only [update, Array.uset, Array.data_length] + simp only [update, Array.uset, Array.length_toList] congr 1 rw [Array.set_set] theorem size_eq (data : Buckets α β) : - size data = .sum (data.1.data.map (·.toList.length)) := rfl + size data = .sum (data.1.toList.map (·.toList.length)) := rfl theorem mk_size (h) : (mk n h : Buckets α β).size = 0 := by simp only [mk, mkArray, size_eq]; clear h @@ -44,7 +46,7 @@ theorem WF.mk' [BEq α] [Hashable α] (h) : (Buckets.mk n h : Buckets α β).WF refine ⟨fun _ h => ?_, fun i h => ?_⟩ · simp only [Buckets.mk, mkArray, List.mem_replicate, ne_eq] at h simp [h, List.Pairwise.nil] - · simp [Buckets.mk, empty', mkArray, Array.getElem_eq_data_getElem, AssocList.All] + · simp [Buckets.mk, empty', mkArray, Array.getElem_eq_getElem_toList, AssocList.All] theorem WF.update [BEq α] [Hashable α] {buckets : Buckets α β} {i d h} (H : buckets.WF) (h₁ : ∀ [PartialEquivBEq α] [LawfulHashable α], @@ -56,20 +58,20 @@ theorem WF.update [BEq α] [Hashable α] {buckets : Buckets α β} {i d h} (H : refine ⟨fun l hl => ?_, fun i hi p hp => ?_⟩ · exact match List.mem_or_eq_of_mem_set hl with | .inl hl => H.1 _ hl - | .inr rfl => h₁ (H.1 _ (Array.getElem_mem_data ..)) + | .inr rfl => h₁ (H.1 _ (Array.getElem_mem_toList ..)) · revert hp - simp only [Array.getElem_eq_data_getElem, update_data, List.getElem_set, Array.data_length, - update_size] + simp only [Array.getElem_eq_getElem_toList, toList_update, List.getElem_set, + Array.length_toList, update_size] split <;> intro hp · next eq => exact eq ▸ h₂ (H.2 _ _) _ hp - · simp only [update_size, Array.data_length] at hi + · simp only [update_size, Array.length_toList] at hi exact H.2 i hi _ hp end Buckets theorem reinsertAux_size [Hashable α] (data : Buckets α β) (a : α) (b : β) : (reinsertAux data a b).size = data.size.succ := by - simp only [reinsertAux, Array.data_length, Array.ugetElem_eq_getElem, Buckets.size_eq, + simp only [reinsertAux, Array.length_toList, Array.ugetElem_eq_getElem, Buckets.size_eq, Nat.succ_eq_add_one] refine have ⟨l₁, l₂, h₁, _, eq⟩ := Buckets.exists_of_update ..; eq ▸ ?_ simp [h₁, Nat.succ_add]; rfl @@ -89,35 +91,35 @@ theorem expand_size [Hashable α] {buckets : Buckets α β} : · rw [Buckets.mk_size]; simp [Buckets.size] · nofun where - go (i source) (target : Buckets α β) (hs : ∀ j < i, source.data[j]?.getD .nil = .nil) : + go (i source) (target : Buckets α β) (hs : ∀ j < i, source.toList[j]?.getD .nil = .nil) : (expand.go i source target).size = - .sum (source.data.map (·.toList.length)) + target.size := by + .sum (source.toList.map (·.toList.length)) + target.size := by unfold expand.go; split · next H => refine (go (i+1) _ _ fun j hj => ?a).trans ?b · case a => - simp only [Array.data_length, Array.data_set] + simp only [Array.length_toList, Array.toList_set] simp [List.getD_eq_getElem?_getD, List.getElem?_set, Option.map_eq_map]; split - · cases source.data[j]? <;> rfl + · cases source.toList[j]? <;> rfl · next H => exact hs _ (Nat.lt_of_le_of_ne (Nat.le_of_lt_succ hj) (Ne.symm H)) · case b => - simp only [Array.data_length, Array.data_set, Array.get_eq_getElem, AssocList.foldl_eq] + simp only [Array.length_toList, Array.toList_set, Array.get_eq_getElem, AssocList.foldl_eq] refine have ⟨l₁, l₂, h₁, _, eq⟩ := List.exists_of_set H; eq ▸ ?_ rw [h₁] simp only [Buckets.size_eq, List.map_append, List.map_cons, AssocList.toList, - List.length_nil, Nat.sum_append, Nat.sum_cons, Nat.zero_add, Array.data_length] + List.length_nil, Nat.sum_append, Nat.sum_cons, Nat.zero_add, Array.length_toList] rw [Nat.add_assoc, Nat.add_assoc, Nat.add_assoc]; congr 1 (conv => rhs; rw [Nat.add_left_comm]); congr 1 - rw [← Array.getElem_eq_data_getElem] + rw [← Array.getElem_eq_getElem_toList] have := @reinsertAux_size α β _; simp [Buckets.size] at this induction source[i].toList generalizing target <;> simp [*, Nat.succ_add]; rfl · next H => rw [(_ : Nat.sum _ = 0), Nat.zero_add] - rw [← (_ : source.data.map (fun _ => .nil) = source.data)] + rw [← (_ : source.toList.map (fun _ => .nil) = source.toList)] · simp only [List.map_map] - induction source.data <;> simp [*] + induction source.toList <;> simp [*] refine List.ext_getElem (by simp) fun j h₁ h₂ => ?_ - simp only [List.getElem_map, Array.data_length] + simp only [List.getElem_map, Array.length_toList] have := (hs j (Nat.lt_of_lt_of_le h₂ (Nat.not_lt.1 H))).symm rwa [List.getElem?_eq_getElem] at this termination_by source.size - i @@ -126,21 +128,21 @@ theorem expand_WF.foldl [BEq α] [Hashable α] (rank : α → Nat) {l : List (α (hl₁ : ∀ [PartialEquivBEq α] [LawfulHashable α], l.Pairwise fun a b => ¬(a.1 == b.1)) (hl₂ : ∀ x ∈ l, rank x.1 = i) {target : Buckets α β} (ht₁ : target.WF) - (ht₂ : ∀ bucket ∈ target.1.data, + (ht₂ : ∀ bucket ∈ target.1.toList, bucket.All fun k _ => rank k ≤ i ∧ ∀ [PartialEquivBEq α] [LawfulHashable α], ∀ x ∈ l, ¬(x.1 == k)) : (l.foldl (fun d x => reinsertAux d x.1 x.2) target).WF ∧ - ∀ bucket ∈ (l.foldl (fun d x => reinsertAux d x.1 x.2) target).1.data, + ∀ bucket ∈ (l.foldl (fun d x => reinsertAux d x.1 x.2) target).1.toList, bucket.All fun k _ => rank k ≤ i := by induction l generalizing target with | nil => exact ⟨ht₁, fun _ h₁ _ h₂ => (ht₂ _ h₁ _ h₂).1⟩ | cons _ _ ih => simp only [List.pairwise_cons, List.mem_cons, forall_eq_or_imp] at hl₁ hl₂ ht₂ refine ih hl₁.2 hl₂.2 - (reinsertAux_WF ht₁ fun _ h => (ht₂ _ (Array.getElem_mem_data ..) _ h).2.1) + (reinsertAux_WF ht₁ fun _ h => (ht₂ _ (Array.getElem_mem_toList ..) _ h).2.1) (fun _ h => ?_) - simp only [reinsertAux, Buckets.update, Array.uset, Array.data_length, - Array.ugetElem_eq_getElem, Array.data_set] at h + simp only [reinsertAux, Buckets.update, Array.uset, Array.length_toList, + Array.ugetElem_eq_getElem, Array.toList_set] at h match List.mem_or_eq_of_mem_set h with | .inl h => intro _ hf @@ -150,7 +152,7 @@ theorem expand_WF.foldl [BEq α] [Hashable α] (rank : α → Nat) {l : List (α | _, .head .. => exact ⟨hl₂.1 ▸ Nat.le_refl _, fun _ h h' => hl₁.1 _ h (PartialEquivBEq.symm h')⟩ | _, .tail _ h => - have ⟨h₁, h₂⟩ := ht₂ _ (Array.getElem_mem_data ..) _ h + have ⟨h₁, h₂⟩ := ht₂ _ (Array.getElem_mem_toList ..) _ h exact ⟨h₁, h₂.2⟩ theorem expand_WF [BEq α] [Hashable α] {buckets : Buckets α β} (H : buckets.WF) : @@ -158,11 +160,11 @@ theorem expand_WF [BEq α] [Hashable α] {buckets : Buckets α β} (H : buckets. go _ H.1 H.2 ⟨.mk' _, fun _ _ _ _ => by simp_all [Buckets.mk, List.mem_replicate]⟩ where go (i) {source : Array (AssocList α β)} - (hs₁ : ∀ [LawfulHashable α] [PartialEquivBEq α], ∀ bucket ∈ source.data, + (hs₁ : ∀ [LawfulHashable α] [PartialEquivBEq α], ∀ bucket ∈ source.toList, bucket.toList.Pairwise fun a b => ¬(a.1 == b.1)) (hs₂ : ∀ (j : Nat) (h : j < source.size), source[j].All fun k _ => ((hash k).toUSize % source.size).toNat = j) - {target : Buckets α β} (ht : target.WF ∧ ∀ bucket ∈ target.1.data, + {target : Buckets α β} (ht : target.WF ∧ ∀ bucket ∈ target.1.toList, bucket.All fun k _ => ((hash k).toUSize % source.size).toNat < i) : (expand.go i source target).WF := by unfold expand.go; split @@ -171,8 +173,8 @@ where · match List.mem_or_eq_of_mem_set hl with | .inl hl => exact hs₁ _ hl | .inr e => exact e ▸ .nil - · simp only [Array.data_length, Array.size_set, Array.getElem_eq_data_getElem, Array.data_set, - List.getElem_set] + · simp only [Array.length_toList, Array.size_set, Array.getElem_eq_getElem_toList, + Array.toList_set, List.getElem_set] split · nofun · exact hs₂ _ (by simp_all) @@ -180,7 +182,7 @@ where have := expand_WF.foldl rank ?_ (hs₂ _ H) ht.1 (fun _ h₁ _ h₂ => ?_) · simp only [Array.get_eq_getElem, AssocList.foldl_eq, Array.size_set] exact ⟨this.1, fun _ h₁ _ h₂ => Nat.lt_succ_of_le (this.2 _ h₁ _ h₂)⟩ - · exact hs₁ _ (Array.getElem_mem_data ..) + · exact hs₁ _ (Array.getElem_mem_toList ..) · have := ht.2 _ h₁ _ h₂ refine ⟨Nat.le_of_lt this, fun _ h h' => Nat.ne_of_lt this ?_⟩ exact LawfulHashable.hash_eq h' ▸ hs₂ _ H _ h @@ -198,7 +200,7 @@ theorem insert_size [BEq α] [Hashable α] {m : Imp α β} {k v} · unfold Buckets.size refine have ⟨_, _, h₁, _, eq⟩ := Buckets.exists_of_update ..; eq ▸ ?_ simp [h, h₁, Buckets.size_eq, Nat.succ_add]; rfl - · rw [expand_size]; simp only [expand, h, Buckets.size, Array.data_length, Buckets.update_size] + · rw [expand_size]; simp only [expand, h, Buckets.size, Array.length_toList, Buckets.update_size] refine have ⟨_, _, h₁, _, eq⟩ := Buckets.exists_of_update ..; eq ▸ ?_ simp [h₁, Buckets.size_eq, Nat.succ_add]; rfl @@ -264,7 +266,7 @@ theorem erase_size [BEq α] [Hashable α] {m : Imp α β} {k} · next H => simp only [h, Buckets.size] refine have ⟨_, _, h₁, _, eq⟩ := Buckets.exists_of_update ..; eq ▸ ?_ - simp only [h₁, Array.data_length, Array.ugetElem_eq_getElem, List.map_append, List.map_cons, + simp only [h₁, Array.length_toList, Array.ugetElem_eq_getElem, List.map_append, List.map_cons, Nat.sum_append, Nat.sum_cons, AssocList.toList_erase] rw [(_ : List.length _ = _ + 1), Nat.add_right_comm]; {rfl} clear h₁ eq @@ -317,8 +319,8 @@ theorem WF.mapVal {α β γ} {f : α → β → γ} [BEq α] [Hashable α] {m : Imp α β} (H : WF m) : WF (mapVal f m) := by have ⟨h₁, h₂⟩ := H.out simp only [Imp.mapVal, h₁, Buckets.mapVal, WF_iff]; refine ⟨?_, ?_, fun i h => ?_⟩ - · simp only [Buckets.size, Array.map_data, List.map_map]; congr; funext l; simp - · simp only [Array.map_data, List.forall_mem_map] + · simp only [Buckets.size, Array.toList_map, List.map_map]; congr; funext l; simp + · simp only [Array.toList_map, List.forall_mem_map] simp only [AssocList.toList_mapVal, List.pairwise_map] exact fun _ => h₂.1 _ · simp only [Array.size_map, AssocList.All, Array.getElem_map, AssocList.toList_mapVal, @@ -361,7 +363,7 @@ theorem WF.filterMap {α β γ} {f : α → β → Option γ} [BEq α] [Hashable suffices ∀ bk sz (h : 0 < bk.length), m.buckets.val.mapM (m := M) (filterMap.go f .nil) ⟨0⟩ = (⟨bk⟩, ⟨sz⟩) → WF ⟨sz, ⟨bk⟩, h⟩ from this _ _ _ rfl - simp only [Array.mapM_eq_mapM_data, bind, StateT.bind, H2, List.map_map, Nat.zero_add, g] + simp only [Array.mapM_eq_mapM_toList, Functor.map, StateT.map, H2, List.map_map, Nat.zero_add, g] intro bk sz h e'; cases e' refine .mk (by simp [Buckets.size]) ⟨?_, fun i h => ?_⟩ · simp only [List.forall_mem_map, List.toList_toAssocList] @@ -369,7 +371,7 @@ theorem WF.filterMap {α β γ} {f : α → β → Option γ} [BEq α] [Hashable have := H.out.2.1 _ h rw [← List.pairwise_map (R := (¬ · == ·))] at this ⊢ exact this.sublist (H3 l.toList) - · simp only [Array.size_mk, List.length_map, Array.data_length, Array.getElem_eq_data_getElem, + · simp only [Array.size_mk, List.length_map, Array.length_toList, Array.getElem_eq_getElem_toList, List.getElem_map] at h ⊢ have := H.out.2.2 _ h simp only [AssocList.All, List.toList_toAssocList, List.mem_reverse, List.mem_filterMap, @@ -385,13 +387,15 @@ variable {_ : BEq α} {_ : Hashable α} @[inline] def mapVal (f : α → β → γ) (self : HashMap α β) : HashMap α γ := ⟨self.1.mapVal f, self.2.mapVal⟩ -/-- -Applies `f` to each key-value pair `a, b` in the map. If it returns `some c` then -`a, c` is pushed into the new map; else the key is removed from the map. --/ -@[inline] def filterMap (f : α → β → Option γ) (self : HashMap α β) : HashMap α γ := - ⟨self.1.filterMap f, self.2.filterMap⟩ +-- Temporarily removed on lean-pr-testing-5403. + +-- /-- +-- Applies `f` to each key-value pair `a, b` in the map. If it returns `some c` then +-- `a, c` is pushed into the new map; else the key is removed from the map. +-- -/ +-- @[inline] def filterMap (f : α → β → Option γ) (self : HashMap α β) : HashMap α γ := +-- ⟨self.1.filterMap f, self.2.filterMap⟩ -/-- Constructs a map with the set of all pairs `a, b` such that `f` returns true. -/ -@[inline] def filter (f : α → β → Bool) (self : HashMap α β) : HashMap α β := - self.filterMap fun a b => bif f a b then some b else none +-- /-- Constructs a map with the set of all pairs `a, b` such that `f` returns true. -/ +-- @[inline] def filter (f : α → β → Bool) (self : HashMap α β) : HashMap α β := +-- self.filterMap fun a b => bif f a b then some b else none diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index 29061c409e..49126faee0 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -86,7 +86,7 @@ drop_while (· != 1) [0, 1, 2, 3] = [1, 2, 3] @[csimp] theorem replaceF_eq_replaceFTR : @replaceF = @replaceFTR := by funext α p l; simp [replaceFTR] - let rec go (acc) : ∀ xs, replaceFTR.go p xs acc = acc.data ++ xs.replaceF p + let rec go (acc) : ∀ xs, replaceFTR.go p xs acc = acc.toList ++ xs.replaceF p | [] => by simp [replaceFTR.go, replaceF] | x::xs => by simp [replaceFTR.go, replaceF]; cases p x <;> simp @@ -149,8 +149,9 @@ def splitOnP (P : α → Bool) (l : List α) : List (List α) := go l [] where @[csimp] theorem splitOnP_eq_splitOnPTR : @splitOnP = @splitOnPTR := by funext α P l; simp [splitOnPTR] - suffices ∀ xs acc r, splitOnPTR.go P xs acc r = r.data ++ splitOnP.go P xs acc.data.reverse from - (this l #[] #[]).symm + suffices ∀ xs acc r, + splitOnPTR.go P xs acc r = r.toList ++ splitOnP.go P xs acc.toList.reverse from + (this l #[] #[]).symm intro xs acc r; induction xs generalizing acc r with simp [splitOnP.go, splitOnPTR.go] | cons x xs IH => cases P x <;> simp [*] @@ -196,7 +197,7 @@ def modifyNthTR (f : α → α) (n : Nat) (l : List α) : List α := go l n #[] | a :: l, 0, acc => acc.toListAppend (f a :: l) | a :: l, n+1, acc => go l n (acc.push a) -theorem modifyNthTR_go_eq : ∀ l n, modifyNthTR.go f l n acc = acc.data ++ modifyNth f n l +theorem modifyNthTR_go_eq : ∀ l n, modifyNthTR.go f l n acc = acc.toList ++ modifyNth f n l | [], n => by cases n <;> simp [modifyNthTR.go, modifyNth] | a :: l, 0 => by simp [modifyNthTR.go, modifyNth] | a :: l, n+1 => by simp [modifyNthTR.go, modifyNth, modifyNthTR_go_eq l] @@ -229,7 +230,7 @@ def insertNth (n : Nat) (a : α) : List α → List α := | _, [], acc => acc.toList | n+1, a :: l, acc => go n l (acc.push a) -theorem insertNthTR_go_eq : ∀ n l, insertNthTR.go a n l acc = acc.data ++ insertNth n a l +theorem insertNthTR_go_eq : ∀ n l, insertNthTR.go a n l acc = acc.toList ++ insertNth n a l | 0, l | _+1, [] => by simp [insertNthTR.go, insertNth] | n+1, a :: l => by simp [insertNthTR.go, insertNth, insertNthTR_go_eq n l] @@ -261,7 +262,7 @@ def takeDTR (n : Nat) (l : List α) (dflt : α) : List α := go n l #[] where | 0, _, acc => acc.toList | n, [], acc => acc.toListAppend (replicate n dflt) -theorem takeDTR_go_eq : ∀ n l, takeDTR.go dflt n l acc = acc.data ++ takeD n l dflt +theorem takeDTR_go_eq : ∀ n l, takeDTR.go dflt n l acc = acc.toList ++ takeD n l dflt | 0, _ => by simp [takeDTR.go] | _+1, [] => by simp [takeDTR.go, replicate_succ] | _+1, _::l => by simp [takeDTR.go, takeDTR_go_eq _ l] @@ -286,7 +287,7 @@ scanl (+) 0 [1, 2, 3] = [0, 1, 3, 6] | [], a, acc => acc.toListAppend [a] | b :: l, a, acc => go l (f a b) (acc.push a) -theorem scanlTR_go_eq : ∀ l, scanlTR.go f l a acc = acc.data ++ scanl f a l +theorem scanlTR_go_eq : ∀ l, scanlTR.go f l a acc = acc.toList ++ scanl f a l | [] => by simp [scanlTR.go, scanl] | a :: l => by simp [scanlTR.go, scanl, scanlTR_go_eq l] @@ -538,8 +539,8 @@ theorem sections_eq_nil_of_isEmpty : ∀ {L}, L.any isEmpty → @sections α L = cases e : L.any isEmpty <;> simp [sections_eq_nil_of_isEmpty, *] clear e; induction L with | nil => rfl | cons l L IH => ?_ simp [IH, sectionsTR.go] - rw [Array.foldl_eq_foldl_data, Array.foldl_data_eq_bind]; rfl - intros; apply Array.foldl_data_eq_map + rw [Array.foldl_eq_foldl_toList, Array.foldl_toList_eq_bind]; rfl + intros; apply Array.foldl_toList_eq_map /-- `extractP p l` returns a pair of an element `a` of `l` satisfying the predicate @@ -576,8 +577,8 @@ def productTR (l₁ : List α) (l₂ : List β) : List (α × β) := @[csimp] theorem product_eq_productTR : @product = @productTR := by funext α β l₁ l₂; simp [product, productTR] - rw [Array.foldl_data_eq_bind]; rfl - intros; apply Array.foldl_data_eq_map + rw [Array.foldl_toList_eq_bind]; rfl + intros; apply Array.foldl_toList_eq_map /-- `sigma l₁ l₂` is the list of dependent pairs `(a, b)` where `a ∈ l₁` and `b ∈ l₂ a`. ``` @@ -592,8 +593,8 @@ def sigmaTR {σ : α → Type _} (l₁ : List α) (l₂ : ∀ a, List (σ a)) : @[csimp] theorem sigma_eq_sigmaTR : @List.sigma = @sigmaTR := by funext α β l₁ l₂; simp [List.sigma, sigmaTR] - rw [Array.foldl_data_eq_bind]; rfl - intros; apply Array.foldl_data_eq_map + rw [Array.foldl_toList_eq_bind]; rfl + intros; apply Array.foldl_toList_eq_map /-- `ofFn f` with `f : fin n → α` returns the list whose ith element is `f i` @@ -765,8 +766,8 @@ theorem dropSlice_zero₂ : ∀ n l, @dropSlice α n 0 l = l funext α n m l; simp [dropSliceTR] split; { rw [dropSlice_zero₂] } rename_i m - let rec go (acc) : ∀ xs n, l = acc.data ++ xs → - dropSliceTR.go l m xs n acc = acc.data ++ xs.dropSlice n (m+1) + let rec go (acc) : ∀ xs n, l = acc.toList ++ xs → + dropSliceTR.go l m xs n acc = acc.toList ++ xs.dropSlice n (m+1) | [], n | _::xs, 0 => fun h => by simp [dropSliceTR.go, dropSlice, h] | x::xs, n+1 => by simp [dropSliceTR.go, dropSlice]; intro h; rw [go _ xs]; {simp}; simp [h] @@ -801,7 +802,7 @@ zipWithLeft' prod.mk [1] ['a', 'b'] = ([(1, some 'a')], ['b']) let rec go (acc) : ∀ as bs, zipWithLeft'TR.go f as bs acc = let (l, r) := as.zipWithLeft' f bs; (acc.toList ++ l, r) | [], bs => by simp [zipWithLeft'TR.go] - | _::_, [] => by simp [zipWithLeft'TR.go, Array.foldl_data_eq_map] + | _::_, [] => by simp [zipWithLeft'TR.go, Array.foldl_toList_eq_map] | a::as, b::bs => by simp [zipWithLeft'TR.go, go _ as bs] simp [zipWithLeft'TR, go] @@ -870,7 +871,7 @@ zipWithLeft f as bs = (zipWithLeft' f as bs).fst funext α β γ f as bs; simp [zipWithLeftTR] let rec go (acc) : ∀ as bs, zipWithLeftTR.go f as bs acc = acc.toList ++ as.zipWithLeft f bs | [], bs => by simp [zipWithLeftTR.go] - | _::_, [] => by simp [zipWithLeftTR.go, Array.foldl_data_eq_map] + | _::_, [] => by simp [zipWithLeftTR.go, Array.foldl_toList_eq_map] | a::as, b::bs => by simp [zipWithLeftTR.go, go _ as bs] simp [zipWithLeftTR, go] @@ -946,7 +947,7 @@ fillNones [none, some 1, none, none] [2, 3] = [2, 1, 3] @[csimp] theorem fillNones_eq_fillNonesTR : @fillNones = @fillNonesTR := by funext α as as'; simp [fillNonesTR] - let rec go (acc) : ∀ as as', @fillNonesTR.go α as as' acc = acc.data ++ as.fillNones as' + let rec go (acc) : ∀ as as', @fillNonesTR.go α as as' acc = acc.toList ++ as.fillNones as' | [], _ => by simp [fillNonesTR.go] | some a :: as, as' => by simp [fillNonesTR.go, go _ as as'] | none :: as, [] => by simp [fillNonesTR.go, reduceOption, filterMap_eq_filterMapTR.go] diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index f03e693d16..1ac4ae3f57 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -10,31 +10,11 @@ import Batteries.Tactic.Alias namespace List -open Nat - -/-! ### mem -/ - -@[simp] theorem mem_toArray {a : α} {l : List α} : a ∈ l.toArray ↔ a ∈ l := by - simp [Array.mem_def] - /-! ### toArray-/ -@[simp] theorem size_toArrayAux (l : List α) (r : Array α) : - (l.toArrayAux r).size = r.size + l.length := by - induction l generalizing r with - | nil => simp [toArrayAux] - | cons a l ih => - simp [ih, List.toArrayAux] - omega - @[simp] theorem getElem_mk {xs : List α} {i : Nat} (h : i < xs.length) : (Array.mk xs)[i] = xs[i] := rfl -@[simp] theorem getElem_toArray (l : List α) (i : Nat) (h : i < l.toArray.size) : - l.toArray[i] = l[i]'(by simpa using h) := by - rw [Array.getElem_eq_data_getElem] - simp - /-! ### next? -/ @[simp] theorem next?_nil : @next? α [] = none := rfl @@ -196,18 +176,18 @@ theorem exists_of_set' {l : List α} (h : n < l.length) : ∃ l₁ a l₂, l = l₁ ++ a :: l₂ ∧ l₁.length = n ∧ l.set n a' = l₁ ++ a' :: l₂ := by rw [set_eq_modifyNth]; exact exists_of_modifyNth _ h -@[deprecated getElem?_set_eq' (since := "2024-06-12")] +@[deprecated getElem?_set_self' (since := "2024-06-12")] theorem get?_set_eq (a : α) (n) (l : List α) : (set l n a).get? n = (fun _ => a) <$> l.get? n := by - simp only [get?_eq_getElem?, getElem?_set_eq', Option.map_eq_map] + simp only [get?_eq_getElem?, getElem?_set_self', Option.map_eq_map] rfl theorem getElem?_set_eq_of_lt (a : α) {n} {l : List α} (h : n < length l) : - (set l n a)[n]? = some a := by rw [getElem?_set_eq', getElem?_eq_getElem h]; rfl + (set l n a)[n]? = some a := by rw [getElem?_set_self', getElem?_eq_getElem h]; rfl @[deprecated getElem?_set_eq_of_lt (since := "2024-06-12")] theorem get?_set_eq_of_lt (a : α) {n} {l : List α} (h : n < length l) : (set l n a).get? n = some a := by - rw [get?_eq_getElem?, getElem?_set_eq', getElem?_eq_getElem h]; rfl + rw [get?_eq_getElem?, getElem?_set_self', getElem?_eq_getElem h]; rfl @[deprecated getElem?_set_ne (since := "2024-06-12")] theorem get?_set_ne (a : α) {m n} (l : List α) (h : m ≠ n) : (set l m a).get? n = l.get? n := by @@ -216,7 +196,7 @@ theorem get?_set_ne (a : α) {m n} (l : List α) (h : m ≠ n) : (set l m a).get @[deprecated getElem?_set (since := "2024-06-12")] theorem get?_set (a : α) {m n} (l : List α) : (set l m a).get? n = if m = n then (fun _ => a) <$> l.get? n else l.get? n := by - simp [getElem?_set'] + simp [getElem?_set']; rfl theorem get?_set_of_lt (a : α) {m n} (l : List α) (h : n < length l) : (set l m a).get? n = if m = n then some a else l.get? n := by @@ -233,18 +213,12 @@ theorem get?_set_of_lt' (a : α) {m n} (l : List α) (h : m < length l) : theorem length_tail_add_one (l : List α) (h : 0 < length l) : (length (tail l)) + 1 = length l := by simp [Nat.sub_add_cancel h] -@[simp] theorem getElem?_tail (l : List α) : l.tail[n]? = l[n + 1]? := by cases l <;> simp - -@[simp] theorem getElem_tail (l : List α) (h : n < l.tail.length) : - l.tail[n] = l[n + 1]'(by simp at h; omega) := by - cases l; contradiction; simp - /-! ### eraseP -/ @[simp] theorem extractP_eq_find?_eraseP (l : List α) : extractP p l = (find? p l, eraseP p l) := by - let rec go (acc) : ∀ xs, l = acc.data ++ xs → - extractP.go p l xs acc = (xs.find? p, acc.data ++ xs.eraseP p) + let rec go (acc) : ∀ xs, l = acc.toList ++ xs → + extractP.go p l xs acc = (xs.find? p, acc.toList ++ xs.eraseP p) | [] => fun h => by simp [extractP.go, find?, eraseP, h] | x::xs => by simp [extractP.go, find?, eraseP]; cases p x <;> simp @@ -419,7 +393,7 @@ theorem pair_mem_product {xs : List α} {ys : List β} {x : α} {y : β} : theorem forIn_eq_bindList [Monad m] [LawfulMonad m] (f : α → β → m (ForInStep β)) (l : List α) (init : β) : forIn l init f = ForInStep.run <$> (ForInStep.yield init).bindList f l := by - induction l generalizing init <;> simp [*, map_eq_pure_bind] + induction l generalizing init <;> simp [*] congr; ext (b | b) <;> simp /-! ### diff -/ @@ -675,37 +649,6 @@ theorem insertP_loop (a : α) (l r : List α) : induction l with simp [insertP, insertP.loop, cond] | cons _ _ ih => split <;> simp [insertP_loop, ih] -/-! ### merge -/ - -theorem cons_merge_cons (s : α → α → Bool) (a b l r) : - merge s (a::l) (b::r) = if s a b then a :: merge s l (b::r) else b :: merge s (a::l) r := by - simp only [merge] - -@[simp] theorem cons_merge_cons_pos (s : α → α → Bool) (l r) (h : s a b) : - merge s (a::l) (b::r) = a :: merge s l (b::r) := by - rw [cons_merge_cons, if_pos h] - -@[simp] theorem cons_merge_cons_neg (s : α → α → Bool) (l r) (h : ¬ s a b) : - merge s (a::l) (b::r) = b :: merge s (a::l) r := by - rw [cons_merge_cons, if_neg h] - -@[simp] theorem length_merge (s : α → α → Bool) (l r) : - (merge s l r).length = l.length + r.length := by - match l, r with - | [], r => simp - | l, [] => simp - | a::l, b::r => - rw [cons_merge_cons] - split - · simp_arith [length_merge s l (b::r)] - · simp_arith [length_merge s (a::l) r] - -theorem mem_merge_left (s : α → α → Bool) (h : x ∈ l) : x ∈ merge s l r := - mem_merge.2 <| .inl h - -theorem mem_merge_right (s : α → α → Bool) (h : x ∈ r) : x ∈ merge s l r := - mem_merge.2 <| .inr h - /-! ### foldlM and foldrM -/ theorem foldlM_map [Monad m] (f : β₁ → β₂) (g : α → β₂ → m α) (l : List β₁) (init : α) : diff --git a/Batteries/Data/List/Perm.lean b/Batteries/Data/List/Perm.lean index 42c4bbc9c2..d924ebf4dd 100644 --- a/Batteries/Data/List/Perm.lean +++ b/Batteries/Data/List/Perm.lean @@ -235,8 +235,8 @@ theorem subperm_append_diff_self_of_count_le {l₁ l₂ : List α} | nil => simp | cons hd tl IH => have : hd ∈ l₂ := by - rw [← count_pos_iff_mem] - exact Nat.lt_of_lt_of_le (count_pos_iff_mem.mpr (.head _)) (h hd (.head _)) + rw [← count_pos_iff] + exact Nat.lt_of_lt_of_le (count_pos_iff.mpr (.head _)) (h hd (.head _)) have := perm_cons_erase this refine Perm.trans ?_ this.symm rw [cons_append, diff_cons, perm_cons] @@ -325,7 +325,7 @@ theorem perm_insertP (p : α → Bool) (a l) : insertP p a l ~ a :: l := by theorem Perm.insertP (p : α → Bool) (a) (h : l₁ ~ l₂) : insertP p a l₁ ~ insertP p a l₂ := Perm.trans (perm_insertP ..) <| Perm.trans (Perm.cons _ h) <| Perm.symm (perm_insertP ..) -theorem perm_merge (s : α → α → Bool) (l r) : merge s l r ~ l ++ r := by +theorem perm_merge (s : α → α → Bool) (l r) : merge l r s ~ l ++ r := by match l, r with | [], r => simp | l, [] => simp @@ -342,5 +342,5 @@ theorem perm_merge (s : α → α → Bool) (l r) : merge s l r ~ l ++ r := by exact Perm.rfl theorem Perm.merge (s₁ s₂ : α → α → Bool) (hl : l₁ ~ l₂) (hr : r₁ ~ r₂) : - merge s₁ l₁ r₁ ~ merge s₂ l₂ r₂ := + merge l₁ r₁ s₁ ~ merge l₂ r₂ s₂ := Perm.trans (perm_merge ..) <| Perm.trans (Perm.append hl hr) <| Perm.symm (perm_merge ..) diff --git a/Batteries/Data/RBMap/Alter.lean b/Batteries/Data/RBMap/Alter.lean index 3f6d25c2b3..fd1e8f205a 100644 --- a/Batteries/Data/RBMap/Alter.lean +++ b/Batteries/Data/RBMap/Alter.lean @@ -206,14 +206,20 @@ theorem Ordered.ins : ∀ {path : Path α} {t : RBNode α}, t.Ordered cmp → path.Ordered cmp → t.All (path.RootOrdered cmp) → (path.ins t).Ordered cmp | .root, _, ht, _, _ => Ordered.setBlack.2 ht | .left red parent x b, a, ha, ⟨hp, xb, xp, bp, hb⟩, H => by - unfold ins; have ⟨ax, ap⟩ := All_and.1 H; exact hp.ins ⟨ax, xb, ha, hb⟩ ⟨xp, ap, bp⟩ + unfold Path.ins + have ⟨ax, ap⟩ := All_and.1 H + exact hp.ins ⟨ax, xb, ha, hb⟩ ⟨xp, ap, bp⟩ | .right red a x parent, b, hb, ⟨hp, ax, xp, ap, ha⟩, H => by - unfold ins; have ⟨xb, bp⟩ := All_and.1 H; exact hp.ins ⟨ax, xb, ha, hb⟩ ⟨xp, ap, bp⟩ + unfold Path.ins + have ⟨xb, bp⟩ := All_and.1 H + exact hp.ins ⟨ax, xb, ha, hb⟩ ⟨xp, ap, bp⟩ | .left black parent x b, a, ha, ⟨hp, xb, xp, bp, hb⟩, H => by - unfold ins; have ⟨ax, ap⟩ := All_and.1 H + unfold Path.ins + have ⟨ax, ap⟩ := All_and.1 H exact hp.ins (ha.balance1 ax xb hb) (balance1_All.2 ⟨xp, ap, bp⟩) | .right black a x parent, b, hb, ⟨hp, ax, xp, ap, ha⟩, H => by - unfold ins; have ⟨xb, bp⟩ := All_and.1 H + unfold Path.ins + have ⟨xb, bp⟩ := All_and.1 H exact hp.ins (ha.balance2 ax xb hb) (balance2_All.2 ⟨xp, ap, bp⟩) theorem Ordered.insertNew {path : Path α} (hp : path.Ordered cmp) (vp : path.RootOrdered cmp v) : @@ -224,14 +230,20 @@ theorem Ordered.del : ∀ {path : Path α} {t : RBNode α} {c}, t.Ordered cmp → path.Ordered cmp → t.All (path.RootOrdered cmp) → (path.del t c).Ordered cmp | .root, _, _, ht, _, _ => Ordered.setBlack.2 ht | .left _ parent x b, a, red, ha, ⟨hp, xb, xp, bp, hb⟩, H => by - unfold del; have ⟨ax, ap⟩ := All_and.1 H; exact hp.del ⟨ax, xb, ha, hb⟩ ⟨xp, ap, bp⟩ + unfold Path.del + have ⟨ax, ap⟩ := All_and.1 H + exact hp.del ⟨ax, xb, ha, hb⟩ ⟨xp, ap, bp⟩ | .right _ a x parent, b, red, hb, ⟨hp, ax, xp, ap, ha⟩, H => by - unfold del; have ⟨xb, bp⟩ := All_and.1 H; exact hp.del ⟨ax, xb, ha, hb⟩ ⟨xp, ap, bp⟩ + unfold Path.del + have ⟨xb, bp⟩ := All_and.1 H + exact hp.del ⟨ax, xb, ha, hb⟩ ⟨xp, ap, bp⟩ | .left _ parent x b, a, black, ha, ⟨hp, xb, xp, bp, hb⟩, H => by - unfold del; have ⟨ax, ap⟩ := All_and.1 H + unfold Path.del + have ⟨ax, ap⟩ := All_and.1 H exact hp.del (ha.balLeft ax xb hb) (ap.balLeft xp bp) | .right _ a x parent, b, black, hb, ⟨hp, ax, xp, ap, ha⟩, H => by - unfold del; have ⟨xb, bp⟩ := All_and.1 H + unfold Path.del + have ⟨xb, bp⟩ := All_and.1 H exact hp.del (ha.balRight ax xb hb) (ap.balRight xp bp) end Path diff --git a/Batteries/Data/RBMap/WF.lean b/Batteries/Data/RBMap/WF.lean index 6fbd669e6e..220852c0ab 100644 --- a/Batteries/Data/RBMap/WF.lean +++ b/Batteries/Data/RBMap/WF.lean @@ -261,7 +261,8 @@ so this is only suitable for use on the root of the tree.) -/ theorem Balanced.insert {t : RBNode α} (h : t.Balanced c n) : ∃ c' n', (insert cmp t v).Balanced c' n' := by - unfold insert; match ins cmp v t, h.ins cmp v with + unfold RBNode.insert + match ins cmp v t, h.ins cmp v with | _, .balanced h => split <;> [exact ⟨_, h.setBlack⟩; exact ⟨_, _, h⟩] | _, .redred _ ha hb => have .node red .. := t; exact ⟨_, _, .black ha hb⟩ diff --git a/Batteries/Data/Rat/Basic.lean b/Batteries/Data/Rat/Basic.lean index f96651d34c..2032b23ecd 100644 --- a/Batteries/Data/Rat/Basic.lean +++ b/Batteries/Data/Rat/Basic.lean @@ -45,23 +45,23 @@ Auxiliary definition for `Rat.normalize`. Constructs `num / den` as a rational n dividing both `num` and `den` by `g` (which is the gcd of the two) if it is not 1. -/ @[inline] def Rat.maybeNormalize (num : Int) (den g : Nat) - (den_nz : den / g ≠ 0) (reduced : (num.div g).natAbs.Coprime (den / g)) : Rat := + (den_nz : den / g ≠ 0) (reduced : (num.tdiv g).natAbs.Coprime (den / g)) : Rat := if hg : g = 1 then { num, den den_nz := by simp [hg] at den_nz; exact den_nz reduced := by simp [hg, Int.natAbs_ofNat] at reduced; exact reduced } - else { num := num.div g, den := den / g, den_nz, reduced } + else { num := num.tdiv g, den := den / g, den_nz, reduced } theorem Rat.normalize.den_nz {num : Int} {den g : Nat} (den_nz : den ≠ 0) (e : g = num.natAbs.gcd den) : den / g ≠ 0 := e ▸ Nat.ne_of_gt (Nat.div_gcd_pos_of_pos_right _ (Nat.pos_of_ne_zero den_nz)) theorem Rat.normalize.reduced {num : Int} {den g : Nat} (den_nz : den ≠ 0) - (e : g = num.natAbs.gcd den) : (num.div g).natAbs.Coprime (den / g) := - have : Int.natAbs (num.div ↑g) = num.natAbs / g := by + (e : g = num.natAbs.gcd den) : (num.tdiv g).natAbs.Coprime (den / g) := + have : Int.natAbs (num.tdiv ↑g) = num.natAbs / g := by match num, num.eq_nat_or_neg with | _, ⟨_, .inl rfl⟩ => rfl - | _, ⟨_, .inr rfl⟩ => rw [Int.neg_div, Int.natAbs_neg, Int.natAbs_neg]; rfl + | _, ⟨_, .inr rfl⟩ => rw [Int.neg_tdiv, Int.natAbs_neg, Int.natAbs_neg]; rfl this ▸ e ▸ Nat.coprime_div_gcd_div_gcd (Nat.gcd_pos_of_pos_right _ (Nat.pos_of_ne_zero den_nz)) /-- @@ -141,12 +141,12 @@ want to unfold it. Use `Rat.mul_def` instead.) -/ @[irreducible] protected def mul (a b : Rat) : Rat := let g1 := Nat.gcd a.num.natAbs b.den let g2 := Nat.gcd b.num.natAbs a.den - { num := (a.num.div g1) * (b.num.div g2) + { num := (a.num.tdiv g1) * (b.num.tdiv g2) den := (a.den / g2) * (b.den / g1) den_nz := Nat.ne_of_gt <| Nat.mul_pos (Nat.div_gcd_pos_of_pos_right _ a.den_pos) (Nat.div_gcd_pos_of_pos_right _ b.den_pos) reduced := by - simp only [Int.natAbs_mul, Int.natAbs_div, Nat.coprime_mul_iff_left] + simp only [Int.natAbs_mul, Int.natAbs_tdiv, Nat.coprime_mul_iff_left] refine ⟨Nat.coprime_mul_iff_right.2 ⟨?_, ?_⟩, Nat.coprime_mul_iff_right.2 ⟨?_, ?_⟩⟩ · exact a.reduced.coprime_div_left (Nat.gcd_dvd_left ..) |>.coprime_div_right (Nat.gcd_dvd_right ..) diff --git a/Batteries/Data/Rat/Lemmas.lean b/Batteries/Data/Rat/Lemmas.lean index 77b911ec94..b3c2d49915 100644 --- a/Batteries/Data/Rat/Lemmas.lean +++ b/Batteries/Data/Rat/Lemmas.lean @@ -23,14 +23,14 @@ theorem ext : {p q : Rat} → p.num = q.num → p.den = q.den → p = q @[simp] theorem maybeNormalize_eq {num den g} (den_nz reduced) : maybeNormalize num den g den_nz reduced = - { num := num.div g, den := den / g, den_nz, reduced } := by + { num := num.tdiv g, den := den / g, den_nz, reduced } := by unfold maybeNormalize; split · subst g; simp · rfl theorem normalize.reduced' {num : Int} {den g : Nat} (den_nz : den ≠ 0) (e : g = num.natAbs.gcd den) : (num / g).natAbs.Coprime (den / g) := by - rw [← Int.div_eq_ediv_of_dvd (e ▸ Int.ofNat_dvd_left.2 (Nat.gcd_dvd_left ..))] + rw [← Int.tdiv_eq_ediv_of_dvd (e ▸ Int.ofNat_dvd_left.2 (Nat.gcd_dvd_left ..))] exact normalize.reduced den_nz e theorem normalize_eq {num den} (den_nz) : normalize num den den_nz = @@ -39,10 +39,10 @@ theorem normalize_eq {num den} (den_nz) : normalize num den den_nz = den_nz := normalize.den_nz den_nz rfl reduced := normalize.reduced' den_nz rfl } := by simp only [normalize, maybeNormalize_eq, - Int.div_eq_ediv_of_dvd (Int.ofNat_dvd_left.2 (Nat.gcd_dvd_left ..))] + Int.tdiv_eq_ediv_of_dvd (Int.ofNat_dvd_left.2 (Nat.gcd_dvd_left ..))] @[simp] theorem normalize_zero (nz) : normalize 0 d nz = 0 := by - simp [normalize, Int.zero_div, Int.natAbs_zero, Nat.div_self (Nat.pos_of_ne_zero nz)]; rfl + simp [normalize, Int.zero_tdiv, Int.natAbs_zero, Nat.div_self (Nat.pos_of_ne_zero nz)]; rfl theorem mk_eq_normalize (num den nz c) : ⟨num, den, nz, c⟩ = normalize num den nz := by simp [normalize_eq, c.gcd_eq_one] @@ -76,7 +76,7 @@ theorem normalize_eq_iff (z₁ : d₁ ≠ 0) (z₂ : d₂ ≠ 0) : theorem maybeNormalize_eq_normalize {num : Int} {den g : Nat} (den_nz reduced) (hn : ↑g ∣ num) (hd : g ∣ den) : maybeNormalize num den g den_nz reduced = normalize num den (mt (by simp [·]) den_nz) := by - simp only [maybeNormalize_eq, mk_eq_normalize, Int.div_eq_ediv_of_dvd hn] + simp only [maybeNormalize_eq, mk_eq_normalize, Int.tdiv_eq_ediv_of_dvd hn] have : g ≠ 0 := mt (by simp [·]) den_nz rw [← normalize_mul_right _ this, Int.ediv_mul_cancel hn] congr 1; exact Nat.div_mul_cancel hd @@ -267,9 +267,9 @@ theorem mul_def (a b : Rat) : have H1 : a.num.natAbs.gcd b.den ≠ 0 := Nat.gcd_ne_zero_right b.den_nz have H2 : b.num.natAbs.gcd a.den ≠ 0 := Nat.gcd_ne_zero_right a.den_nz rw [mk_eq_normalize, ← normalize_mul_right _ (Nat.mul_ne_zero H1 H2)]; congr 1 - · rw [Int.ofNat_mul, ← Int.mul_assoc, Int.mul_right_comm (Int.div ..), - Int.div_mul_cancel (Int.ofNat_dvd_left.2 <| Nat.gcd_dvd_left ..), Int.mul_assoc, - Int.div_mul_cancel (Int.ofNat_dvd_left.2 <| Nat.gcd_dvd_left ..)] + · rw [Int.ofNat_mul, ← Int.mul_assoc, Int.mul_right_comm (Int.tdiv ..), + Int.tdiv_mul_cancel (Int.ofNat_dvd_left.2 <| Nat.gcd_dvd_left ..), Int.mul_assoc, + Int.tdiv_mul_cancel (Int.ofNat_dvd_left.2 <| Nat.gcd_dvd_left ..)] · rw [← Nat.mul_assoc, Nat.mul_right_comm, Nat.mul_right_comm (_/_), Nat.div_mul_cancel (Nat.gcd_dvd_right ..), Nat.mul_assoc, Nat.div_mul_cancel (Nat.gcd_dvd_right ..)] diff --git a/Batteries/Data/String/Lemmas.lean b/Batteries/Data/String/Lemmas.lean index ddd43d32a4..490c34908b 100644 --- a/Batteries/Data/String/Lemmas.lean +++ b/Batteries/Data/String/Lemmas.lean @@ -15,11 +15,11 @@ namespace String attribute [ext (iff := false)] ext theorem lt_trans {s₁ s₂ s₃ : String} : s₁ < s₂ → s₂ < s₃ → s₁ < s₃ := - List.lt_trans' (α := Char) Nat.lt_trans + List.lt_trans (α := Char) Nat.lt_trans (fun h1 h2 => Nat.not_lt.2 <| Nat.le_trans (Nat.not_lt.1 h2) (Nat.not_lt.1 h1)) theorem lt_antisymm {s₁ s₂ : String} (h₁ : ¬s₁ < s₂) (h₂ : ¬s₂ < s₁) : s₁ = s₂ := - ext <| List.lt_antisymm' (α := Char) + ext <| List.lt_antisymm (α := Char) (fun h1 h2 => Char.le_antisymm (Nat.not_lt.1 h2) (Nat.not_lt.1 h1)) h₁ h₂ instance : Batteries.TransOrd String := .compareOfLessAndEq @@ -209,7 +209,7 @@ theorem next_of_valid (cs : List Char) (c : Char) (cs' : List Char) : next ⟨cs ++ c :: cs'⟩ ⟨utf8Len cs⟩ = ⟨utf8Len cs + c.utf8Size⟩ := next_of_valid' .. @[simp] theorem atEnd_iff (s : String) (p : Pos) : atEnd s p ↔ s.endPos ≤ p := - decide_eq_true_iff _ + decide_eq_true_iff theorem valid_next {p : Pos} (h : p.Valid s) (h₂ : p < s.endPos) : (next s p).Valid s := by match s, p, h with diff --git a/Batteries/Data/UnionFind/Basic.lean b/Batteries/Data/UnionFind/Basic.lean index 549c9e1cde..4fff679993 100644 --- a/Batteries/Data/UnionFind/Basic.lean +++ b/Batteries/Data/UnionFind/Basic.lean @@ -128,7 +128,7 @@ abbrev parent (self : UnionFind) (i : Nat) : Nat := parentD self.arr i theorem parent'_lt (self : UnionFind) (i : Fin self.size) : (self.arr.get i).parent < self.size := by - simp only [← parentD_eq, parentD_lt, Fin.is_lt, Array.data_length] + simp only [← parentD_eq, parentD_lt, Fin.is_lt, Array.length_toList] theorem parent_lt (self : UnionFind) (i : Nat) : self.parent i < self.size ↔ i < self.size := by simp only [parentD]; split <;> simp only [*, parent'_lt] @@ -151,8 +151,8 @@ theorem rank'_lt_rankMax (self : UnionFind) (i : Fin self.size) : let rec go : ∀ {l} {x : UFNode}, x ∈ l → x.rank ≤ List.foldr (max ·.rank) 0 l | a::l, _, List.Mem.head _ => by dsimp; apply Nat.le_max_left | a::l, _, .tail _ h => by dsimp; exact Nat.le_trans (go h) (Nat.le_max_right ..) - simp only [Array.get_eq_getElem, rankMax, Array.foldr_eq_foldr_data] - exact Nat.lt_succ.2 <| go (self.arr.data.get_mem i.1 i.2) + simp only [Array.get_eq_getElem, rankMax, Array.foldr_eq_foldr_toList] + exact Nat.lt_succ.2 <| go (self.arr.toList.get_mem i.1 i.2) theorem rankD_lt_rankMax (self : UnionFind) (i : Nat) : rankD self.arr i < self.rankMax := by @@ -217,7 +217,7 @@ theorem parent_rootD (self : UnionFind) (x : Nat) : @[nolint unusedHavesSuffices] theorem rootD_parent (self : UnionFind) (x : Nat) : self.rootD (self.parent x) = self.rootD x := by - simp only [rootD, Array.data_length, parent_lt] + simp only [rootD, Array.length_toList, parent_lt] split · simp only [parentD, ↓reduceDIte, *] (conv => rhs; rw [root]); split @@ -226,7 +226,7 @@ theorem rootD_parent (self : UnionFind) (x : Nat) : self.rootD (self.parent x) = · simp only [not_false_eq_true, parentD_of_not_lt, *] theorem rootD_lt {self : UnionFind} {x : Nat} : self.rootD x < self.size ↔ x < self.size := by - simp only [rootD, Array.data_length]; split <;> simp [*] + simp only [rootD, Array.length_toList]; split <;> simp [*] @[nolint unusedHavesSuffices] theorem rootD_eq_self {self : UnionFind} {x : Nat} : self.rootD x = x ↔ self.parent x = x := by @@ -284,7 +284,7 @@ termination_by self.rankMax - self.rank x theorem findAux_root {self : UnionFind} {x : Fin self.size} : (findAux self x).root = self.root x := by rw [findAux, root] - simp only [Array.data_length, Array.get_eq_getElem, dite_eq_ite] + simp only [Array.length_toList, Array.get_eq_getElem, dite_eq_ite] split <;> simp only have := Nat.sub_lt_sub_left (self.lt_rankMax x) (self.rank'_lt _ ‹_›) exact findAux_root @@ -298,7 +298,7 @@ theorem findAux_s {self : UnionFind} {x : Fin self.size} : rw [show self.rootD _ = (self.findAux ⟨_, self.parent'_lt x⟩).root from _] · rw [findAux]; split <;> rfl · rw [← rootD_parent, parent, parentD_eq] - simp only [rootD, Array.get_eq_getElem, Array.data_length, findAux_root] + simp only [rootD, Array.get_eq_getElem, Array.length_toList, findAux_root] apply dif_pos exact parent'_lt .. @@ -309,8 +309,7 @@ theorem rankD_findAux {self : UnionFind} {x : Fin self.size} : rw [findAux_s]; split <;> [rfl; skip] have := Nat.sub_lt_sub_left (self.lt_rankMax x) (self.rank'_lt _ ‹_›) have := lt_of_parentD (by rwa [parentD_eq]) - rw [rankD_eq' (by simp [FindAux.size_eq, h])] - rw [Array.get_modify (by rwa [FindAux.size_eq])] + rw [rankD_eq' (by simp [FindAux.size_eq, h]), Array.get_modify] split <;> simp [← rankD_eq, rankD_findAux (x := ⟨_, self.parent'_lt x⟩), -Array.get_eq_getElem] else simp only [rankD, Array.data_length, Array.get_eq_getElem, rank] @@ -364,7 +363,7 @@ theorem parentD_findAux_or (self : UnionFind) (x : Fin self.size) (i) : · simp [*] · have := Nat.sub_lt_sub_left (self.lt_rankMax x) (self.rank'_lt _ ‹_›) exact (parentD_findAux_or self ⟨_, self.parent'_lt x⟩ i).imp_left <| .imp_right fun h => by - simp only [h, ← parentD_eq, rootD_parent, Array.data_length] + simp only [h, ← parentD_eq, rootD_parent, Array.length_toList] termination_by self.rankMax - self.rank x theorem lt_rankD_findAux {self : UnionFind} {x : Fin self.size} : @@ -386,7 +385,7 @@ def find (self : UnionFind) (x : Fin self.size) : { 1.arr := r.s 2.1.val := r.root 1.parentD_lt := fun h => by - simp only [Array.data_length, FindAux.size_eq] at * + simp only [Array.length_toList, FindAux.size_eq] at * exact parentD_findAux_lt h 1.rankD_lt := fun h => by rw [rankD_findAux, rankD_findAux]; exact lt_rankD_findAux h 2.1.isLt := show _ < r.s.size by rw [r.size_eq]; exact r.root.2 @@ -420,7 +419,7 @@ def findD (self : UnionFind) (x : Nat) : UnionFind × Nat := @[simp] theorem find_parent_1 (self : UnionFind) (x : Fin self.size) : (self.find x).1.parent x = self.rootD x := by - simp only [parent, Array.data_length, find] + simp only [parent, Array.length_toList, find] rw [parentD_findAux, if_pos rfl] theorem find_parent_or (self : UnionFind) (x : Fin self.size) (i) : @@ -500,7 +499,7 @@ theorem setParent_rankD_lt {arr : Array UFNode} {x y : Fin arr.size} def link (self : UnionFind) (x y : Fin self.size) (yroot : self.parent y = y) : UnionFind where arr := linkAux self.arr x y parentD_lt h := by - simp only [Array.data_length, linkAux_size] at * + simp only [Array.length_toList, linkAux_size] at * simp only [linkAux, Array.get_eq_getElem] split <;> [skip; split <;> [skip; split]] · exact self.parentD_lt h @@ -522,7 +521,7 @@ def link (self : UnionFind) (x y : Fin self.size) (yroot : self.parent y = y) : simp only [rankD_set, Fin.eta, Array.get_eq_getElem] split · simp_all - · simp_all only [Array.get_eq_getElem, Array.data_length, Nat.lt_irrefl, not_false_eq_true, + · simp_all only [Array.get_eq_getElem, Array.length_toList, Nat.lt_irrefl, not_false_eq_true, and_true, ite_false, ite_eq_right_iff] rintro rfl simp [rankD_eq, *] diff --git a/Batteries/Data/Vector/Basic.lean b/Batteries/Data/Vector/Basic.lean index 7726a5b9e5..534f94030a 100644 --- a/Batteries/Data/Vector/Basic.lean +++ b/Batteries/Data/Vector/Basic.lean @@ -263,7 +263,7 @@ alias take := shrink if and only if `p a[i] b[i]` holds true for all valid indices `i`. -/ @[inline] def isEqv (a b : Vector α n) (p : α → α → Bool) : Bool := - Array.isEqvAux a.toArray b.toArray (a.size_eq.trans b.size_eq.symm) p 0 + Array.isEqvAux a.toArray b.toArray (a.size_eq.trans b.size_eq.symm) p 0 (Nat.zero_le _) instance [BEq α] : BEq (Vector α n) := ⟨fun a b => isEqv a b BEq.beq⟩ diff --git a/Batteries/Lean/HashSet.lean b/Batteries/Lean/HashSet.lean index 4b6df398b3..2a6ee47909 100644 --- a/Batteries/Lean/HashSet.lean +++ b/Batteries/Lean/HashSet.lean @@ -23,13 +23,6 @@ def anyM [Monad m] (s : HashSet α) (f : α → m Bool) : m Bool := do return true return false -/-- -`O(n)`. Returns `true` if `f` returns `true` for any element of the set. --/ -@[inline] -def any (s : HashSet α) (f : α → Bool) : Bool := - Id.run <| s.anyM f - /-- `O(n)`. Returns `true` if `f` returns `true` for all elements of the set. -/ @@ -40,13 +33,6 @@ def allM [Monad m] (s : HashSet α) (f : α → m Bool) : m Bool := do return false return true -/-- -`O(n)`. Returns `true` if `f` returns `true` for all elements of the set. --/ -@[inline] -def all (s : HashSet α) (f : α → Bool) : Bool := - Id.run <| s.allM f - instance : BEq (HashSet α) where beq s t := s.all (t.contains ·) && t.all (s.contains ·) @@ -59,10 +45,3 @@ def insert' (s : HashSet α) (a : α) : HashSet α × Bool := let oldSize := s.size let s := s.insert a (s, s.size == oldSize) - -/-- -`O(n)`. Obtain a `HashSet` from an array. --/ -@[inline] -protected def ofArray [BEq α] [Hashable α] (as : Array α) : HashSet α := - HashSet.empty.insertMany as diff --git a/Batteries/Lean/Meta/AssertHypotheses.lean b/Batteries/Lean/Meta/AssertHypotheses.lean deleted file mode 100644 index 6275032063..0000000000 --- a/Batteries/Lean/Meta/AssertHypotheses.lean +++ /dev/null @@ -1,40 +0,0 @@ -/- -Copyright (c) 2022 Jannis Limperg. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jannis Limperg --/ - -import Lean.Meta.Tactic.Assert - -open Lean Lean.Meta - -namespace Lean.Meta - -/-- -Description of a hypothesis for `Lean.MVarId.assertHypotheses'`. --/ -structure Hypothesis' extends Hypothesis where - /-- The hypothesis' `BinderInfo` -/ - binderInfo : BinderInfo - /-- The hypothesis' `LocalDeclKind` -/ - kind : LocalDeclKind - -/-- -Add the given hypotheses to the local context. This is a generalisation of -`Lean.MVarId.assertHypotheses` which lets you specify --/ -def _root_.Lean.MVarId.assertHypotheses' (mvarId : MVarId) - (hs : Array Hypothesis') : MetaM (Array FVarId × MVarId) := do - let (fvarIds, mvarId) ← mvarId.assertHypotheses $ hs.map (·.toHypothesis) - mvarId.modifyLCtx fun lctx => Id.run do - let mut lctx := lctx - for h : i in [:hs.size] do - let h := hs[i] - if h.kind != .default then - lctx := lctx.setKind fvarIds[i]! h.kind - if h.binderInfo != .default then - lctx := lctx.setBinderInfo fvarIds[i]! h.binderInfo - pure lctx - return (fvarIds, mvarId) - -end Lean.Meta diff --git a/Batteries/Lean/Meta/Clear.lean b/Batteries/Lean/Meta/Clear.lean deleted file mode 100644 index 66f38b7798..0000000000 --- a/Batteries/Lean/Meta/Clear.lean +++ /dev/null @@ -1,26 +0,0 @@ -/- -Copyright (c) 2022 Jannis Limperg. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jannis Limperg --/ - -import Batteries.Lean.Meta.Basic -import Lean.Meta.Tactic.Clear - -open Lean Lean.Meta - -/-- -Try to clear the given fvars from the local context. Returns the new goal and -the hypotheses that were cleared. Unlike `Lean.MVarId.tryClearMany`, this -function does not require the `hyps` to be given in the order in which they -appear in the local context. --/ -def Lean.MVarId.tryClearMany' (goal : MVarId) (hyps : Array FVarId) : - MetaM (MVarId × Array FVarId) := - goal.withContext do - let hyps ← sortFVarsByContextOrder hyps - hyps.foldrM (init := (goal, Array.mkEmpty hyps.size)) - fun h (goal, cleared) => do - let goal' ← goal.tryClear h - let cleared := if goal == goal' then cleared else cleared.push h - return (goal', cleared) diff --git a/lean-toolchain b/lean-toolchain index 89985206ac..a00797801f 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.12.0 +leanprover/lean4:v4.13.0-rc1 diff --git a/test/lint_lean.lean b/test/lint_lean.lean index 46eb6591ff..f5657a5b96 100644 --- a/test/lint_lean.lean +++ b/test/lint_lean.lean @@ -15,14 +15,13 @@ but it is useful to run locally to see what the linters would catch. -- attribute [nolint dupNamespace] Lean.Elab.Tactic.Tactic -- attribute [nolint dupNamespace] Lean.Parser.Parser Lean.Parser.Parser.rec Lean.Parser.Parser.mk -- Lean.Parser.Parser.info Lean.Parser.Parser.fn +-- attribute [nolint explicitVarsOfIff] Iff.refl /-! Failing lints that need work. -/ --- #lint only explicitVarsOfIff in all -- Found 156 errors - --- Many fixed in https://github.com/leanprover/lean4/pull/4620 +-- Many fixed in https://github.com/leanprover/lean4/pull/4620 and subsequent PRs -- and should be checked again. --- #lint only simpNF in all -- Found 12 errors +-- #lint only simpNF in all -- Found 22 errors /-! Lints that fail, but that we're not intending to do anything about. -/ @@ -41,6 +40,7 @@ but it is useful to run locally to see what the linters would catch. /-! Lints that have succeeded in the past, and hopefully still do! -/ +-- #lint only explicitVarsOfIff in all -- Found 1 errors, `Iff.refl`, which could be nolinted. -- #lint only impossibleInstance in all -- Found 0 errors -- #lint only simpVarHead in all -- Found 0 error -- #lint only unusedHavesSuffices in all -- Found 0 errors From 63c1c38b123b0741b7b7fd56fb8510f95bfd0e55 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Thu, 3 Oct 2024 14:21:23 +0200 Subject: [PATCH 105/177] fix: botched merge --- Batteries/Data/HashMap/WF.lean | 22 ++++++++++------------ 1 file changed, 10 insertions(+), 12 deletions(-) diff --git a/Batteries/Data/HashMap/WF.lean b/Batteries/Data/HashMap/WF.lean index 11b68de28d..c95e3ad133 100644 --- a/Batteries/Data/HashMap/WF.lean +++ b/Batteries/Data/HashMap/WF.lean @@ -387,15 +387,13 @@ variable {_ : BEq α} {_ : Hashable α} @[inline] def mapVal (f : α → β → γ) (self : HashMap α β) : HashMap α γ := ⟨self.1.mapVal f, self.2.mapVal⟩ --- Temporarily removed on lean-pr-testing-5403. - --- /-- --- Applies `f` to each key-value pair `a, b` in the map. If it returns `some c` then --- `a, c` is pushed into the new map; else the key is removed from the map. --- -/ --- @[inline] def filterMap (f : α → β → Option γ) (self : HashMap α β) : HashMap α γ := --- ⟨self.1.filterMap f, self.2.filterMap⟩ - --- /-- Constructs a map with the set of all pairs `a, b` such that `f` returns true. -/ --- @[inline] def filter (f : α → β → Bool) (self : HashMap α β) : HashMap α β := --- self.filterMap fun a b => bif f a b then some b else none +/-- +Applies `f` to each key-value pair `a, b` in the map. If it returns `some c` then +`a, c` is pushed into the new map; else the key is removed from the map. +-/ +@[inline] def filterMap (f : α → β → Option γ) (self : HashMap α β) : HashMap α γ := + ⟨self.1.filterMap f, self.2.filterMap⟩ + +/-- Constructs a map with the set of all pairs `a, b` such that `f` returns true. -/ +@[inline] def filter (f : α → β → Bool) (self : HashMap α β) : HashMap α β := + self.filterMap fun a b => bif f a b then some b else none From 13f9b00769bdac2c0041406a6c2524a361e8d660 Mon Sep 17 00:00:00 2001 From: Mac Malone Date: Fri, 4 Oct 2024 02:53:03 -0400 Subject: [PATCH 106/177] fix: use user's Lean search path in linter (#980) --- scripts/runLinter.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/runLinter.lean b/scripts/runLinter.lean index 89976f0e47..e22dc36da5 100644 --- a/scripts/runLinter.lean +++ b/scripts/runLinter.lean @@ -37,7 +37,7 @@ unsafe def main (args : List String) : IO Unit := do | name => some name | _ => none | IO.eprintln "Usage: runLinter [--update] [Batteries.Data.Nat.Basic]" *> IO.Process.exit 1 - searchPathRef.set compile_time_search_path% + initSearchPath (← findSysroot) let mFile ← findOLean module unless (← mFile.pathExists) do -- run `lake build module` (and ignore result) if the file hasn't been built yet From daf1ed91789811cf6bbb7bf2f4dad6b3bad8fbf4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sat, 5 Oct 2024 10:58:02 -0400 Subject: [PATCH 107/177] chore: move to v4.13.0-rc3 (#981) --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index a00797801f..eff86fd63d 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.13.0-rc1 +leanprover/lean4:v4.13.0-rc3 From 5ac298e7c5ab9d24ba15967d28818a34a21a46c0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 14 Oct 2024 06:02:34 +0200 Subject: [PATCH 108/177] =?UTF-8?q?feat:=20`l=20<+~=20[]=20=E2=86=94=20l?= =?UTF-8?q?=20=3D=20[]`=20(#972)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: F. G. Dorais --- Batteries/Data/List/Perm.lean | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/Batteries/Data/List/Perm.lean b/Batteries/Data/List/Perm.lean index d924ebf4dd..e72ab8f624 100644 --- a/Batteries/Data/List/Perm.lean +++ b/Batteries/Data/List/Perm.lean @@ -25,7 +25,7 @@ open Perm (swap) section Subperm -theorem nil_subperm {l : List α} : [] <+~ l := ⟨[], Perm.nil, by simp⟩ +@[simp] theorem nil_subperm {l : List α} : [] <+~ l := ⟨[], Perm.nil, by simp⟩ theorem Perm.subperm_left {l l₁ l₂ : List α} (p : l₁ ~ l₂) : l <+~ l₁ ↔ l <+~ l₂ := suffices ∀ {l₁ l₂ : List α}, l₁ ~ l₂ → l <+~ l₁ → l <+~ l₂ from ⟨this p, this p.symm⟩ @@ -47,6 +47,8 @@ theorem Subperm.trans {l₁ l₂ l₃ : List α} (s₁₂ : l₁ <+~ l₂) (s₂ let ⟨l₁', p₁, s₁⟩ := p₂.subperm_left.2 s₁₂ ⟨l₁', p₁, s₁.trans s₂⟩ +theorem Subperm.cons_self : l <+~ a :: l := ⟨l, .refl _, sublist_cons_self ..⟩ + theorem Subperm.cons_right {α : Type _} {l l' : List α} (x : α) (h : l <+~ l') : l <+~ x :: l' := h.trans (sublist_cons_self x l').subperm @@ -67,6 +69,9 @@ theorem Subperm.filter (p : α → Bool) ⦃l l' : List α⦄ (h : l <+~ l') : let ⟨xs, hp, h⟩ := h exact ⟨_, hp.filter p, h.filter p⟩ +@[simp] theorem subperm_nil : l <+~ [] ↔ l = [] := + ⟨fun h => length_eq_zero.1 $ Nat.le_zero.1 h.length_le, by rintro rfl; rfl⟩ + @[simp] theorem singleton_subperm_iff {α} {l : List α} {a : α} : [a] <+~ l ↔ a ∈ l := by refine ⟨fun ⟨s, hla, h⟩ => ?_, fun h => ⟨[a], .rfl, singleton_sublist.mpr h⟩⟩ rwa [perm_singleton.mp hla, singleton_sublist] at h From c8dcca4553b2c8da9b3c03ff75e26de95fb6cb7f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 14 Oct 2024 01:44:48 -0400 Subject: [PATCH 109/177] feat: add bisection algorithm (#890) --- Batteries/Data/Nat.lean | 1 + Batteries/Data/Nat/Bisect.lean | 137 +++++++++++++++++++++++++++++++++ 2 files changed, 138 insertions(+) create mode 100644 Batteries/Data/Nat/Bisect.lean diff --git a/Batteries/Data/Nat.lean b/Batteries/Data/Nat.lean index f1a6ddcca1..dbf3161213 100644 --- a/Batteries/Data/Nat.lean +++ b/Batteries/Data/Nat.lean @@ -1,3 +1,4 @@ import Batteries.Data.Nat.Basic +import Batteries.Data.Nat.Bisect import Batteries.Data.Nat.Gcd import Batteries.Data.Nat.Lemmas diff --git a/Batteries/Data/Nat/Bisect.lean b/Batteries/Data/Nat/Bisect.lean new file mode 100644 index 0000000000..f692675944 --- /dev/null +++ b/Batteries/Data/Nat/Bisect.lean @@ -0,0 +1,137 @@ +/- +Copyright (c) 2024 François G. Dorais. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: François G. Dorais +-/ +import Batteries.Tactic.Basic +import Batteries.Data.Nat.Basic + +namespace Nat + +/-- Average of two natural numbers rounded toward zero. -/ +abbrev avg (a b : Nat) := (a + b) / 2 + +theorem avg_comm (a b : Nat) : avg a b = avg b a := by + rw [avg, Nat.add_comm] + +theorem avg_le_left (h : b ≤ a) : avg a b ≤ a := by + apply Nat.div_le_of_le_mul; simp_arith [*] + +theorem avg_le_right (h : a ≤ b) : avg a b ≤ b := by + apply Nat.div_le_of_le_mul; simp_arith [*] + +theorem avg_lt_left (h : b < a) : avg a b < a := by + apply Nat.div_lt_of_lt_mul; omega + +theorem avg_lt_right (h : a < b) : avg a b < b := by + apply Nat.div_lt_of_lt_mul; omega + +theorem le_avg_left (h : a ≤ b) : a ≤ avg a b := by + apply (Nat.le_div_iff_mul_le Nat.zero_lt_two).mpr; simp_arith [*] + +theorem le_avg_right (h : b ≤ a) : b ≤ avg a b := by + apply (Nat.le_div_iff_mul_le Nat.zero_lt_two).mpr; simp_arith [*] + +theorem le_add_one_of_avg_eq_left (h : avg a b = a) : b ≤ a + 1 := by + cases Nat.lt_or_ge b (a+2) with + | inl hlt => exact Nat.le_of_lt_add_one hlt + | inr hge => + absurd Nat.lt_irrefl a + conv => rhs; rw [← h] + rw [← Nat.add_one_le_iff, Nat.le_div_iff_mul_le Nat.zero_lt_two] + omega + +theorem le_add_one_of_avg_eq_right (h : avg a b = b) : a ≤ b + 1 := by + cases Nat.lt_or_ge a (b+2) with + | inl hlt => exact Nat.le_of_lt_add_one hlt + | inr hge => + absurd Nat.lt_irrefl b + conv => rhs; rw [← h] + rw [← Nat.add_one_le_iff, Nat.le_div_iff_mul_le Nat.zero_lt_two] + omega + +/-- +Given natural numbers `a < b` such that `p a = true` and `p b = false`, `bisect` finds a natural +number `a ≤ c < b` such that `p c = true` and `p (c+1) = false`. +-/ +def bisect {p : Nat → Bool} (h : start < stop) (hstart : p start = true) (hstop : p stop = false) := + let mid := avg start stop + have hmidstop : mid < stop := by apply Nat.div_lt_of_lt_mul; omega + if hstartmid : start < mid then + match hmid : p mid with + | false => bisect hstartmid hstart hmid + | true => bisect hmidstop hmid hstop + else + mid +termination_by stop - start + +theorem bisect_lt_stop {p : Nat → Bool} (h : start < stop) (hstart : p start = true) + (hstop : p stop = false) : bisect h hstart hstop < stop := by + unfold bisect + simp only; split + · split + · next h' _ => + have : avg start stop - start < stop - start := by + apply Nat.sub_lt_sub_right + · exact Nat.le_of_lt h' + · exact Nat.avg_lt_right h + apply Nat.lt_trans + · exact bisect_lt_stop .. + · exact avg_lt_right h + · exact bisect_lt_stop .. + · exact avg_lt_right h + +theorem start_le_bisect {p : Nat → Bool} (h : start < stop) (hstart : p start = true) + (hstop : p stop = false) : start ≤ bisect h hstart hstop := by + unfold bisect + simp only; split + · split + · next h' _ => + have : avg start stop - start < stop - start := by + apply Nat.sub_lt_sub_right + · exact Nat.le_of_lt h' + · exact avg_lt_right h + exact start_le_bisect .. + · next h' _ => + apply Nat.le_trans + · exact Nat.le_of_lt h' + · exact start_le_bisect .. + · exact le_avg_left (Nat.le_of_lt h) + +theorem bisect_true {p : Nat → Bool} (h : start < stop) (hstart : p start = true) + (hstop : p stop = false) : p (bisect h hstart hstop) = true := by + unfold bisect + simp only; split + · split + · have : avg start stop - start < stop - start := by + apply Nat.sub_lt_sub_right + · exact Nat.le_avg_left (Nat.le_of_lt h) + · exact Nat.avg_lt_right h + exact bisect_true .. + · exact bisect_true .. + · next h' => + rw [← hstart]; congr + apply Nat.le_antisymm + · exact Nat.le_of_not_gt h' + · exact Nat.le_avg_left (Nat.le_of_lt h) + +theorem bisect_add_one_false {p : Nat → Bool} (h : start < stop) (hstart : p start = true) + (hstop : p stop = false) : p (bisect h hstart hstop + 1) = false := by + unfold bisect + simp only; split + · split + · have : avg start stop - start < stop - start := by + apply Nat.sub_lt_sub_right + · exact Nat.le_avg_left (Nat.le_of_lt h) + · exact Nat.avg_lt_right h + exact bisect_add_one_false .. + · exact bisect_add_one_false .. + · next h' => + have heq : avg start stop = start := by + apply Nat.le_antisymm + · exact Nat.le_of_not_gt h' + · exact Nat.le_avg_left (Nat.le_of_lt h) + rw [← hstop, heq]; congr + apply Nat.le_antisymm + · exact Nat.succ_le_of_lt h + · exact Nat.le_add_one_of_avg_eq_left heq From 405f9496b3cc4e19b72a8c3f5e889f99abd543cd Mon Sep 17 00:00:00 2001 From: Bulhwi Cha Date: Mon, 14 Oct 2024 05:45:47 +0000 Subject: [PATCH 110/177] refactor: move theorems about lists from mathlib (#756) Co-authored-by: Kim Morrison --- Batteries/Data/List/Lemmas.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 1ac4ae3f57..96a79c5bfa 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -15,6 +15,10 @@ namespace List @[simp] theorem getElem_mk {xs : List α} {i : Nat} (h : i < xs.length) : (Array.mk xs)[i] = xs[i] := rfl +/-! ### isEmpty -/ + +theorem isEmpty_iff_eq_nil {l : List α} : l.isEmpty ↔ l = [] := by cases l <;> simp [isEmpty] + /-! ### next? -/ @[simp] theorem next?_nil : @next? α [] = none := rfl @@ -48,6 +52,11 @@ theorem get?_inj apply getElem?_inj h₀ h₁ simp_all +/-! ### modifyHead -/ + +@[simp] theorem modifyHead_modifyHead (l : List α) (f g : α → α) : + (l.modifyHead f).modifyHead g = l.modifyHead (g ∘ f) := by cases l <;> simp [modifyHead] + /-! ### modifyNth -/ @[simp] theorem modifyNth_nil (f : α → α) (n) : [].modifyNth f n = [] := by cases n <;> rfl From 9efd9c267ad7a71c5e3a83e8fbbd446fe61ef119 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 14 Oct 2024 01:49:34 -0400 Subject: [PATCH 111/177] feat: `Fin.foldlM` and `Fin.foldrM` (#814) --- Batteries/Data/Fin/Basic.lean | 54 ++++++++++++++ Batteries/Data/Fin/Lemmas.lean | 132 ++++++++++++++++++++++++++------- 2 files changed, 159 insertions(+), 27 deletions(-) diff --git a/Batteries/Data/Fin/Basic.lean b/Batteries/Data/Fin/Basic.lean index 346f3006e5..b61481e33a 100644 --- a/Batteries/Data/Fin/Basic.lean +++ b/Batteries/Data/Fin/Basic.lean @@ -14,3 +14,57 @@ def enum (n) : Array (Fin n) := Array.ofFn id /-- `list n` is the list of all elements of `Fin n` in order -/ def list (n) : List (Fin n) := (enum n).toList + +/-- +Folds a monadic function over `Fin n` from left to right: +``` +Fin.foldlM n f x₀ = do + let x₁ ← f x₀ 0 + let x₂ ← f x₁ 1 + ... + let xₙ ← f xₙ₋₁ (n-1) + pure xₙ +``` +-/ +@[inline] def foldlM [Monad m] (n) (f : α → Fin n → m α) (init : α) : m α := loop init 0 where + /-- + Inner loop for `Fin.foldlM`. + ``` + Fin.foldlM.loop n f xᵢ i = do + let xᵢ₊₁ ← f xᵢ i + ... + let xₙ ← f xₙ₋₁ (n-1) + pure xₙ + ``` + -/ + loop (x : α) (i : Nat) : m α := do + if h : i < n then f x ⟨i, h⟩ >>= (loop · (i+1)) else pure x + termination_by n - i + +/-- +Folds a monadic function over `Fin n` from right to left: +``` +Fin.foldrM n f xₙ = do + let xₙ₋₁ ← f (n-1) xₙ + let xₙ₋₂ ← f (n-2) xₙ₋₁ + ... + let x₀ ← f 0 x₁ + pure x₀ +``` +-/ +@[inline] def foldrM [Monad m] (n) (f : Fin n → α → m α) (init : α) : m α := + loop ⟨n, Nat.le_refl n⟩ init where + /-- + Inner loop for `Fin.foldrM`. + ``` + Fin.foldrM.loop n f i xᵢ = do + let xᵢ₋₁ ← f (i-1) xᵢ + ... + let x₁ ← f 1 x₂ + let x₀ ← f 0 x₁ + pure x₀ + ``` + -/ + loop : {i // i ≤ n} → α → m α + | ⟨0, _⟩, x => pure x + | ⟨i+1, h⟩, x => f ⟨i, h⟩ x >>= loop ⟨i, Nat.le_of_lt h⟩ diff --git a/Batteries/Data/Fin/Lemmas.lean b/Batteries/Data/Fin/Lemmas.lean index 33534b4549..5010e1310f 100644 --- a/Batteries/Data/Fin/Lemmas.lean +++ b/Batteries/Data/Fin/Lemmas.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Batteries.Data.Fin.Basic +import Batteries.Data.List.Lemmas namespace Fin @@ -53,28 +54,102 @@ theorem list_reverse (n) : (list n).reverse = (list n).map rev := by conv => rhs; rw [list_succ] simp [← List.map_reverse, ih, Function.comp_def, rev_succ] +/-! ### foldlM -/ + +theorem foldlM_loop_lt [Monad m] (f : α → Fin n → m α) (x) (h : i < n) : + foldlM.loop n f x i = f x ⟨i, h⟩ >>= (foldlM.loop n f . (i+1)) := by + rw [foldlM.loop, dif_pos h] + +theorem foldlM_loop_eq [Monad m] (f : α → Fin n → m α) (x) : foldlM.loop n f x n = pure x := by + rw [foldlM.loop, dif_neg (Nat.lt_irrefl _)] + +theorem foldlM_loop [Monad m] (f : α → Fin (n+1) → m α) (x) (h : i < n+1) : + foldlM.loop (n+1) f x i = f x ⟨i, h⟩ >>= (foldlM.loop n (fun x j => f x j.succ) . i) := by + if h' : i < n then + rw [foldlM_loop_lt _ _ h] + congr; funext + rw [foldlM_loop_lt _ _ h', foldlM_loop]; rfl + else + cases Nat.le_antisymm (Nat.le_of_lt_succ h) (Nat.not_lt.1 h') + rw [foldlM_loop_lt] + congr; funext + rw [foldlM_loop_eq, foldlM_loop_eq] +termination_by n - i + +@[simp] theorem foldlM_zero [Monad m] (f : α → Fin 0 → m α) (x) : foldlM 0 f x = pure x := + foldlM_loop_eq .. + +theorem foldlM_succ [Monad m] (f : α → Fin (n+1) → m α) (x) : + foldlM (n+1) f x = f x 0 >>= foldlM n (fun x j => f x j.succ) := foldlM_loop .. + +theorem foldlM_eq_foldlM_list [Monad m] (f : α → Fin n → m α) (x) : + foldlM n f x = (list n).foldlM f x := by + induction n generalizing x with + | zero => simp + | succ n ih => + rw [foldlM_succ, list_succ, List.foldlM_cons] + congr; funext + rw [List.foldlM_map, ih] + +/-! ### foldrM -/ + +theorem foldrM_loop_zero [Monad m] (f : Fin n → α → m α) (x) : + foldrM.loop n f ⟨0, Nat.zero_le _⟩ x = pure x := by + rw [foldrM.loop] + +theorem foldrM_loop_succ [Monad m] (f : Fin n → α → m α) (x) (h : i < n) : + foldrM.loop n f ⟨i+1, h⟩ x = f ⟨i, h⟩ x >>= foldrM.loop n f ⟨i, Nat.le_of_lt h⟩ := by + rw [foldrM.loop] + +theorem foldrM_loop [Monad m] [LawfulMonad m] (f : Fin (n+1) → α → m α) (x) (h : i+1 ≤ n+1) : + foldrM.loop (n+1) f ⟨i+1, h⟩ x = + foldrM.loop n (fun j => f j.succ) ⟨i, Nat.le_of_succ_le_succ h⟩ x >>= f 0 := by + induction i generalizing x with + | zero => + rw [foldrM_loop_zero, foldrM_loop_succ, pure_bind] + conv => rhs; rw [←bind_pure (f 0 x)] + congr; funext; exact foldrM_loop_zero .. + | succ i ih => + rw [foldrM_loop_succ, foldrM_loop_succ, bind_assoc] + congr; funext; exact ih .. + +@[simp] theorem foldrM_zero [Monad m] (f : Fin 0 → α → m α) (x) : foldrM 0 f x = pure x := + foldrM_loop_zero .. + +theorem foldrM_succ [Monad m] [LawfulMonad m] (f : Fin (n+1) → α → m α) (x) : + foldrM (n+1) f x = foldrM n (fun i => f i.succ) x >>= f 0 := foldrM_loop .. + +theorem foldrM_eq_foldrM_list [Monad m] [LawfulMonad m] (f : Fin n → α → m α) (x) : + foldrM n f x = (list n).foldrM f x := by + induction n with + | zero => simp + | succ n ih => rw [foldrM_succ, ih, list_succ, List.foldrM_cons, List.foldrM_map] + /-! ### foldl -/ -theorem foldl_loop_lt (f : α → Fin n → α) (x) (h : m < n) : - foldl.loop n f x m = foldl.loop n f (f x ⟨m, h⟩) (m+1) := by +theorem foldl_loop_lt (f : α → Fin n → α) (x) (h : i < n) : + foldl.loop n f x i = foldl.loop n f (f x ⟨i, h⟩) (i+1) := by rw [foldl.loop, dif_pos h] theorem foldl_loop_eq (f : α → Fin n → α) (x) : foldl.loop n f x n = x := by rw [foldl.loop, dif_neg (Nat.lt_irrefl _)] -theorem foldl_loop (f : α → Fin (n+1) → α) (x) (h : m < n+1) : - foldl.loop (n+1) f x m = foldl.loop n (fun x i => f x i.succ) (f x ⟨m, h⟩) m := by - if h' : m < n then - rw [foldl_loop_lt _ _ h, foldl_loop_lt _ _ h', foldl_loop]; rfl +theorem foldl_loop (f : α → Fin (n+1) → α) (x) (h : i < n+1) : + foldl.loop (n+1) f x i = foldl.loop n (fun x j => f x j.succ) (f x ⟨i, h⟩) i := by + if h' : i < n then + rw [foldl_loop_lt _ _ h] + rw [foldl_loop_lt _ _ h', foldl_loop]; rfl else cases Nat.le_antisymm (Nat.le_of_lt_succ h) (Nat.not_lt.1 h') - rw [foldl_loop_lt, foldl_loop_eq, foldl_loop_eq] -termination_by n - m + rw [foldl_loop_lt] + rw [foldl_loop_eq, foldl_loop_eq] -@[simp] theorem foldl_zero (f : α → Fin 0 → α) (x) : foldl 0 f x = x := by simp [foldl, foldl.loop] +@[simp] theorem foldl_zero (f : α → Fin 0 → α) (x) : foldl 0 f x = x := + foldl_loop_eq .. theorem foldl_succ (f : α → Fin (n+1) → α) (x) : - foldl (n+1) f x = foldl n (fun x i => f x i.succ) (f x 0) := foldl_loop .. + foldl (n+1) f x = foldl n (fun x i => f x i.succ) (f x 0) := + foldl_loop .. theorem foldl_succ_last (f : α → Fin (n+1) → α) (x) : foldl (n+1) f x = f (foldl n (f · ·.castSucc) x) (last n) := by @@ -83,6 +158,10 @@ theorem foldl_succ_last (f : α → Fin (n+1) → α) (x) : | zero => simp [foldl_succ, Fin.last] | succ n ih => rw [foldl_succ, ih (f · ·.succ), foldl_succ]; simp [succ_castSucc] +theorem foldl_eq_foldlM (f : α → Fin n → α) (x) : + foldl n f x = foldlM (m:=Id) n f x := by + induction n generalizing x <;> simp [foldl_succ, foldlM_succ, *] + theorem foldl_eq_foldl_list (f : α → Fin n → α) (x) : foldl n f x = (list n).foldl f x := by induction n generalizing x with | zero => rw [foldl_zero, list_zero, List.foldl_nil] @@ -90,24 +169,21 @@ theorem foldl_eq_foldl_list (f : α → Fin n → α) (x) : foldl n f x = (list /-! ### foldr -/ -unseal foldr.loop in -theorem foldr_loop_zero (f : Fin n → α → α) (x) : foldr.loop n f ⟨0, Nat.zero_le _⟩ x = x := - rfl +theorem foldr_loop_zero (f : Fin n → α → α) (x) : + foldr.loop n f ⟨0, Nat.zero_le _⟩ x = x := by + rw [foldr.loop] -unseal foldr.loop in -theorem foldr_loop_succ (f : Fin n → α → α) (x) (h : m < n) : - foldr.loop n f ⟨m+1, h⟩ x = foldr.loop n f ⟨m, Nat.le_of_lt h⟩ (f ⟨m, h⟩ x) := - rfl +theorem foldr_loop_succ (f : Fin n → α → α) (x) (h : i < n) : + foldr.loop n f ⟨i+1, h⟩ x = foldr.loop n f ⟨i, Nat.le_of_lt h⟩ (f ⟨i, h⟩ x) := by + rw [foldr.loop] -theorem foldr_loop (f : Fin (n+1) → α → α) (x) (h : m+1 ≤ n+1) : - foldr.loop (n+1) f ⟨m+1, h⟩ x = - f 0 (foldr.loop n (fun i => f i.succ) ⟨m, Nat.le_of_succ_le_succ h⟩ x) := by - induction m generalizing x with - | zero => simp [foldr_loop_zero, foldr_loop_succ] - | succ m ih => rw [foldr_loop_succ, ih, foldr_loop_succ, Fin.succ] +theorem foldr_loop (f : Fin (n+1) → α → α) (x) (h : i+1 ≤ n+1) : + foldr.loop (n+1) f ⟨i+1, h⟩ x = + f 0 (foldr.loop n (fun j => f j.succ) ⟨i, Nat.le_of_succ_le_succ h⟩ x) := by + induction i generalizing x <;> simp [foldr_loop_zero, foldr_loop_succ, *] -@[simp] theorem foldr_zero (f : Fin 0 → α → α) (x) : - foldr 0 f x = x := foldr_loop_zero .. +@[simp] theorem foldr_zero (f : Fin 0 → α → α) (x) : foldr 0 f x = x := + foldr_loop_zero .. theorem foldr_succ (f : Fin (n+1) → α → α) (x) : foldr (n+1) f x = f 0 (foldr n (fun i => f i.succ) x) := foldr_loop .. @@ -118,13 +194,15 @@ theorem foldr_succ_last (f : Fin (n+1) → α → α) (x) : | zero => simp [foldr_succ, Fin.last] | succ n ih => rw [foldr_succ, ih (f ·.succ), foldr_succ]; simp [succ_castSucc] +theorem foldr_eq_foldrM (f : Fin n → α → α) (x) : + foldr n f x = foldrM (m:=Id) n f x := by + induction n <;> simp [foldr_succ, foldrM_succ, *] + theorem foldr_eq_foldr_list (f : Fin n → α → α) (x) : foldr n f x = (list n).foldr f x := by induction n with | zero => rw [foldr_zero, list_zero, List.foldr_nil] | succ n ih => rw [foldr_succ, ih, list_succ, List.foldr_cons, List.foldr_map] -/-! ### foldl/foldr -/ - theorem foldl_rev (f : Fin n → α → α) (x) : foldl n (fun x i => f i.rev x) x = foldr n f x := by induction n generalizing x with From 5f963d5d06cc3c2d3abd0806891133137a59d7eb Mon Sep 17 00:00:00 2001 From: Bulhwi Cha Date: Mon, 14 Oct 2024 07:38:45 +0000 Subject: [PATCH 112/177] chore: remove duplicate theorem about lists (#986) --- Batteries/Data/List/Lemmas.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 96a79c5bfa..6a956f9b1d 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -17,7 +17,7 @@ namespace List /-! ### isEmpty -/ -theorem isEmpty_iff_eq_nil {l : List α} : l.isEmpty ↔ l = [] := by cases l <;> simp [isEmpty] +@[deprecated (since := "2024-08-15")] alias isEmpty_iff_eq_nil := isEmpty_iff /-! ### next? -/ From 0d328c81cfd4d0d79b235bfbbe7e4ac5b0dd2b4c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 14 Oct 2024 06:40:05 -0400 Subject: [PATCH 113/177] chore: add missing simp for array size lemmas (#982) --- Batteries/Data/Array/Lemmas.lean | 24 ++++++++++++++---------- 1 file changed, 14 insertions(+), 10 deletions(-) diff --git a/Batteries/Data/Array/Lemmas.lean b/Batteries/Data/Array/Lemmas.lean index 3f28087212..e6b1cf385a 100644 --- a/Batteries/Data/Array/Lemmas.lean +++ b/Batteries/Data/Array/Lemmas.lean @@ -78,7 +78,7 @@ theorem toList_zipWith (f : α → β → γ) (as : Array α) (bs : Array β) : @[deprecated (since := "2024-09-09")] alias data_zipWith := toList_zipWith @[deprecated (since := "2024-08-13")] alias zipWith_eq_zipWith_data := data_zipWith -theorem size_zipWith (as : Array α) (bs : Array β) (f : α → β → γ) : +@[simp] theorem size_zipWith (as : Array α) (bs : Array β) (f : α → β → γ) : (as.zipWith bs f).size = min as.size bs.size := by rw [size_eq_length_toList, toList_zipWith, List.length_zipWith] @@ -88,7 +88,7 @@ theorem toList_zip (as : Array α) (bs : Array β) : @[deprecated (since := "2024-09-09")] alias data_zip := toList_zip @[deprecated (since := "2024-08-13")] alias zip_eq_zip_data := data_zip -theorem size_zip (as : Array α) (bs : Array β) : +@[simp] theorem size_zip (as : Array α) (bs : Array β) : (as.zip bs).size = min as.size bs.size := as.size_zipWith bs Prod.mk @@ -150,19 +150,23 @@ where /-! ### shrink -/ theorem size_shrink_loop (a : Array α) (n) : (shrink.loop n a).size = a.size - n := by - induction n generalizing a with simp[shrink.loop] - | succ n ih => - simp[ih] - omega + induction n generalizing a with simp only [shrink.loop, Nat.sub_zero] + | succ n ih => simp only [ih, size_pop]; omega -theorem size_shrink (a : Array α) (n) : (a.shrink n).size = min a.size n := by +@[simp] theorem size_shrink (a : Array α) (n) : (a.shrink n).size = min a.size n := by simp [shrink, size_shrink_loop] omega /-! ### set -/ -theorem size_set! (a : Array α) (i v) : (a.set! i v).size = a.size := by - rw [set!_is_setD, size_setD] +theorem size_set! (a : Array α) (i v) : (a.set! i v).size = a.size := by simp + +/-! ### swapAt -/ + +theorem size_swapAt (a : Array α) (x i) : (a.swapAt i x).snd.size = a.size := by simp + +@[simp] theorem size_swapAt! (a : Array α) (x) (h : i < a.size) : + (a.swapAt! i x).snd.size = a.size := by simp [h] /-! ### map -/ @@ -190,7 +194,7 @@ private theorem size_insertAt_loop (as : Array α) (i : Fin (as.size+1)) (j : Fi · rw [size_insertAt_loop, size_swap] · rfl -theorem size_insertAt (as : Array α) (i : Fin (as.size+1)) (v : α) : +@[simp] theorem size_insertAt (as : Array α) (i : Fin (as.size+1)) (v : α) : (as.insertAt i v).size = as.size + 1 := by rw [insertAt, size_insertAt_loop, size_push] From ad3ba5ff13913874b80146b54d0a4e5b9b739451 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 14 Oct 2024 06:40:35 -0400 Subject: [PATCH 114/177] chore: move panicWith out of UnionFind (#983) --- Batteries.lean | 1 + Batteries/Data/UnionFind/Basic.lean | 10 +++++----- Batteries/Util/Panic.lean | 12 ++++++++++++ 3 files changed, 18 insertions(+), 5 deletions(-) create mode 100644 Batteries/Util/Panic.lean diff --git a/Batteries.lean b/Batteries.lean index f04a7a1f76..26e0d102d1 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -97,6 +97,7 @@ import Batteries.Test.Internal.DummyLabelAttr import Batteries.Util.Cache import Batteries.Util.ExtendedBinder import Batteries.Util.LibraryNote +import Batteries.Util.Panic import Batteries.Util.Pickle import Batteries.Util.ProofWanted import Batteries.WF diff --git a/Batteries/Data/UnionFind/Basic.lean b/Batteries/Data/UnionFind/Basic.lean index 4fff679993..ce9c99a77e 100644 --- a/Batteries/Data/UnionFind/Basic.lean +++ b/Batteries/Data/UnionFind/Basic.lean @@ -3,10 +3,15 @@ Copyright (c) 2021 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Batteries.Tactic.Alias import Batteries.Tactic.Lint.Misc import Batteries.Tactic.SeqFocus +import Batteries.Util.Panic import Batteries.Data.Array.Lemmas +@[deprecated (since := "2024-10-05")] +protected alias Batteries.UnionFind.panicWith := Batteries.panicWith + namespace Batteries /-- Union-find node type -/ @@ -18,11 +23,6 @@ structure UFNode where namespace UnionFind -/-- Panic with return value -/ -def panicWith (v : α) (msg : String) : α := @panic α ⟨v⟩ msg - -@[simp] theorem panicWith_eq (v : α) (msg) : panicWith v msg = v := rfl - /-- Parent of a union-find node, defaults to self when the node is a root -/ def parentD (arr : Array UFNode) (i : Nat) : Nat := if h : i < arr.size then (arr.get ⟨i, h⟩).parent else i diff --git a/Batteries/Util/Panic.lean b/Batteries/Util/Panic.lean new file mode 100644 index 0000000000..e6317bee71 --- /dev/null +++ b/Batteries/Util/Panic.lean @@ -0,0 +1,12 @@ +/- +Copyright (c) 2024 François G. Dorais. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: François G. Dorais +-/ + +namespace Batteries + +/-- Panic with a specific default value `v`. -/ +def panicWith (v : α) (msg : String) : α := @panic α ⟨v⟩ msg + +@[simp] theorem panicWith_eq (v : α) (msg) : panicWith v msg = v := rfl From 0ccda640fd39bf95aeb0bcf91e206b14dd6a397f Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 16 Oct 2024 00:12:38 +0100 Subject: [PATCH 115/177] feat: characterize `OfScientific.ofScientific` on `Rat` (#990) --- Batteries/Data/Rat/Basic.lean | 3 ++- Batteries/Data/Rat/Lemmas.lean | 6 ++++++ 2 files changed, 8 insertions(+), 1 deletion(-) diff --git a/Batteries/Data/Rat/Basic.lean b/Batteries/Data/Rat/Basic.lean index 2032b23ecd..01cbd07859 100644 --- a/Batteries/Data/Rat/Basic.lean +++ b/Batteries/Data/Rat/Basic.lean @@ -112,7 +112,8 @@ def divInt : Int → Int → Rat else (m * 10 ^ e : Nat) -instance : OfScientific Rat where ofScientific := Rat.ofScientific +instance : OfScientific Rat where + ofScientific m s e := Rat.ofScientific (OfNat.ofNat m) s (OfNat.ofNat e) /-- Rational number strictly less than relation, as a `Bool`. -/ protected def blt (a b : Rat) : Bool := diff --git a/Batteries/Data/Rat/Lemmas.lean b/Batteries/Data/Rat/Lemmas.lean index b3c2d49915..844dc7b552 100644 --- a/Batteries/Data/Rat/Lemmas.lean +++ b/Batteries/Data/Rat/Lemmas.lean @@ -332,6 +332,12 @@ theorem ofScientific_def : Rat.ofScientific m s e = if s then mkRat m (10 ^ e) else (m * 10 ^ e : Nat) := by cases s; exact ofScientific_false_def; exact ofScientific_true_def +/-- `Rat.ofScientific` applied to numeric literals is the same as a scientific literal. -/ +@[simp] +theorem ofScientific_ofNat_ofNat : + Rat.ofScientific (no_index (OfNat.ofNat m)) s (no_index (OfNat.ofNat e)) + = OfScientific.ofScientific m s e := rfl + @[simp] theorem intCast_den (a : Int) : (a : Rat).den = 1 := rfl @[simp] theorem intCast_num (a : Int) : (a : Rat).num = a := rfl From b731e84cb99d738b8d9710a0ba02bf8ec8d7fd26 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 16 Oct 2024 01:56:55 +0200 Subject: [PATCH 116/177] feat: upstream DList results from Mathlib (#989) Co-authored-by: Kim Morrison --- Batteries/Data/DList.lean | 72 +--------------------------- Batteries/Data/DList/Basic.lean | 80 ++++++++++++++++++++++++++++++++ Batteries/Data/DList/Lemmas.lean | 63 +++++++++++++++++++++++++ Batteries/StdDeprecations.lean | 2 +- 4 files changed, 146 insertions(+), 71 deletions(-) create mode 100644 Batteries/Data/DList/Basic.lean create mode 100644 Batteries/Data/DList/Lemmas.lean diff --git a/Batteries/Data/DList.lean b/Batteries/Data/DList.lean index 6a5450f0ba..00a2d9fadf 100644 --- a/Batteries/Data/DList.lean +++ b/Batteries/Data/DList.lean @@ -1,70 +1,2 @@ -/- -Copyright (c) 2018 Microsoft Corporation. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Leonardo de Moura --/ -namespace Batteries -/-- -A difference List is a Function that, given a List, returns the original -contents of the difference List prepended to the given List. -This structure supports `O(1)` `append` and `concat` operations on lists, making it -useful for append-heavy uses such as logging and pretty printing. --/ -structure DList (α : Type u) where - /-- "Run" a `DList` by appending it on the right by a `List α` to get another `List α`. -/ - apply : List α → List α - /-- The `apply` function of a `DList` is completely determined by the list `apply []`. -/ - invariant : ∀ l, apply l = apply [] ++ l - -namespace DList -variable {α : Type u} -open List - -/-- `O(1)` (`apply` is `O(|l|)`). Convert a `List α` into a `DList α`. -/ -def ofList (l : List α) : DList α := - ⟨(l ++ ·), fun t => by simp⟩ - -/-- `O(1)` (`apply` is `O(1)`). Return an empty `DList α`. -/ -def empty : DList α := - ⟨id, fun _ => rfl⟩ - -instance : EmptyCollection (DList α) := ⟨DList.empty⟩ - -/-- `O(apply())`. Convert a `DList α` into a `List α` by running the `apply` function. -/ -def toList : DList α → List α - | ⟨f, _⟩ => f [] - -/-- `O(1)` (`apply` is `O(1)`). A `DList α` corresponding to the list `[a]`. -/ -def singleton (a : α) : DList α where - apply := fun t => a :: t - invariant := fun _ => rfl - -/-- `O(1)` (`apply` is `O(1)`). Prepend `a` on a `DList α`. -/ -def cons : α → DList α → DList α - | a, ⟨f, h⟩ => { - apply := fun t => a :: f t - invariant := by intro t; simp; rw [h] - } - -/-- `O(1)` (`apply` is `O(1)`). Append two `DList α`. -/ -def append : DList α → DList α → DList α - | ⟨f, h₁⟩, ⟨g, h₂⟩ => { - apply := f ∘ g - invariant := by - intro t - show f (g t) = (f (g [])) ++ t - rw [h₁ (g t), h₂ t, ← append_assoc (f []) (g []) t, ← h₁ (g [])] - } - -/-- `O(1)` (`apply` is `O(1)`). Append an element at the end of a `DList α`. -/ -def push : DList α → α → DList α - | ⟨f, h⟩, a => { - apply := fun t => f (a :: t) - invariant := by - intro t - show f (a :: t) = f (a :: nil) ++ t - rw [h [a], h (a::t), append_assoc (f []) [a] t] - rfl - } - -instance : Append (DList α) := ⟨DList.append⟩ +import Batteries.Data.DList.Basic +import Batteries.Data.DList.Lemmas diff --git a/Batteries/Data/DList/Basic.lean b/Batteries/Data/DList/Basic.lean new file mode 100644 index 0000000000..51228eb4b4 --- /dev/null +++ b/Batteries/Data/DList/Basic.lean @@ -0,0 +1,80 @@ +/- +Copyright (c) 2018 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura +-/ +namespace Batteries +/-- +A difference List is a Function that, given a List, returns the original +contents of the difference List prepended to the given List. +This structure supports `O(1)` `append` and `push` operations on lists, making it +useful for append-heavy uses such as logging and pretty printing. +-/ +structure DList (α : Type u) where + /-- "Run" a `DList` by appending it on the right by a `List α` to get another `List α`. -/ + apply : List α → List α + /-- The `apply` function of a `DList` is completely determined by the list `apply []`. -/ + invariant : ∀ l, apply l = apply [] ++ l + +namespace DList +variable {α : Type u} +open List + +/-- `O(1)` (`apply` is `O(|l|)`). Convert a `List α` into a `DList α`. -/ +def ofList (l : List α) : DList α := + ⟨(l ++ ·), fun t => by simp⟩ + +/-- `O(1)` (`apply` is `O(1)`). Return an empty `DList α`. -/ +def empty : DList α := + ⟨id, fun _ => rfl⟩ + +instance : EmptyCollection (DList α) := ⟨DList.empty⟩ + +/-- `O(apply())`. Convert a `DList α` into a `List α` by running the `apply` function. -/ +def toList : DList α → List α + | ⟨f, _⟩ => f [] + +/-- `O(1)` (`apply` is `O(1)`). A `DList α` corresponding to the list `[a]`. -/ +def singleton (a : α) : DList α where + apply := fun t => a :: t + invariant := fun _ => rfl + +/-- `O(1)` (`apply` is `O(1)`). Prepend `a` on a `DList α`. -/ +def cons : α → DList α → DList α + | a, ⟨f, h⟩ => { + apply := fun t => a :: f t + invariant := by intro t; simp; rw [h] + } + +/-- `O(1)` (`apply` is `O(1)`). Append two `DList α`. -/ +def append : DList α → DList α → DList α + | ⟨f, h₁⟩, ⟨g, h₂⟩ => { + apply := f ∘ g + invariant := by + intro t + show f (g t) = (f (g [])) ++ t + rw [h₁ (g t), h₂ t, ← append_assoc (f []) (g []) t, ← h₁ (g [])] + } + +/-- `O(1)` (`apply` is `O(1)`). Append an element at the end of a `DList α`. -/ +def push : DList α → α → DList α + | ⟨f, h⟩, a => { + apply := fun t => f (a :: t) + invariant := by + intro t + show f (a :: t) = f (a :: nil) ++ t + rw [h [a], h (a::t), append_assoc (f []) [a] t] + rfl + } + +instance : Append (DList α) := ⟨DList.append⟩ + +/-- Convert a lazily-evaluated `List` to a `DList` -/ +def lazy_ofList (l : Thunk (List α)) : DList α := + ⟨fun xs => l.get ++ xs, fun t => by simp⟩ + +/-- Concatenates a list of difference lists to form a single difference list. Similar to +`List.join`. -/ +def join {α : Type _} : List (DList α) → DList α + | [] => DList.empty + | x :: xs => x ++ DList.join xs diff --git a/Batteries/Data/DList/Lemmas.lean b/Batteries/Data/DList/Lemmas.lean new file mode 100644 index 0000000000..61f8885864 --- /dev/null +++ b/Batteries/Data/DList/Lemmas.lean @@ -0,0 +1,63 @@ +/- +Copyright (c) 2017 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura +-/ +import Batteries.Data.DList.Basic + +/-! +# Difference list + +This file provides a few results about `DList`. + +A difference list is a function that, given a list, returns the original content of the +difference list prepended to the given list. It is useful to represent elements of a given type +as `a₁ + ... + aₙ` where `+ : α → α → α` is any operation, without actually computing. + +This structure supports `O(1)` `append` and `push` operations on lists, making it +useful for append-heavy uses such as logging and pretty printing. +-/ + +universe u + +namespace Batteries.DList + +open Function + +variable {α : Type u} + +attribute [local simp] Function.comp + +attribute [local simp] ofList toList empty singleton cons push append + +theorem toList_ofList (l : List α) : DList.toList (DList.ofList l) = l := by + cases l; rfl; simp only [DList.toList, DList.ofList, List.cons_append, List.append_nil] + +theorem ofList_toList (l : DList α) : DList.ofList (DList.toList l) = l := by + obtain ⟨app, inv⟩ := l + simp only [ofList, toList, mk.injEq] + funext x + rw [(inv x)] + +theorem toList_empty : toList (@empty α) = [] := by simp + +theorem toList_singleton (x : α) : toList (singleton x) = [x] := by simp + +theorem toList_append (l₁ l₂ : DList α) : toList (l₁ ++ l₂) = toList l₁ ++ toList l₂ := by + obtain ⟨_, l₁_invariant⟩ := l₁; cases l₂; simp; rw [l₁_invariant] + +theorem toList_cons (x : α) (l : DList α) : toList (cons x l) = x :: toList l := by + cases l; simp + +theorem toList_push (x : α) (l : DList α) : toList (push l x) = toList l ++ [x] := by + obtain ⟨_, l_invariant⟩ := l; simp; rw [l_invariant] + +@[simp] +theorem DList_singleton {α : Type _} {a : α} : singleton a = lazy_ofList [a] := + rfl + +@[simp] +theorem DList_lazy {α : Type _} {l : List α} : lazy_ofList l = ofList l := + rfl + +end Batteries.DList diff --git a/Batteries/StdDeprecations.lean b/Batteries/StdDeprecations.lean index cb49a3bdbf..74a00f4f5a 100644 --- a/Batteries/StdDeprecations.lean +++ b/Batteries/StdDeprecations.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Batteries.Tactic.Alias -import Batteries.Data.DList +import Batteries.Data.DList.Basic import Batteries.Data.PairingHeap import Batteries.Data.BinomialHeap.Basic import Batteries.Data.HashMap.Basic From c0b3791156c4a2eb0b311e650f9eb86a4d31e366 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Tue, 15 Oct 2024 20:39:32 -0400 Subject: [PATCH 117/177] feat: add `Array.eraseIdx!` and some lemmas (#988) --- Batteries/Data/Array/Basic.lean | 10 ++++++++++ Batteries/Data/Array/Lemmas.lean | 7 +++++++ 2 files changed, 17 insertions(+) diff --git a/Batteries/Data/Array/Basic.lean b/Batteries/Data/Array/Basic.lean index 7fb0cf8cf8..c74b8f4750 100644 --- a/Batteries/Data/Array/Basic.lean +++ b/Batteries/Data/Array/Basic.lean @@ -172,6 +172,16 @@ greater than i. abbrev eraseIdxN (a : Array α) (i : Nat) (h : i < a.size := by get_elem_tactic) : Array α := a.feraseIdx ⟨i, h⟩ +/-- +Remove the element at a given index from an array, panics if index is out of bounds. +-/ +def eraseIdx! (a : Array α) (i : Nat) : Array α := + if h : i < a.size then + a.feraseIdx ⟨i, h⟩ + else + have : Inhabited (Array α) := ⟨a⟩ + panic! s!"index {i} out of bounds" + end Array diff --git a/Batteries/Data/Array/Lemmas.lean b/Batteries/Data/Array/Lemmas.lean index e6b1cf385a..124f8d412c 100644 --- a/Batteries/Data/Array/Lemmas.lean +++ b/Batteries/Data/Array/Lemmas.lean @@ -147,6 +147,13 @@ where @[simp] proof_wanted toList_erase [BEq α] {l : Array α} {a : α} : (l.erase a).toList = l.toList.erase a +@[simp] theorem eraseIdx!_eq_eraseIdx (a : Array α) (i : Nat) : + a.eraseIdx! i = a.eraseIdx i := rfl + +@[simp] theorem size_eraseIdx (a : Array α) (i : Nat) : + (a.eraseIdx i).size = if i < a.size then a.size-1 else a.size := by + simp only [eraseIdx]; split; simp; rfl + /-! ### shrink -/ theorem size_shrink_loop (a : Array α) (n) : (shrink.loop n a).size = a.size - n := by From 4e80cf3f982a6f93d32917d06dc9f77d02770854 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 16 Oct 2024 11:49:29 +1100 Subject: [PATCH 118/177] chore: cleanup of DList api (#991) Co-authored-by: F. G. Dorais --- Batteries/Data/DList/Basic.lean | 12 ++++++++++-- Batteries/Data/DList/Lemmas.lean | 27 +++++++++++---------------- 2 files changed, 21 insertions(+), 18 deletions(-) diff --git a/Batteries/Data/DList/Basic.lean b/Batteries/Data/DList/Basic.lean index 51228eb4b4..b3eb0bbc97 100644 --- a/Batteries/Data/DList/Basic.lean +++ b/Batteries/Data/DList/Basic.lean @@ -3,6 +3,8 @@ Copyright (c) 2018 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ +import Batteries.Tactic.Alias + namespace Batteries /-- A difference List is a Function that, given a List, returns the original @@ -16,6 +18,8 @@ structure DList (α : Type u) where /-- The `apply` function of a `DList` is completely determined by the list `apply []`. -/ invariant : ∀ l, apply l = apply [] ++ l +attribute [simp] DList.apply + namespace DList variable {α : Type u} open List @@ -30,8 +34,10 @@ def empty : DList α := instance : EmptyCollection (DList α) := ⟨DList.empty⟩ +instance : Inhabited (DList α) := ⟨DList.empty⟩ + /-- `O(apply())`. Convert a `DList α` into a `List α` by running the `apply` function. -/ -def toList : DList α → List α +@[simp] def toList : DList α → List α | ⟨f, _⟩ => f [] /-- `O(1)` (`apply` is `O(1)`). A `DList α` corresponding to the list `[a]`. -/ @@ -70,9 +76,11 @@ def push : DList α → α → DList α instance : Append (DList α) := ⟨DList.append⟩ /-- Convert a lazily-evaluated `List` to a `DList` -/ -def lazy_ofList (l : Thunk (List α)) : DList α := +def ofThunk (l : Thunk (List α)) : DList α := ⟨fun xs => l.get ++ xs, fun t => by simp⟩ +@[deprecated (since := "2024-10-16")] alias lazy_ofList := ofThunk + /-- Concatenates a list of difference lists to form a single difference list. Similar to `List.join`. -/ def join {α : Type _} : List (DList α) → DList α diff --git a/Batteries/Data/DList/Lemmas.lean b/Batteries/Data/DList/Lemmas.lean index 61f8885864..6c17e9e00a 100644 --- a/Batteries/Data/DList/Lemmas.lean +++ b/Batteries/Data/DList/Lemmas.lean @@ -18,20 +18,12 @@ This structure supports `O(1)` `append` and `push` operations on lists, making i useful for append-heavy uses such as logging and pretty printing. -/ -universe u - namespace Batteries.DList open Function -variable {α : Type u} - -attribute [local simp] Function.comp - -attribute [local simp] ofList toList empty singleton cons push append - theorem toList_ofList (l : List α) : DList.toList (DList.ofList l) = l := by - cases l; rfl; simp only [DList.toList, DList.ofList, List.cons_append, List.append_nil] + cases l; rfl; simp [ofList] theorem ofList_toList (l : DList α) : DList.ofList (DList.toList l) = l := by obtain ⟨app, inv⟩ := l @@ -39,25 +31,28 @@ theorem ofList_toList (l : DList α) : DList.ofList (DList.toList l) = l := by funext x rw [(inv x)] -theorem toList_empty : toList (@empty α) = [] := by simp +theorem toList_empty : toList (@empty α) = [] := by simp [empty] -theorem toList_singleton (x : α) : toList (singleton x) = [x] := by simp +theorem toList_singleton (x : α) : toList (singleton x) = [x] := by simp [singleton] theorem toList_append (l₁ l₂ : DList α) : toList (l₁ ++ l₂) = toList l₁ ++ toList l₂ := by - obtain ⟨_, l₁_invariant⟩ := l₁; cases l₂; simp; rw [l₁_invariant] + simp only [toList, append, Function.comp]; rw [invariant] theorem toList_cons (x : α) (l : DList α) : toList (cons x l) = x :: toList l := by - cases l; simp + cases l; simp [cons] theorem toList_push (x : α) (l : DList α) : toList (push l x) = toList l ++ [x] := by - obtain ⟨_, l_invariant⟩ := l; simp; rw [l_invariant] + simp only [toList, push]; rw [invariant] @[simp] -theorem DList_singleton {α : Type _} {a : α} : singleton a = lazy_ofList [a] := +theorem singleton_eq_ofThunk {α : Type _} {a : α} : singleton a = ofThunk [a] := rfl @[simp] -theorem DList_lazy {α : Type _} {l : List α} : lazy_ofList l = ofList l := +theorem ofThunk_coe {α : Type _} {l : List α} : ofThunk l = ofList l := rfl +@[deprecated (since := "2024-10-16")] alias DList_singleton := singleton_eq_ofThunk +@[deprecated (since := "2024-10-16")] alias DList_lazy := ofThunk_coe + end Batteries.DList From 422d1a5f608fccafeddab9748e8038ef346b59bf Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Wed, 16 Oct 2024 00:34:13 -0400 Subject: [PATCH 119/177] feat: add `#help` commands (#969) --- Batteries.lean | 1 + Batteries/Tactic/HelpCmd.lean | 317 +++++++++++++++++++++++++++++++ test/help_cmd.lean | 344 ++++++++++++++++++++++++++++++++++ 3 files changed, 662 insertions(+) create mode 100644 Batteries/Tactic/HelpCmd.lean create mode 100644 test/help_cmd.lean diff --git a/Batteries.lean b/Batteries.lean index 26e0d102d1..7d3486f15b 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -74,6 +74,7 @@ import Batteries.Tactic.Case import Batteries.Tactic.Classical import Batteries.Tactic.Congr import Batteries.Tactic.Exact +import Batteries.Tactic.HelpCmd import Batteries.Tactic.Init import Batteries.Tactic.Instances import Batteries.Tactic.Lemma diff --git a/Batteries/Tactic/HelpCmd.lean b/Batteries/Tactic/HelpCmd.lean new file mode 100644 index 0000000000..cf1c0f2646 --- /dev/null +++ b/Batteries/Tactic/HelpCmd.lean @@ -0,0 +1,317 @@ +/- +Copyright (c) 2022 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ +import Lean.Elab.Syntax +import Lean.DocString + +/-! + +# The `#help` command + +The `#help` command can be used to list all definitions in a variety of extensible aspects of lean. + +* `#help option` lists options (used in `set_option myOption`) +* `#help attr` lists attributes (used in `@[myAttr] def foo := ...`) +* `#help cats` lists syntax categories (like `term`, `tactic`, `stx` etc) +* `#help cat C` lists elements of syntax category C + * `#help term`, `#help tactic`, `#help conv`, `#help command` + are shorthand for `#help cat term` etc. + * `#help cat+ C` also shows `elab` and `macro` definitions associated to the syntaxes + +All forms take an optional identifier to narrow the search; for example `#help option pp` shows +only `pp.*` options. + +-/ + +namespace Batteries.Tactic +open Lean Meta Elab Tactic Command + +/-- +The command `#help option` shows all options that have been defined in the current environment. +Each option has a format like: +``` +option pp.all : Bool := false + (pretty printer) display coercions, implicit parameters, proof terms, fully qualified names, + universe, and disable beta reduction and notations during pretty printing +``` +This says that `pp.all` is an option which can be set to a `Bool` value, and the default value is +`false`. If an option has been modified from the default using e.g. `set_option pp.all true`, +it will appear as a `(currently: true)` note next to the option. + +The form `#help option id` will show only options that begin with `id`. +-/ +syntax withPosition("#help " colGt &"option" (colGt ppSpace Parser.rawIdent)?) : command + +private def elabHelpOption (id : Option Ident) : CommandElabM Unit := do + let id := id.map (·.raw.getId.toString false) + let mut decls : Lean.RBMap _ _ compare := {} + for (name, decl) in show Lean.RBMap .. from ← getOptionDecls do + let name := name.toString false + if let some id := id then + if !id.isPrefixOf name then + continue + decls := decls.insert name decl + let mut msg := Format.nil + let opts ← getOptions + if decls.isEmpty then + match id with + | some id => throwError "no options start with {id}" + | none => throwError "no options found (!)" + for (name, decl) in decls do + let mut msg1 := match decl.defValue with + | .ofString val => s!"String := {repr val}" + | .ofBool val => s!"Bool := {repr val}" + | .ofName val => s!"Name := {repr val}" + | .ofNat val => s!"Nat := {repr val}" + | .ofInt val => s!"Int := {repr val}" + | .ofSyntax val => s!"Syntax := {repr val}" + if let some val := opts.find (.mkSimple name) then + msg1 := s!"{msg1} (currently: {val})" + msg := msg ++ .nest 2 (f!"option {name} : {msg1}" ++ .line ++ decl.descr) ++ .line ++ .line + logInfo msg + +elab_rules : command | `(#help option $(id)?) => elabHelpOption id + +/-- +The command `#help attribute` (or the short form `#help attr`) shows all attributes that have been +defined in the current environment. +Each attribute has a format like: +``` +[inline]: mark definition to always be inlined +``` +This says that `inline` is an attribute that can be placed on definitions like +`@[inline] def foo := 1`. (Individual attributes may have restrictions on where they can be +applied; see the attribute's documentation for details.) Both the attribute's `descr` field as well +as the docstring will be displayed here. + +The form `#help attr id` will show only attributes that begin with `id`. +-/ +syntax withPosition("#help " colGt (&"attr" <|> &"attribute") + (colGt ppSpace Parser.rawIdent)?) : command + +private def elabHelpAttr (id : Option Ident) : CommandElabM Unit := do + let id := id.map (·.raw.getId.toString false) + let mut decls : Lean.RBMap _ _ compare := {} + /- + #adaptation_note + On nightly-2024-06-21, added the `.toList` here: + without it the requisite `ForIn` instance can't be found. + -/ + for (name, decl) in (← attributeMapRef.get).toList do + let name := name.toString false + if let some id := id then + if !id.isPrefixOf name then + continue + decls := decls.insert name decl + let mut msg := Format.nil + let env ← getEnv + if decls.isEmpty then + match id with + | some id => throwError "no attributes start with {id}" + | none => throwError "no attributes found (!)" + for (name, decl) in decls do + let mut msg1 := s!"[{name}]: {decl.descr}" + if let some doc ← findDocString? env decl.ref then + msg1 := s!"{msg1}\n{doc.trim}" + msg := msg ++ .nest 2 msg1 ++ .line ++ .line + logInfo msg + +elab_rules : command + | `(#help attr $(id)?) => elabHelpAttr id + | `(#help attribute $(id)?) => elabHelpAttr id + +/-- Gets the initial string token in a parser description. For example, for a declaration like +`syntax "bla" "baz" term : tactic`, it returns `some "bla"`. Returns `none` for syntax declarations +that don't start with a string constant. -/ +partial def getHeadTk (e : Expr) : Option String := + match e.getAppFnArgs with + | (``ParserDescr.node, #[_, _, p]) + | (``ParserDescr.trailingNode, #[_, _, _, p]) + | (``ParserDescr.unary, #[.app _ (.lit (.strVal "withPosition")), p]) + | (``ParserDescr.unary, #[.app _ (.lit (.strVal "atomic")), p]) + | (``ParserDescr.unary, #[.app _ (.lit (.strVal "ppRealGroup")), p]) + | (``ParserDescr.unary, #[.app _ (.lit (.strVal "ppRealFill")), p]) + | (``Parser.ppRealFill, #[p]) + | (``Parser.withAntiquot, #[_, p]) + | (``Parser.leadingNode, #[_, _, p]) + | (``Parser.trailingNode, #[_, _, _, p]) + | (``Parser.group, #[p]) + | (``Parser.withCache, #[_, p]) + | (``Parser.withResetCache, #[p]) + | (``Parser.withPosition, #[p]) + | (``Parser.withOpen, #[p]) + | (``Parser.withPositionAfterLinebreak, #[p]) + | (``Parser.suppressInsideQuot, #[p]) + | (``Parser.ppRealGroup, #[p]) + | (``Parser.ppIndent, #[p]) + | (``Parser.ppDedent, #[p]) + => getHeadTk p + | (``ParserDescr.binary, #[.app _ (.lit (.strVal "andthen")), p, q]) + | (``HAndThen.hAndThen, #[_, _, _, _, p, .lam _ _ q _]) + => getHeadTk p <|> getHeadTk q + | (``ParserDescr.nonReservedSymbol, #[.lit (.strVal tk), _]) + | (``ParserDescr.symbol, #[.lit (.strVal tk)]) + | (``Parser.nonReservedSymbol, #[.lit (.strVal tk), _]) + | (``Parser.symbol, #[.lit (.strVal tk)]) + | (``Parser.unicodeSymbol, #[.lit (.strVal tk), _]) + => pure tk + | _ => none + +/-- +The command `#help cats` shows all syntax categories that have been defined in the +current environment. +Each syntax has a format like: +``` +category command [Lean.Parser.initFn✝] +``` +The name of the syntax category in this case is `command`, and `Lean.Parser.initFn✝` is the +name of the declaration that introduced it. (It is often an anonymous declaration like this, +but you can click to go to the definition.) It also shows the doc string if available. + +The form `#help cats id` will show only syntax categories that begin with `id`. +-/ +syntax withPosition("#help " colGt &"cats" (colGt ppSpace Parser.rawIdent)?) : command + +private def elabHelpCats (id : Option Ident) : CommandElabM Unit := do + let id := id.map (·.raw.getId.toString false) + let mut decls : Lean.RBMap _ _ compare := {} + for (name, cat) in (Parser.parserExtension.getState (← getEnv)).categories do + let name := name.toString false + if let some id := id then + if !id.isPrefixOf name then + continue + decls := decls.insert name cat + let mut msg := MessageData.nil + let env ← getEnv + if decls.isEmpty then + match id with + | some id => throwError "no syntax categories start with {id}" + | none => throwError "no syntax categories found (!)" + for (name, cat) in decls do + let mut msg1 := m!"category {name} [{mkConst cat.declName}]" + if let some doc ← findDocString? env cat.declName then + msg1 := msg1 ++ Format.line ++ doc.trim + msg := msg ++ .nest 2 msg1 ++ (.line ++ .line : Format) + logInfo msg + +elab_rules : command | `(#help cats $(id)?) => elabHelpCats id + +/-- +The command `#help cat C` shows all syntaxes that have been defined in syntax category `C` in the +current environment. +Each syntax has a format like: +``` +syntax "first"... [Parser.tactic.first] + `first | tac | ...` runs each `tac` until one succeeds, or else fails. +``` +The quoted string is the leading token of the syntax, if applicable. It is followed by the full +name of the syntax (which you can also click to go to the definition), and the documentation. + +* The form `#help cat C id` will show only attributes that begin with `id`. +* The form `#help cat+ C` will also show information about any `macro`s and `elab`s + associated to the listed syntaxes. +-/ +syntax withPosition("#help " colGt &"cat" "+"? colGt ident + (colGt ppSpace (Parser.rawIdent <|> str))?) : command + +private def elabHelpCat (more : Option Syntax) (catStx : Ident) (id : Option String) : + CommandElabM Unit := do + let mut decls : Lean.RBMap _ _ compare := {} + let mut rest : Lean.RBMap _ _ compare := {} + let catName := catStx.getId.eraseMacroScopes + let some cat := (Parser.parserExtension.getState (← getEnv)).categories.find? catName + | throwErrorAt catStx "{catStx} is not a syntax category" + liftTermElabM <| Term.addCategoryInfo catStx catName + let env ← getEnv + for (k, _) in cat.kinds do + let mut used := false + if let some tk := do getHeadTk (← (← env.find? k).value?) then + let tk := tk.trim + if let some id := id then + if !id.isPrefixOf tk then + continue + used := true + decls := decls.insert tk ((decls.findD tk #[]).push k) + if !used && id.isNone then + rest := rest.insert (k.toString false) k + let mut msg := MessageData.nil + if decls.isEmpty && rest.isEmpty then + match id with + | some id => throwError "no {catName} declarations start with {id}" + | none => throwError "no {catName} declarations found" + let env ← getEnv + let addMsg (k : SyntaxNodeKind) (msg msg1 : MessageData) : CommandElabM MessageData := do + let mut msg1 := msg1 + if let some doc ← findDocString? env k then + msg1 := msg1 ++ Format.line ++ doc.trim + msg1 := .nest 2 msg1 + if more.isSome then + let addElabs {α} (type : String) (attr : KeyedDeclsAttribute α) + (msg : MessageData) : CommandElabM MessageData := do + let mut msg := msg + for e in attr.getEntries env k do + let x := e.declName + msg := msg ++ Format.line ++ m!"+ {type} {mkConst x}" + if let some doc ← findDocString? env x then + msg := msg ++ .nest 2 (Format.line ++ doc.trim) + pure msg + msg1 ← addElabs "macro" macroAttribute msg1 + match catName with + | `term => msg1 ← addElabs "term elab" Term.termElabAttribute msg1 + | `command => msg1 ← addElabs "command elab" commandElabAttribute msg1 + | `tactic | `conv => msg1 ← addElabs "tactic elab" tacticElabAttribute msg1 + | _ => pure () + return msg ++ msg1 ++ (.line ++ .line : Format) + for (name, ks) in decls do + for k in ks do + msg ← addMsg k msg m!"syntax {repr name}... [{mkConst k}]" + for (_, k) in rest do + msg ← addMsg k msg m!"syntax ... [{mkConst k}]" + logInfo msg + +elab_rules : command + | `(#help cat $[+%$more]? $cat) => elabHelpCat more cat none + | `(#help cat $[+%$more]? $cat $id:ident) => elabHelpCat more cat (id.getId.toString false) + | `(#help cat $[+%$more]? $cat $id:str) => elabHelpCat more cat id.getString + +/-- +The command `#help term` shows all term syntaxes that have been defined in the current environment. +See `#help cat` for more information. +-/ +syntax withPosition("#help " colGt &"term" "+"? + (colGt ppSpace (Parser.rawIdent <|> str))?) : command +macro_rules + | `(#help term%$tk $[+%$more]? $(id)?) => + `(#help cat$[+%$more]? $(mkIdentFrom tk `term) $(id)?) + +/-- +The command `#help tactic` shows all tactics that have been defined in the current environment. +See `#help cat` for more information. +-/ +syntax withPosition("#help " colGt &"tactic" "+"? + (colGt ppSpace (Parser.rawIdent <|> str))?) : command +macro_rules + | `(#help tactic%$tk $[+%$more]? $(id)?) => + `(#help cat$[+%$more]? $(mkIdentFrom tk `tactic) $(id)?) + +/-- +The command `#help conv` shows all tactics that have been defined in the current environment. +See `#help cat` for more information. +-/ +syntax withPosition("#help " colGt &"conv" "+"? + (colGt ppSpace (Parser.rawIdent <|> str))?) : command +macro_rules + | `(#help conv%$tk $[+%$more]? $(id)?) => + `(#help cat$[+%$more]? $(mkIdentFrom tk `conv) $(id)?) + +/-- +The command `#help command` shows all commands that have been defined in the current environment. +See `#help cat` for more information. +-/ +syntax withPosition("#help " colGt &"command" "+"? + (colGt ppSpace (Parser.rawIdent <|> str))?) : command +macro_rules + | `(#help command%$tk $[+%$more]? $(id)?) => + `(#help cat$[+%$more]? $(mkIdentFrom tk `command) $(id)?) diff --git a/test/help_cmd.lean b/test/help_cmd.lean new file mode 100644 index 0000000000..57e214cafd --- /dev/null +++ b/test/help_cmd.lean @@ -0,0 +1,344 @@ +import Batteries.Tactic.HelpCmd + +/-! The `#help` command + +The `#help` command family currently contains these subcommands: + +* `#help attr` / `#help attribute` +* `#help cat` +* `#help cats` +* `#help command` (abbrev for `#help cat command`) +* `#help conv` (abbrev for `#help cat conv`) +* `#help option` +* `#help tactic` (abbrev for `#help cat tactic`) +* `#help term` (abbrev for `#help cat term`) + +All forms take an optional identifier prefix to narrow the search. The `#help cat` command has a +variant `#help cat+` that displays additional information, similarly for commands derived from +`#help cat`. + +WARNING: Some of these tests will need occasional updates when new features are added and even when +some documentation is edited. This type of break will be unexpected but the fix will not be +unexpected! Just update the guard text to match the output after your addition. +-/ + +/-! `#help attr` -/ + +-- this is a long and constantly updated listing, we don't check the output +#guard_msgs(error, drop info) in +#help attr + +/-- +error: no attributes start with foobarbaz +-/ +#guard_msgs in +#help attr foobarbaz + +/-- +info: +[inline]: mark definition to be inlined + +[inline_if_reduce]: mark definition to be inlined when resultant term after reduction is not a +`cases_on` application +-/ +#guard_msgs in +#help attr inl + +/-! `#help cat` -/ + +-- this is a long and constantly updated listing, we don't check the output +#guard_msgs(error, drop info) in +#help cat term + +/-- +error: foobarbaz is not a syntax category +-/ +#guard_msgs in +#help cat foobarbaz + +/-- +info: +syntax "("... [«prec(_)»] + Parentheses are used for grouping precedence expressions. + +syntax "+"... [Lean.Parser.Syntax.addPrec] + Addition of precedences. This is normally used only for offsetting, e.g. `max + 1`. + +syntax "-"... [Lean.Parser.Syntax.subPrec] + Subtraction of precedences. This is normally used only for offsetting, e.g. `max - 1`. + +syntax "arg"... [precArg] + Precedence used for application arguments (`do`, `by`, ...). + +syntax "lead"... [precLead] + Precedence used for terms not supposed to be used as arguments (`let`, `have`, ...). + +syntax "max"... [precMax] + Maximum precedence used in term parsers, in particular for terms in + function position (`ident`, `paren`, ...) + +syntax "min"... [precMin] + Minimum precedence used in term parsers. + +syntax "min1"... [precMin1] + `(min+1)` (we can only write `min+1` after `Meta.lean`) + +syntax ... [Lean.Parser.Syntax.numPrec] +-/ +#guard_msgs in +#help cat prec + +/-- +info: +syntax "("... [«prec(_)»] + Parentheses are used for grouping precedence expressions. ++ macro «_aux_Init_Notation___macroRules_prec(_)_1» + Parentheses are used for grouping precedence expressions. + +syntax "+"... [Lean.Parser.Syntax.addPrec] + Addition of precedences. This is normally used only for offsetting, e.g. `max + 1`. ++ macro Lean._aux_Init_Meta___macroRules_Lean_Parser_Syntax_addPrec_1 + +syntax "-"... [Lean.Parser.Syntax.subPrec] + Subtraction of precedences. This is normally used only for offsetting, e.g. `max - 1`. ++ macro Lean._aux_Init_Meta___macroRules_Lean_Parser_Syntax_subPrec_1 + +syntax "arg"... [precArg] + Precedence used for application arguments (`do`, `by`, ...). ++ macro _aux_Init_Notation___macroRules_precArg_1 + Precedence used for application arguments (`do`, `by`, ...). + +syntax "lead"... [precLead] + Precedence used for terms not supposed to be used as arguments (`let`, `have`, ...). ++ macro _aux_Init_Notation___macroRules_precLead_1 + Precedence used for terms not supposed to be used as arguments (`let`, `have`, ...). + +syntax "max"... [precMax] + Maximum precedence used in term parsers, in particular for terms in + function position (`ident`, `paren`, ...) ++ macro _aux_Init_Notation___macroRules_precMax_1 + Maximum precedence used in term parsers, in particular for terms in + function position (`ident`, `paren`, ...) + +syntax "min"... [precMin] + Minimum precedence used in term parsers. ++ macro _aux_Init_Notation___macroRules_precMin_1 + Minimum precedence used in term parsers. + +syntax "min1"... [precMin1] + `(min+1)` (we can only write `min+1` after `Meta.lean`) ++ macro _aux_Init_Notation___macroRules_precMin1_1 + `(min+1)` (we can only write `min+1` after `Meta.lean`) + +syntax ... [Lean.Parser.Syntax.numPrec] +-/ +#guard_msgs in +#help cat+ prec + +/-! `#help cats` -/ + +-- this is a long and constantly updated listing, we don't check the output +#guard_msgs(error, drop info) in +#help cats + +/-- +error: no syntax categories start with foobarbaz +-/ +#guard_msgs in +#help cats foobarbaz + +/-- +info: +category prec [Lean.Parser.Category.prec] + `prec` is a builtin syntax category for precedences. A precedence is a value + that expresses how tightly a piece of syntax binds: for example `1 + 2 * 3` is + parsed as `1 + (2 * 3)` because `*` has a higher pr0ecedence than `+`. + Higher numbers denote higher precedence. + In addition to literals like `37`, there are some special named priorities: + * `arg` for the precedence of function arguments + * `max` for the highest precedence used in term parsers (not actually the maximum possible value) + * `lead` for the precedence of terms not supposed to be used as arguments + and you can also add and subtract precedences. + +category prio [Lean.Parser.Category.prio] + `prio` is a builtin syntax category for priorities. + Priorities are used in many different attributes. + Higher numbers denote higher priority, and for example typeclass search will + try high priority instances before low priority. + In addition to literals like `37`, you can also use `low`, `mid`, `high`, as well as + add and subtract priorities. +-/ +#guard_msgs in +#help cats pr + +/-! `#help command` -/ + +-- this is a long and constantly updated listing, we don't check the output +#guard_msgs(error, drop info) in +#help command + +/-- +error: no command declarations start with foobarbaz +-/ +#guard_msgs in +#help command foobarbaz + +/-- +info: +syntax "#eval"... [Lean.Parser.Command.eval] + +syntax "#eval!"... [Lean.Parser.Command.evalBang] + +syntax "#exit"... [Lean.Parser.Command.exit] +-/ +#guard_msgs in +#help command "#e" + +/-- +info: +syntax "#eval"... [Lean.Parser.Command.eval] ++ command elab Lean.Elab.Command.elabEval + +syntax "#eval!"... [Lean.Parser.Command.evalBang] ++ command elab Lean.Elab.Command.elabEvalBang + +syntax "#exit"... [Lean.Parser.Command.exit] ++ command elab Lean.Elab.Command.elabExit +-/ +#guard_msgs in +#help command+ "#e" + +/-! #help conv -/ + +-- this is a long and constantly updated listing, we don't check the output +#guard_msgs(error, drop info) in +#help conv + +/-- +error: no conv declarations start with foobarbaz +-/ +#guard_msgs in +#help conv foobarbaz + +/-- +info: +syntax "reduce"... [Lean.Parser.Tactic.Conv.reduce] + Puts term in normal form, this tactic is meant for debugging purposes only. + +syntax "repeat"... [Lean.Parser.Tactic.Conv.convRepeat_] + `repeat convs` runs the sequence `convs` repeatedly until it fails to apply. + +syntax "rewrite"... [Lean.Parser.Tactic.Conv.rewrite] + `rw [thm]` rewrites the target using `thm`. See the `rw` tactic for more information. +-/ +#guard_msgs in +#help conv "re" + +/-- +info: +syntax "reduce"... [Lean.Parser.Tactic.Conv.reduce] + Puts term in normal form, this tactic is meant for debugging purposes only. ++ tactic elab Lean.Elab.Tactic.Conv.evalReduce + +syntax "repeat"... [Lean.Parser.Tactic.Conv.convRepeat_] + `repeat convs` runs the sequence `convs` repeatedly until it fails to apply. ++ macro Lean.Parser.Tactic.Conv._aux_Init_Conv___macroRules_Lean_Parser_Tactic_Conv_convRepeat__1 + +syntax "rewrite"... [Lean.Parser.Tactic.Conv.rewrite] + `rw [thm]` rewrites the target using `thm`. See the `rw` tactic for more information. ++ tactic elab Lean.Elab.Tactic.Conv.evalRewrite +-/ +#guard_msgs in +#help conv+ "re" + +/-! `#help option` -/ + +-- this is a long and constantly updated listing, we don't check the output +#guard_msgs(error, drop info) in +#help option + +/-- +error: no options start with foobarbaz +-/ +#guard_msgs in +#help option foobarbaz + +/-- +info: +option pp.instanceTypes : Bool := false + (pretty printer) when printing explicit applications, show the types of inst-implicit arguments + +option pp.instances : Bool := true + (pretty printer) if set to false, replace inst-implicit arguments to explicit applications with +placeholders + +option pp.instantiateMVars : Bool := true + (pretty printer) instantiate mvars before delaborating +-/ +#guard_msgs in +#help option pp.ins + +/-! `#help tactic` -/ + +-- this is a long and constantly updated listing, we don't check the output +#guard_msgs(error, drop info) in +#help tactic + +/-- +error: no tactic declarations start with foobarbaz +-/ +#guard_msgs in +#help tactic foobarbaz + +/-- +info: +syntax "by_cases"... [«tacticBy_cases_:_»] + `by_cases (h :)? p` splits the main goal into two cases, assuming `h : p` in the first branch, +and `h : ¬ p` in the second branch. +-/ +#guard_msgs in +#help tactic by + +/-- +info: +syntax "by_cases"... [«tacticBy_cases_:_»] + `by_cases (h :)? p` splits the main goal into two cases, assuming `h : p` in the first branch, and `h : ¬ p` in the second branch. ++ macro «_aux_Init_ByCases___macroRules_tacticBy_cases_:__2» ++ macro «_aux_Init_ByCases___macroRules_tacticBy_cases_:__1» +-/ +#guard_msgs in +#help tactic+ by + +/-! #help term -/ + +-- this is a long and constantly updated listing, we don't check the output +#guard_msgs(error, drop info) in +#help term + +/-- +error: no term declarations start with foobarbaz +-/ +#guard_msgs in +#help term foobarbaz + +/-- +info: +syntax "decl_name%"... [Lean.Parser.Term.declName] + A macro which evaluates to the name of the currently elaborating declaration. + +syntax "default_or_ofNonempty%"... [Lean.Parser.Term.defaultOrOfNonempty] +-/ +#guard_msgs in +#help term de + +/-- +info: +syntax "decl_name%"... [Lean.Parser.Term.declName] + A macro which evaluates to the name of the currently elaborating declaration. ++ term elab Lean.Elab.Term.elabDeclName + +syntax "default_or_ofNonempty%"... [Lean.Parser.Term.defaultOrOfNonempty] ++ term elab Lean.Elab.Term.Op.elabDefaultOrNonempty +-/ +#guard_msgs in +#help term+ de From d011c00010429728fab0bb3675b13f4bce63539e Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Wed, 16 Oct 2024 21:29:16 -0700 Subject: [PATCH 120/177] fix: make `proof_wanted` not report unused variables (#997) --- Batteries/Util/ProofWanted.lean | 9 ++++++--- test/proof_wanted.lean | 15 +++++++++++++++ 2 files changed, 21 insertions(+), 3 deletions(-) create mode 100644 test/proof_wanted.lean diff --git a/Batteries/Util/ProofWanted.lean b/Batteries/Util/ProofWanted.lean index 97f1473f70..3a6fd90981 100644 --- a/Batteries/Util/ProofWanted.lean +++ b/Batteries/Util/ProofWanted.lean @@ -33,7 +33,10 @@ elaboration, but it's then removed from the environment. def elabProofWanted : CommandElab | `($mods:declModifiers proof_wanted $name $args* : $res) => withoutModifyingEnv do -- The helper axiom is used instead of `sorry` to avoid spurious warnings - elabCommand <| ← `(axiom helper (p : Prop) : p - $mods:declModifiers - theorem $name $args* : $res := helper _) + elabCommand <| ← `( + section + set_option linter.unusedVariables false + axiom helper {α : Sort _} : α + $mods:declModifiers theorem $name $args* : $res := helper + end) | _ => throwUnsupportedSyntax diff --git a/test/proof_wanted.lean b/test/proof_wanted.lean new file mode 100644 index 0000000000..38f5714eb3 --- /dev/null +++ b/test/proof_wanted.lean @@ -0,0 +1,15 @@ +import Batteries.Util.ProofWanted + +/-! +No unused variable warnings. +-/ +#guard_msgs in proof_wanted foo (x : Nat) : True + +/-! +When not a proposition, rely on `theorem` command failing. +-/ +/-- +error: type of theorem 'foo' is not a proposition + Nat → Nat +-/ +#guard_msgs in proof_wanted foo (x : Nat) : Nat From cc0bc876eeef0518ddc1c8d3bd6f48cc83e68901 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 17 Oct 2024 17:42:39 +1100 Subject: [PATCH 121/177] chore: attempt to restore Mathlib CI (#999) --- .github/workflows/test_mathlib.yml | 52 +++++++++++++----------------- 1 file changed, 22 insertions(+), 30 deletions(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index 8eca882a4c..efce8db636 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -12,15 +12,24 @@ jobs: runs-on: ubuntu-latest if: github.event.workflow_run.conclusion == 'success' && github.event.workflow_run.event == 'pull_request' && github.repository == 'leanprover-community/batteries' steps: - - name: Retrieve information about the original workflow - uses: potiuk/get-workflow-origin@v1_1 - id: workflow-info + - name: Retrieve PR information + id: pr-info + uses: actions/github-script@v6 with: - token: ${{ secrets.GITHUB_TOKEN }} - sourceRunId: ${{ github.event.workflow_run.id }} + script: | + const prNumber = context.payload.workflow_run.pull_requests[0].number; + const { data: pr } = await github.rest.pulls.get({ + owner: context.repo.owner, + repo: context.repo.name, + pull_number: prNumber + }); + core.setOutput('targetBranch', pr.base.ref); + core.setOutput('pullRequestNumber', pr.number); + core.exportVariable('HEAD_REPO', pr.head.repo.full_name); + core.exportVariable('HEAD_BRANCH', pr.head.ref); - name: Checkout mathlib4 repository - if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.workflow-info.outputs.targetBranch == 'main' + if: steps.pr-info.outputs.pullRequestNumber != '' && steps.pr-info.outputs.targetBranch == 'main' uses: actions/checkout@v4 with: repository: leanprover-community/mathlib4 @@ -28,36 +37,19 @@ jobs: ref: master fetch-depth: 0 - - name: install elan - if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.workflow-info.outputs.targetBranch == 'main' + - name: Install elan + if: steps.pr-info.outputs.pullRequestNumber != '' && steps.pr-info.outputs.targetBranch == 'main' run: | set -o pipefail curl -sSfL https://github.com/leanprover/elan/releases/download/v3.0.0/elan-x86_64-unknown-linux-gnu.tar.gz | tar xz ./elan-init -y --default-toolchain none echo "$HOME/.elan/bin" >> "${GITHUB_PATH}" - - name: Retrieve PR information - if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.workflow-info.outputs.targetBranch == 'main' - id: pr-info - uses: actions/github-script@v6 - env: - PR_NUMBER: ${{ steps.workflow-info.outputs.pullRequestNumber }} - with: - script: | - const prNumber = process.env.PR_NUMBER; - const { data: pr } = await github.rest.pulls.get({ - owner: context.repo.owner, - repo: context.repo.repo, - pull_number: prNumber - }); - core.exportVariable('HEAD_REPO', pr.head.repo.full_name); - core.exportVariable('HEAD_BRANCH', pr.head.ref); - - name: Check if tag exists - if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.workflow-info.outputs.targetBranch == 'main' + if: steps.pr-info.outputs.pullRequestNumber != '' && steps.pr-info.outputs.targetBranch == 'main' id: check_mathlib_tag env: - PR_NUMBER: ${{ steps.workflow-info.outputs.pullRequestNumber }} + PR_NUMBER: ${{ steps.pr-info.outputs.pullRequestNumber }} HEAD_REPO: ${{ env.HEAD_REPO }} HEAD_BRANCH: ${{ env.HEAD_BRANCH }} run: | @@ -75,7 +67,7 @@ jobs: echo "Branch does not exist, creating it." git switch -c batteries-pr-testing-$PR_NUMBER "$BASE" - # Use the fork and branch name to modify the lakefile.lean + # Modify the lakefile.lean with the fork and branch name sed -i "s,require \"leanprover-community\" / \"batteries\" @ git \".\+\",require \"leanprover-community\" / \"batteries\" from git \"https://github.com/$HEAD_REPO\" @ \"$HEAD_BRANCH\",g" lakefile.lean lake update batteries @@ -91,8 +83,8 @@ jobs: fi - name: Push changes - if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.workflow-info.outputs.targetBranch == 'main' + if: steps.pr-info.outputs.pullRequestNumber != '' && steps.pr-info.outputs.targetBranch == 'main' env: - PR_NUMBER: ${{ steps.workflow-info.outputs.pullRequestNumber }} + PR_NUMBER: ${{ steps.pr-info.outputs.pullRequestNumber }} run: | git push origin batteries-pr-testing-$PR_NUMBER From c521f0185f4dd42b6aa4898010d5ba5357c57c9f Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Thu, 17 Oct 2024 13:40:51 +0200 Subject: [PATCH 122/177] chore: deprecate alias `Function.funext_iff` in favor of `funext_iff` (#998) --- Batteries/Logic.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Batteries/Logic.lean b/Batteries/Logic.lean index f8fe5a6e38..70091ae973 100644 --- a/Batteries/Logic.lean +++ b/Batteries/Logic.lean @@ -54,6 +54,7 @@ theorem funext₃ {β : α → Sort _} {γ : ∀ a, β a → Sort _} {δ : ∀ a {f g : ∀ a b c, δ a b c} (h : ∀ a b c, f a b c = g a b c) : f = g := funext fun _ => funext₂ <| h _ +@[deprecated (since := "2024-10-17")] protected alias Function.funext_iff := funext_iff theorem ne_of_apply_ne {α β : Sort _} (f : α → β) {x y : α} : f x ≠ f y → x ≠ y := From 9e3d0d810e9021767c360756f066574f72e8fcce Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Fri, 18 Oct 2024 16:43:01 -0400 Subject: [PATCH 123/177] fix: Mathlib CI typo (#1002) --- .github/workflows/test_mathlib.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index efce8db636..a1af0d771d 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -20,7 +20,7 @@ jobs: const prNumber = context.payload.workflow_run.pull_requests[0].number; const { data: pr } = await github.rest.pulls.get({ owner: context.repo.owner, - repo: context.repo.name, + repo: context.repo.repo, pull_number: prNumber }); core.setOutput('targetBranch', pr.base.ref); From 6f569686cca3d733804fd07f5adf4b90ac57bb90 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 19 Oct 2024 10:54:47 +1100 Subject: [PATCH 124/177] chore: deprecate `Expr.lambdaArity` in favor of existing upstream version (#992) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: François G. Dorais --- Batteries/Lean/Expr.lean | 13 ++++++------- 1 file changed, 6 insertions(+), 7 deletions(-) diff --git a/Batteries/Lean/Expr.lean b/Batteries/Lean/Expr.lean index 20338d3424..d1c568f3e4 100644 --- a/Batteries/Lean/Expr.lean +++ b/Batteries/Lean/Expr.lean @@ -24,6 +24,9 @@ def toSyntax (e : Expr) : TermElabM Syntax.Term := withFreshMacroScope do mvar.mvarId!.assign e pure stx +@[deprecated (since := "2024-10-16"), inherit_doc getNumHeadLambdas] +abbrev lambdaArity := @getNumHeadLambdas + /-- Returns the number of leading `∀` binders of an expression. Ignores metadata. -/ @@ -32,13 +35,9 @@ def forallArity : Expr → Nat | forallE _ _ body _ => 1 + forallArity body | _ => 0 -/-- -Returns the number of leading `λ` binders of an expression. Ignores metadata. --/ -def lambdaArity : Expr → Nat - | mdata _ b => lambdaArity b - | lam _ _ b _ => 1 + lambdaArity b - | _ => 0 +-- TODO: replace `forallArity` after https://github.com/leanprover/lean4/pull/5729 +-- @[deprecated (since := "2024-10-16"), inherit_doc getNumHeadForalls] +-- abbrev forallArity := @getNumHeadForalls /-- Like `getAppNumArgs` but ignores metadata. -/ def getAppNumArgs' (e : Expr) : Nat := From 1e0bf50b357069e1d658512a579a5faac6587c40 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Fri, 18 Oct 2024 20:00:36 -0400 Subject: [PATCH 125/177] fix: print prefix bug (#996) Co-authored-by: Mario Carneiro --- Batteries/Tactic/PrintPrefix.lean | 18 +++++++++--------- test/print_prefix.lean | 4 ++++ 2 files changed, 13 insertions(+), 9 deletions(-) diff --git a/Batteries/Tactic/PrintPrefix.lean b/Batteries/Tactic/PrintPrefix.lean index b83e7b0ad8..dbedad072f 100644 --- a/Batteries/Tactic/PrintPrefix.lean +++ b/Batteries/Tactic/PrintPrefix.lean @@ -108,12 +108,12 @@ by setting `showTypes` to `false`: The complete set of flags can be seen in the documentation for `Lean.Elab.Command.PrintPrefixConfig`. -/ -elab (name := printPrefix) "#print" tk:"prefix" - cfg:(Lean.Parser.Tactic.config)? name:ident : command => liftTermElabM do - let nameId := name.getId - let opts ← elabPrintPrefixConfig (mkOptionalNode cfg) - let mut msgs ← matchingConstants opts nameId - if msgs.isEmpty then - if let [name] ← resolveGlobalConst name then - msgs ← matchingConstants opts name - logInfoAt tk (.joinSep msgs.toList "") +elab (name := printPrefix) tk:"#print " colGt "prefix" + cfg:(Lean.Parser.Tactic.config)? name:(ident)? : command => liftTermElabM do + if let some name := name then + let opts ← elabPrintPrefixConfig (mkOptionalNode cfg) + let mut msgs ← matchingConstants opts name.getId + if msgs.isEmpty then + if let [name] ← resolveGlobalConst name then + msgs ← matchingConstants opts name + logInfoAt tk (.joinSep msgs.toList "") diff --git a/test/print_prefix.lean b/test/print_prefix.lean index 77f89437c1..0d14e95aa3 100644 --- a/test/print_prefix.lean +++ b/test/print_prefix.lean @@ -171,3 +171,7 @@ TestInd.toCtorIdx : TestInd → Nat -/ #guard_msgs in #print prefix TestInd + +-- `#print prefix` does nothing if no identifier is provided +#guard_msgs in +#print prefix From 10130798199d306703dee5ab2567961444ebbd04 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sat, 19 Oct 2024 03:18:55 -0400 Subject: [PATCH 126/177] feat: add `trans` tactic (#1001) --- Batteries.lean | 1 + Batteries/Tactic/Trans.lean | 218 ++++++++++++++++++++++++++++++++++++ test/trans.lean | 107 ++++++++++++++++++ 3 files changed, 326 insertions(+) create mode 100644 Batteries/Tactic/Trans.lean create mode 100644 test/trans.lean diff --git a/Batteries.lean b/Batteries.lean index 7d3486f15b..5c54caecd3 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -92,6 +92,7 @@ import Batteries.Tactic.PrintPrefix import Batteries.Tactic.SeqFocus import Batteries.Tactic.ShowUnused import Batteries.Tactic.SqueezeScope +import Batteries.Tactic.Trans import Batteries.Tactic.Unreachable import Batteries.Tactic.Where import Batteries.Test.Internal.DummyLabelAttr diff --git a/Batteries/Tactic/Trans.lean b/Batteries/Tactic/Trans.lean new file mode 100644 index 0000000000..7070c843ac --- /dev/null +++ b/Batteries/Tactic/Trans.lean @@ -0,0 +1,218 @@ +/- +Copyright (c) 2022 Siddhartha Gadgil. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Siddhartha Gadgil, Mario Carneiro +-/ +import Lean.Elab.Tactic.ElabTerm +import Batteries.Tactic.Alias + +/-! +# `trans` tactic + +This implements the `trans` tactic, which can apply transitivity theorems with an optional middle +variable argument. +-/ + +/-- Compose using transitivity, homogeneous case. -/ +def Trans.simple {r : α → α → Sort _} [Trans r r r] : r a b → r b c → r a c := trans + +@[deprecated (since := "2024-10-18")] +alias Trans.heq := Trans.trans + +namespace Batteries.Tactic +open Lean Meta Elab + +initialize registerTraceClass `Tactic.trans + +/-- Discrimation tree settings for the `trans` extension. -/ +def transExt.config : WhnfCoreConfig := {} + +/-- Environment extension storing transitivity lemmas -/ +initialize transExt : + SimpleScopedEnvExtension (Name × Array DiscrTree.Key) (DiscrTree Name) ← + registerSimpleScopedEnvExtension { + addEntry := fun dt (n, ks) => dt.insertCore ks n + initial := {} + } + +initialize registerBuiltinAttribute { + name := `trans + descr := "transitive relation" + add := fun decl _ kind => MetaM.run' do + let declTy := (← getConstInfo decl).type + let (xs, _, targetTy) ← withReducible <| forallMetaTelescopeReducing declTy + let fail := throwError + "@[trans] attribute only applies to lemmas proving + x ∼ y → y ∼ z → x ∼ z, got {indentExpr declTy} with target {indentExpr targetTy}" + let .app (.app rel _) _ := targetTy | fail + let some yzHyp := xs.back? | fail + let some xyHyp := xs.pop.back? | fail + let .app (.app _ _) _ ← inferType yzHyp | fail + let .app (.app _ _) _ ← inferType xyHyp | fail + let key ← withReducible <| DiscrTree.mkPath rel transExt.config + transExt.add (decl, key) kind +} + +open Lean.Elab.Tactic + +/-- solving `e ← mkAppM' f #[x]` -/ +def getExplicitFuncArg? (e : Expr) : MetaM (Option <| Expr × Expr) := do + match e with + | Expr.app f a => do + if ← isDefEq (← mkAppM' f #[a]) e then + return some (f, a) + else + getExplicitFuncArg? f + | _ => return none + +/-- solving `tgt ← mkAppM' rel #[x, z]` given `tgt = f z` -/ +def getExplicitRelArg? (tgt f z : Expr) : MetaM (Option <| Expr × Expr) := do + match f with + | Expr.app rel x => do + let check: Bool ← do + try + let folded ← mkAppM' rel #[x, z] + isDefEq folded tgt + catch _ => + pure false + if check then + return some (rel, x) + else + getExplicitRelArg? tgt rel z + | _ => return none + +/-- refining `tgt ← mkAppM' rel #[x, z]` dropping more arguments if possible -/ +def getExplicitRelArgCore (tgt rel x z : Expr) : MetaM (Expr × Expr) := do + match rel with + | Expr.app rel' _ => do + let check: Bool ← do + try + let folded ← mkAppM' rel' #[x, z] + isDefEq folded tgt + catch _ => + pure false + if !check then + return (rel, x) + else + getExplicitRelArgCore tgt rel' x z + | _ => return (rel ,x) + +/-- Internal definition for `trans` tactic. Either a binary relation or a non-dependent +arrow. -/ +inductive TransRelation + /-- Expression for transitive relation. -/ + | app (rel : Expr) + /-- Constant name for transitive relation. -/ + | implies (name : Name) (bi : BinderInfo) + +/-- Finds an explicit binary relation in the argument, if possible. -/ +def getRel (tgt : Expr) : MetaM (Option (TransRelation × Expr × Expr)) := do + match tgt with + | .forallE name binderType body info => return .some (.implies name info, binderType, body) + | .app f z => + match (← getExplicitRelArg? tgt f z) with + | some (rel, x) => + let (rel, x) ← getExplicitRelArgCore tgt rel x z + return some (.app rel, x, z) + | none => + return none + | _ => return none + +/-- +`trans` applies to a goal whose target has the form `t ~ u` where `~` is a transitive relation, +that is, a relation which has a transitivity lemma tagged with the attribute [trans]. + +* `trans s` replaces the goal with the two subgoals `t ~ s` and `s ~ u`. +* If `s` is omitted, then a metavariable is used instead. + +Additionally, `trans` also applies to a goal whose target has the form `t → u`, +in which case it replaces the goal with `t → s` and `s → u`. +-/ +elab "trans" t?:(ppSpace colGt term)? : tactic => withMainContext do + let tgt := (← instantiateMVars (← (← getMainGoal).getType)).cleanupAnnotations + let .some (rel, x, z) ← getRel tgt | + throwError (m!"transitivity lemmas only apply to binary relations and " ++ + m!"non-dependent arrows, not {indentExpr tgt}") + match rel with + | .implies name info => + -- only consider non-dependent arrows + if z.hasLooseBVars then + throwError "`trans` is not implemented for dependent arrows{indentExpr tgt}" + -- parse the intermeditate term + let middleType ← mkFreshExprMVar none + let t'? ← t?.mapM (elabTermWithHoles · middleType (← getMainTag)) + let middle ← (t'?.map (pure ·.1)).getD (mkFreshExprMVar middleType) + liftMetaTactic fun goal => do + -- create two new goals + let g₁ ← mkFreshExprMVar (some <| .forallE name x middle info) .synthetic + let g₂ ← mkFreshExprMVar (some <| .forallE name middle z info) .synthetic + -- close the original goal with `fun x => g₂ (g₁ x)` + goal.assign (.lam name x (.app g₂ (.app g₁ (.bvar 0))) .default) + pure <| [g₁.mvarId!, g₂.mvarId!] ++ if let some (_, gs') := t'? then gs' else [middle.mvarId!] + return + | .app rel => + trace[Tactic.trans]"goal decomposed" + trace[Tactic.trans]"rel: {indentExpr rel}" + trace[Tactic.trans]"x: {indentExpr x}" + trace[Tactic.trans]"z: {indentExpr z}" + -- first trying the homogeneous case + try + let ty ← inferType x + let t'? ← t?.mapM (elabTermWithHoles · ty (← getMainTag)) + let s ← saveState + trace[Tactic.trans]"trying homogeneous case" + let lemmas := + (← (transExt.getState (← getEnv)).getUnify rel transExt.config).push ``Trans.simple + for lem in lemmas do + trace[Tactic.trans]"trying lemma {lem}" + try + liftMetaTactic fun g => do + let lemTy ← inferType (← mkConstWithLevelParams lem) + let arity ← withReducible <| forallTelescopeReducing lemTy fun es _ => pure es.size + let y ← (t'?.map (pure ·.1)).getD (mkFreshExprMVar ty) + let g₁ ← mkFreshExprMVar (some <| ← mkAppM' rel #[x, y]) .synthetic + let g₂ ← mkFreshExprMVar (some <| ← mkAppM' rel #[y, z]) .synthetic + g.assign (← mkAppOptM lem (mkArray (arity - 2) none ++ #[some g₁, some g₂])) + pure <| [g₁.mvarId!, g₂.mvarId!] ++ + if let some (_, gs') := t'? then gs' else [y.mvarId!] + return + catch _ => s.restore + pure () + catch _ => + trace[Tactic.trans]"trying heterogeneous case" + let t'? ← t?.mapM (elabTermWithHoles · none (← getMainTag)) + let s ← saveState + for lem in (← (transExt.getState (← getEnv)).getUnify rel transExt.config).push + ``HEq.trans |>.push ``Trans.trans do + try + liftMetaTactic fun g => do + trace[Tactic.trans]"trying lemma {lem}" + let lemTy ← inferType (← mkConstWithLevelParams lem) + let arity ← withReducible <| forallTelescopeReducing lemTy fun es _ => pure es.size + trace[Tactic.trans]"arity: {arity}" + trace[Tactic.trans]"lemma-type: {lemTy}" + let y ← (t'?.map (pure ·.1)).getD (mkFreshExprMVar none) + trace[Tactic.trans]"obtained y: {y}" + trace[Tactic.trans]"rel: {indentExpr rel}" + trace[Tactic.trans]"x:{indentExpr x}" + trace[Tactic.trans]"z: {indentExpr z}" + let g₂ ← mkFreshExprMVar (some <| ← mkAppM' rel #[y, z]) .synthetic + trace[Tactic.trans]"obtained g₂: {g₂}" + let g₁ ← mkFreshExprMVar (some <| ← mkAppM' rel #[x, y]) .synthetic + trace[Tactic.trans]"obtained g₁: {g₁}" + g.assign (← mkAppOptM lem (mkArray (arity - 2) none ++ #[some g₁, some g₂])) + pure <| [g₁.mvarId!, g₂.mvarId!] ++ if let some (_, gs') := t'? then gs' else [y.mvarId!] + return + catch e => + trace[Tactic.trans]"failed: {e.toMessageData}" + s.restore + throwError m!"no applicable transitivity lemma found for {indentExpr tgt}" + +/-- Synonym for `trans` tactic. -/ +syntax "transitivity" (ppSpace colGt term)? : tactic +set_option hygiene false in +macro_rules + | `(tactic| transitivity) => `(tactic| trans) + | `(tactic| transitivity $e) => `(tactic| trans $e) + +end Batteries.Tactic diff --git a/test/trans.lean b/test/trans.lean new file mode 100644 index 0000000000..d05874399b --- /dev/null +++ b/test/trans.lean @@ -0,0 +1,107 @@ +import Batteries.Tactic.Trans + +-- testing that the attribute is recognized and used +def nleq (a b : Nat) : Prop := a ≤ b + +@[trans] def nleq_trans : nleq a b → nleq b c → nleq a c := Nat.le_trans + +example (a b c : Nat) : nleq a b → nleq b c → nleq a c := by + intro h₁ h₂ + trans b + assumption + assumption + +example (a b c : Nat) : nleq a b → nleq b c → nleq a c := by intros; trans <;> assumption + +-- using `Trans` typeclass +@[trans] def eq_trans {a b c : α} : a = b → b = c → a = c := by + intro h₁ h₂ + apply Eq.trans h₁ h₂ + +example (a b c : Nat) : a = b → b = c → a = c := by intros; trans <;> assumption + +example (a b c : Nat) : a = b → b = c → a = c := by + intro h₁ h₂ + trans b + assumption + assumption + +example : @Trans Nat Nat Nat (· ≤ ·) (· ≤ ·) (· ≤ ·) := inferInstance + +example (a b c : Nat) : a ≤ b → b ≤ c → a ≤ c := by + intros h₁ h₂ + trans ?b + case b => exact b + exact h₁ + exact h₂ + +example (a b c : α) (R : α → α → Prop) [Trans R R R] : R a b → R b c → R a c := by + intros h₁ h₂ + trans ?b + case b => exact b + exact h₁ + exact h₂ + +example (a b c : Nat) : a ≤ b → b ≤ c → a ≤ c := by + intros h₁ h₂ + trans + exact h₁ + exact h₂ + +example (a b c : Nat) : a ≤ b → b ≤ c → a ≤ c := by intros; trans <;> assumption + +example (a b c : Nat) : a < b → b < c → a < c := by + intro h₁ h₂ + trans b + assumption + assumption + +example (a b c : Nat) : a < b → b < c → a < c := by intros; trans <;> assumption + +example (x n p : Nat) (h₁ : n * Nat.succ p ≤ x) : n * p ≤ x := by + trans + · apply Nat.mul_le_mul_left; apply Nat.le_succ + · apply h₁ + +example (a : α) (c : γ) : ∀ b : β, HEq a b → HEq b c → HEq a c := by + intro b h₁ h₂ + trans b + assumption + assumption + +def MyLE (n m : Nat) := ∃ k, n + k = m + +@[trans] theorem MyLE.trans {n m k : Nat} (h1 : MyLE n m) (h2 : MyLE m k) : MyLE n k := by + cases h1 + cases h2 + subst_vars + exact ⟨_, Eq.symm <| Nat.add_assoc _ _ _⟩ + +example {n m k : Nat} (h1 : MyLE n m) (h2 : MyLE m k) : MyLE n k := by + trans <;> assumption + +/-- `trans` for implications. -/ +example {A B C : Prop} (h : A → B) (g : B → C) : A → C := by + trans B + · guard_target =ₛ A → B -- ensure we have `B` and not a free metavariable. + exact h + · guard_target =ₛ B → C + exact g + +/-- `trans` for arrows between types. -/ +example {A B C : Type} (h : A → B) (g : B → C) : A → C := by + trans + rotate_right + · exact B + · exact h + · exact g + +universe u v w + +/-- `trans` for arrows between types. -/ +example {A : Type u} {B : Type v} {C : Type w} (h : A → B) (g : B → C) : A → C := by + trans + rotate_right + · exact B + · exact h + · exact g From dd6b1019b5cef990161bf3edfebeb6b0be78044a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 21 Oct 2024 16:45:24 +1100 Subject: [PATCH 127/177] chore: rename List.modifyNth->modify and insertNth->insertIdx (#1003) --- Batteries/Data/List/Basic.lean | 56 ++++---- Batteries/Data/List/Lemmas.lean | 240 ++++++++++++++++---------------- Batteries/Data/List/Perm.lean | 10 +- 3 files changed, 159 insertions(+), 147 deletions(-) diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index 49126faee0..fa1de49516 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -170,10 +170,12 @@ using `f` if the index is larger than the length of the List. modifyNthTail f 2 [a, b, c] = [a, b] ++ f [c] ``` -/ -@[simp] def modifyNthTail (f : List α → List α) : Nat → List α → List α +@[simp] def modifyTailIdx (f : List α → List α) : Nat → List α → List α | 0, l => f l | _+1, [] => [] - | n+1, a :: l => a :: modifyNthTail f n l + | n+1, a :: l => a :: modifyTailIdx f n l + +@[deprecated (since := "2024-10-21")] alias modifyNthTail := modifyTailIdx /-- Apply `f` to the head of the list, if it exists. -/ @[inline] def modifyHead (f : α → α) : List α → List α @@ -185,25 +187,29 @@ modifyNthTail f 2 [a, b, c] = [a, b] ++ f [c] @[simp] theorem modifyHead_cons (a : α) (l : List α) (f : α → α) : (a :: l).modifyHead f = f a :: l := by rw [modifyHead] -/-- Apply `f` to the nth element of the list, if it exists. -/ -def modifyNth (f : α → α) : Nat → List α → List α := - modifyNthTail (modifyHead f) +/-- +Apply `f` to the nth element of the list, if it exists, replacing that element with the result. +-/ +def modify (f : α → α) : Nat → List α → List α := + modifyTailIdx (modifyHead f) -/-- Tail-recursive version of `modifyNth`. -/ -def modifyNthTR (f : α → α) (n : Nat) (l : List α) : List α := go l n #[] where - /-- Auxiliary for `modifyNthTR`: `modifyNthTR.go f l n acc = acc.toList ++ modifyNth f n l`. -/ +@[deprecated (since := "2024-10-21")] alias modifyNth := modify + +/-- Tail-recursive version of `modify`. -/ +def modifyTR (f : α → α) (n : Nat) (l : List α) : List α := go l n #[] where + /-- Auxiliary for `modifyTR`: `modifyTR.go f l n acc = acc.toList ++ modify f n l`. -/ go : List α → Nat → Array α → List α | [], _, acc => acc.toList | a :: l, 0, acc => acc.toListAppend (f a :: l) | a :: l, n+1, acc => go l n (acc.push a) -theorem modifyNthTR_go_eq : ∀ l n, modifyNthTR.go f l n acc = acc.toList ++ modifyNth f n l - | [], n => by cases n <;> simp [modifyNthTR.go, modifyNth] - | a :: l, 0 => by simp [modifyNthTR.go, modifyNth] - | a :: l, n+1 => by simp [modifyNthTR.go, modifyNth, modifyNthTR_go_eq l] +theorem modifyTR_go_eq : ∀ l n, modifyTR.go f l n acc = acc.toList ++ modify f n l + | [], n => by cases n <;> simp [modifyTR.go, modify] + | a :: l, 0 => by simp [modifyTR.go, modify] + | a :: l, n+1 => by simp [modifyTR.go, modify, modifyTR_go_eq l] -@[csimp] theorem modifyNth_eq_modifyNthTR : @modifyNth = @modifyNthTR := by - funext α f n l; simp [modifyNthTR, modifyNthTR_go_eq] +@[csimp] theorem modify_eq_modifyTR : @modify = @modifyTR := by + funext α f n l; simp [modifyTR, modifyTR_go_eq] /-- Apply `f` to the last element of `l`, if it exists. -/ @[inline] def modifyLast (f : α → α) (l : List α) : List α := go l #[] where @@ -219,23 +225,25 @@ theorem modifyNthTR_go_eq : ∀ l n, modifyNthTR.go f l n acc = acc.toList ++ mo insertNth 2 1 [1, 2, 3, 4] = [1, 2, 1, 3, 4] ``` -/ -def insertNth (n : Nat) (a : α) : List α → List α := - modifyNthTail (cons a) n +def insertIdx (n : Nat) (a : α) : List α → List α := + modifyTailIdx (cons a) n + +@[deprecated (since := "2024-10-21")] alias insertNth := insertIdx -/-- Tail-recursive version of `insertNth`. -/ -@[inline] def insertNthTR (n : Nat) (a : α) (l : List α) : List α := go n l #[] where - /-- Auxiliary for `insertNthTR`: `insertNthTR.go a n l acc = acc.toList ++ insertNth n a l`. -/ +/-- Tail-recursive version of `insertIdx`. -/ +@[inline] def insertIdxTR (n : Nat) (a : α) (l : List α) : List α := go n l #[] where + /-- Auxiliary for `insertIdxTR`: `insertIdxTR.go a n l acc = acc.toList ++ insertIdx n a l`. -/ go : Nat → List α → Array α → List α | 0, l, acc => acc.toListAppend (a :: l) | _, [], acc => acc.toList | n+1, a :: l, acc => go n l (acc.push a) -theorem insertNthTR_go_eq : ∀ n l, insertNthTR.go a n l acc = acc.toList ++ insertNth n a l - | 0, l | _+1, [] => by simp [insertNthTR.go, insertNth] - | n+1, a :: l => by simp [insertNthTR.go, insertNth, insertNthTR_go_eq n l] +theorem insertIdxTR_go_eq : ∀ n l, insertIdxTR.go a n l acc = acc.toList ++ insertIdx n a l + | 0, l | _+1, [] => by simp [insertIdxTR.go, insertIdx] + | n+1, a :: l => by simp [insertIdxTR.go, insertIdx, insertIdxTR_go_eq n l] -@[csimp] theorem insertNth_eq_insertNthTR : @insertNth = @insertNthTR := by - funext α f n l; simp [insertNthTR, insertNthTR_go_eq] +@[csimp] theorem insertIdx_eq_insertIdxTR : @insertIdx = @insertIdxTR := by + funext α f n l; simp [insertIdxTR, insertIdxTR_go_eq] theorem headD_eq_head? (l) (a : α) : headD l a = (head? l).getD a := by cases l <;> rfl diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 6a956f9b1d..5179c81ba0 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -15,10 +15,6 @@ namespace List @[simp] theorem getElem_mk {xs : List α} {i : Nat} (h : i < xs.length) : (Array.mk xs)[i] = xs[i] := rfl -/-! ### isEmpty -/ - -@[deprecated (since := "2024-08-15")] alias isEmpty_iff_eq_nil := isEmpty_iff - /-! ### next? -/ @[simp] theorem next?_nil : @next? α [] = none := rfl @@ -40,173 +36,114 @@ theorem dropLast_eq_eraseIdx {xs : List α} {i : Nat} (last_idx : i + 1 = xs.len exact ih last_idx exact fun _ => nomatch xs -/-! ### get? -/ - -@[deprecated getElem_eq_iff (since := "2024-06-12")] -theorem get_eq_iff : List.get l n = x ↔ l.get? n.1 = some x := by - simp - -@[deprecated getElem?_inj (since := "2024-06-12")] -theorem get?_inj - (h₀ : i < xs.length) (h₁ : Nodup xs) (h₂ : xs.get? i = xs.get? j) : i = j := by - apply getElem?_inj h₀ h₁ - simp_all - /-! ### modifyHead -/ @[simp] theorem modifyHead_modifyHead (l : List α) (f g : α → α) : (l.modifyHead f).modifyHead g = l.modifyHead (g ∘ f) := by cases l <;> simp [modifyHead] -/-! ### modifyNth -/ +/-! ### modify -/ -@[simp] theorem modifyNth_nil (f : α → α) (n) : [].modifyNth f n = [] := by cases n <;> rfl +@[simp] theorem modify_nil (f : α → α) (n) : [].modify f n = [] := by cases n <;> rfl -@[simp] theorem modifyNth_zero_cons (f : α → α) (a : α) (l : List α) : - (a :: l).modifyNth f 0 = f a :: l := rfl +@[simp] theorem modify_zero_cons (f : α → α) (a : α) (l : List α) : + (a :: l).modify f 0 = f a :: l := rfl -@[simp] theorem modifyNth_succ_cons (f : α → α) (a : α) (l : List α) (n) : - (a :: l).modifyNth f (n + 1) = a :: l.modifyNth f n := by rfl +@[simp] theorem modify_succ_cons (f : α → α) (a : α) (l : List α) (n) : + (a :: l).modify f (n + 1) = a :: l.modify f n := by rfl -theorem modifyNthTail_id : ∀ n (l : List α), l.modifyNthTail id n = l +theorem modifyTailIdx_id : ∀ n (l : List α), l.modifyTailIdx id n = l | 0, _ => rfl | _+1, [] => rfl - | n+1, a :: l => congrArg (cons a) (modifyNthTail_id n l) + | n+1, a :: l => congrArg (cons a) (modifyTailIdx_id n l) -theorem eraseIdx_eq_modifyNthTail : ∀ n (l : List α), eraseIdx l n = modifyNthTail tail n l +theorem eraseIdx_eq_modifyTailIdx : ∀ n (l : List α), eraseIdx l n = modifyTailIdx tail n l | 0, l => by cases l <;> rfl | _+1, [] => rfl - | _+1, _ :: _ => congrArg (cons _) (eraseIdx_eq_modifyNthTail _ _) - -@[deprecated (since := "2024-05-06")] alias removeNth_eq_nth_tail := eraseIdx_eq_modifyNthTail + | _+1, _ :: _ => congrArg (cons _) (eraseIdx_eq_modifyTailIdx _ _) -theorem getElem?_modifyNth (f : α → α) : - ∀ n (l : List α) m, (modifyNth f n l)[m]? = (fun a => if n = m then f a else a) <$> l[m]? +theorem getElem?_modify (f : α → α) : + ∀ n (l : List α) m, (modify f n l)[m]? = (fun a => if n = m then f a else a) <$> l[m]? | n, l, 0 => by cases l <;> cases n <;> simp | n, [], _+1 => by cases n <;> rfl - | 0, _ :: l, m+1 => by cases h : l[m]? <;> simp [h, modifyNth, m.succ_ne_zero.symm] + | 0, _ :: l, m+1 => by cases h : l[m]? <;> simp [h, modify, m.succ_ne_zero.symm] | n+1, a :: l, m+1 => by - simp only [modifyNth_succ_cons, getElem?_cons_succ, Nat.reduceEqDiff, Option.map_eq_map] - refine (getElem?_modifyNth f n l m).trans ?_ + simp only [modify_succ_cons, getElem?_cons_succ, Nat.reduceEqDiff, Option.map_eq_map] + refine (getElem?_modify f n l m).trans ?_ cases h' : l[m]? <;> by_cases h : n = m <;> simp [h, if_pos, if_neg, Option.map, mt Nat.succ.inj, not_false_iff, h'] -@[deprecated getElem?_modifyNth (since := "2024-06-12")] -theorem get?_modifyNth (f : α → α) (n) (l : List α) (m) : - (modifyNth f n l).get? m = (fun a => if n = m then f a else a) <$> l.get? m := by - simp [getElem?_modifyNth] - -theorem length_modifyNthTail (f : List α → List α) (H : ∀ l, length (f l) = length l) : - ∀ n l, length (modifyNthTail f n l) = length l +theorem length_modifyTailIdx (f : List α → List α) (H : ∀ l, length (f l) = length l) : + ∀ n l, length (modifyTailIdx f n l) = length l | 0, _ => H _ | _+1, [] => rfl - | _+1, _ :: _ => congrArg (·+1) (length_modifyNthTail _ H _ _) - -@[deprecated (since := "2024-06-07")] alias modifyNthTail_length := length_modifyNthTail + | _+1, _ :: _ => congrArg (·+1) (length_modifyTailIdx _ H _ _) -theorem modifyNthTail_add (f : List α → List α) (n) (l₁ l₂ : List α) : - modifyNthTail f (l₁.length + n) (l₁ ++ l₂) = l₁ ++ modifyNthTail f n l₂ := by +theorem modifyTailIdx_add (f : List α → List α) (n) (l₁ l₂ : List α) : + modifyTailIdx f (l₁.length + n) (l₁ ++ l₂) = l₁ ++ modifyTailIdx f n l₂ := by induction l₁ <;> simp [*, Nat.succ_add] -theorem exists_of_modifyNthTail (f : List α → List α) {n} {l : List α} (h : n ≤ l.length) : - ∃ l₁ l₂, l = l₁ ++ l₂ ∧ l₁.length = n ∧ modifyNthTail f n l = l₁ ++ f l₂ := +theorem exists_of_modifyTailIdx (f : List α → List α) {n} {l : List α} (h : n ≤ l.length) : + ∃ l₁ l₂, l = l₁ ++ l₂ ∧ l₁.length = n ∧ modifyTailIdx f n l = l₁ ++ f l₂ := have ⟨_, _, eq, hl⟩ : ∃ l₁ l₂, l = l₁ ++ l₂ ∧ l₁.length = n := ⟨_, _, (take_append_drop n l).symm, length_take_of_le h⟩ - ⟨_, _, eq, hl, hl ▸ eq ▸ modifyNthTail_add (n := 0) ..⟩ + ⟨_, _, eq, hl, hl ▸ eq ▸ modifyTailIdx_add (n := 0) ..⟩ -@[simp] theorem length_modifyNth (f : α → α) : ∀ n l, length (modifyNth f n l) = length l := - length_modifyNthTail _ fun l => by cases l <;> rfl +@[simp] theorem length_modify (f : α → α) : ∀ n l, length (modify f n l) = length l := + length_modifyTailIdx _ fun l => by cases l <;> rfl -@[deprecated (since := "2024-06-07")] alias modify_get?_length := length_modifyNth +@[simp] theorem getElem?_modify_eq (f : α → α) (n) (l : List α) : + (modify f n l)[n]? = f <$> l[n]? := by + simp only [getElem?_modify, if_pos] -@[simp] theorem getElem?_modifyNth_eq (f : α → α) (n) (l : List α) : - (modifyNth f n l)[n]? = f <$> l[n]? := by - simp only [getElem?_modifyNth, if_pos] +@[simp] theorem getElem?_modify_ne (f : α → α) {m n} (l : List α) (h : m ≠ n) : + (modify f m l)[n]? = l[n]? := by + simp only [getElem?_modify, if_neg h, id_map'] -@[deprecated getElem?_modifyNth_eq (since := "2024-06-12")] -theorem get?_modifyNth_eq (f : α → α) (n) (l : List α) : - (modifyNth f n l).get? n = f <$> l.get? n := by - simp [getElem?_modifyNth_eq] - -@[simp] theorem getElem?_modifyNth_ne (f : α → α) {m n} (l : List α) (h : m ≠ n) : - (modifyNth f m l)[n]? = l[n]? := by - simp only [getElem?_modifyNth, if_neg h, id_map'] - -@[deprecated getElem?_modifyNth_ne (since := "2024-06-12")] -theorem get?_modifyNth_ne (f : α → α) {m n} (l : List α) (h : m ≠ n) : - (modifyNth f m l).get? n = l.get? n := by - simp [h] - -theorem exists_of_modifyNth (f : α → α) {n} {l : List α} (h : n < l.length) : - ∃ l₁ a l₂, l = l₁ ++ a :: l₂ ∧ l₁.length = n ∧ modifyNth f n l = l₁ ++ f a :: l₂ := - match exists_of_modifyNthTail _ (Nat.le_of_lt h) with +theorem exists_of_modify (f : α → α) {n} {l : List α} (h : n < l.length) : + ∃ l₁ a l₂, l = l₁ ++ a :: l₂ ∧ l₁.length = n ∧ modify f n l = l₁ ++ f a :: l₂ := + match exists_of_modifyTailIdx _ (Nat.le_of_lt h) with | ⟨_, _::_, eq, hl, H⟩ => ⟨_, _, _, eq, hl, H⟩ | ⟨_, [], eq, hl, _⟩ => nomatch Nat.ne_of_gt h (eq ▸ append_nil _ ▸ hl) -theorem modifyNthTail_eq_take_drop (f : List α → List α) (H : f [] = []) : - ∀ n l, modifyNthTail f n l = take n l ++ f (drop n l) +theorem modifyTailIdx_eq_take_drop (f : List α → List α) (H : f [] = []) : + ∀ n l, modifyTailIdx f n l = take n l ++ f (drop n l) | 0, _ => rfl | _ + 1, [] => H.symm - | n + 1, b :: l => congrArg (cons b) (modifyNthTail_eq_take_drop f H n l) + | n + 1, b :: l => congrArg (cons b) (modifyTailIdx_eq_take_drop f H n l) -theorem modifyNth_eq_take_drop (f : α → α) : - ∀ n l, modifyNth f n l = take n l ++ modifyHead f (drop n l) := - modifyNthTail_eq_take_drop _ rfl +theorem modify_eq_take_drop (f : α → α) : + ∀ n l, modify f n l = take n l ++ modifyHead f (drop n l) := + modifyTailIdx_eq_take_drop _ rfl -theorem modifyNth_eq_take_cons_drop (f : α → α) {n l} (h : n < length l) : - modifyNth f n l = take n l ++ f l[n] :: drop (n + 1) l := by - rw [modifyNth_eq_take_drop, drop_eq_getElem_cons h]; rfl +theorem modify_eq_take_cons_drop (f : α → α) {n l} (h : n < length l) : + modify f n l = take n l ++ f l[n] :: drop (n + 1) l := by + rw [modify_eq_take_drop, drop_eq_getElem_cons h]; rfl /-! ### set -/ -theorem set_eq_modifyNth (a : α) : ∀ n (l : List α), set l n a = modifyNth (fun _ => a) n l +theorem set_eq_modify (a : α) : ∀ n (l : List α), set l n a = modify (fun _ => a) n l | 0, l => by cases l <;> rfl | _+1, [] => rfl - | _+1, _ :: _ => congrArg (cons _) (set_eq_modifyNth _ _ _) + | _+1, _ :: _ => congrArg (cons _) (set_eq_modify _ _ _) theorem set_eq_take_cons_drop (a : α) {n l} (h : n < length l) : set l n a = take n l ++ a :: drop (n + 1) l := by - rw [set_eq_modifyNth, modifyNth_eq_take_cons_drop _ h] + rw [set_eq_modify, modify_eq_take_cons_drop _ h] -theorem modifyNth_eq_set_get? (f : α → α) : - ∀ n (l : List α), l.modifyNth f n = ((fun a => l.set n (f a)) <$> l.get? n).getD l +theorem modify_eq_set_get? (f : α → α) : + ∀ n (l : List α), l.modify f n = ((fun a => l.set n (f a)) <$> l.get? n).getD l | 0, l => by cases l <;> rfl | _+1, [] => rfl | n+1, b :: l => - (congrArg (cons _) (modifyNth_eq_set_get? ..)).trans <| by cases h : l[n]? <;> simp [h] - -theorem modifyNth_eq_set_get (f : α → α) {n} {l : List α} (h) : - l.modifyNth f n = l.set n (f (l.get ⟨n, h⟩)) := by - rw [modifyNth_eq_set_get?, get?_eq_get h]; rfl + (congrArg (cons _) (modify_eq_set_get? ..)).trans <| by cases h : l[n]? <;> simp [h] --- The naming of `exists_of_set'` and `exists_of_set` have been swapped. --- If no one complains, we will remove this version later. -@[deprecated exists_of_set (since := "2024-07-04")] -theorem exists_of_set' {l : List α} (h : n < l.length) : - ∃ l₁ a l₂, l = l₁ ++ a :: l₂ ∧ l₁.length = n ∧ l.set n a' = l₁ ++ a' :: l₂ := by - rw [set_eq_modifyNth]; exact exists_of_modifyNth _ h - -@[deprecated getElem?_set_self' (since := "2024-06-12")] -theorem get?_set_eq (a : α) (n) (l : List α) : (set l n a).get? n = (fun _ => a) <$> l.get? n := by - simp only [get?_eq_getElem?, getElem?_set_self', Option.map_eq_map] - rfl +theorem modify_eq_set_get (f : α → α) {n} {l : List α} (h) : + l.modify f n = l.set n (f (l.get ⟨n, h⟩)) := by + rw [modify_eq_set_get?, get?_eq_get h]; rfl theorem getElem?_set_eq_of_lt (a : α) {n} {l : List α} (h : n < length l) : (set l n a)[n]? = some a := by rw [getElem?_set_self', getElem?_eq_getElem h]; rfl -@[deprecated getElem?_set_eq_of_lt (since := "2024-06-12")] -theorem get?_set_eq_of_lt (a : α) {n} {l : List α} (h : n < length l) : - (set l n a).get? n = some a := by - rw [get?_eq_getElem?, getElem?_set_self', getElem?_eq_getElem h]; rfl - -@[deprecated getElem?_set_ne (since := "2024-06-12")] -theorem get?_set_ne (a : α) {m n} (l : List α) (h : m ≠ n) : (set l m a).get? n = l.get? n := by - simp [h] - -@[deprecated getElem?_set (since := "2024-06-12")] -theorem get?_set (a : α) {m n} (l : List α) : - (set l m a).get? n = if m = n then (fun _ => a) <$> l.get? n else l.get? n := by - simp [getElem?_set']; rfl - theorem get?_set_of_lt (a : α) {m n} (l : List α) (h : n < length l) : (set l m a).get? n = if m = n then some a else l.get? n := by simp [getElem?_set', getElem?_eq_getElem h] @@ -215,8 +152,6 @@ theorem get?_set_of_lt' (a : α) {m n} (l : List α) (h : m < length l) : (set l m a).get? n = if m = n then some a else l.get? n := by simp [getElem?_set]; split <;> subst_vars <;> simp [*, getElem?_eq_getElem h] -@[deprecated (since := "2024-05-06")] alias length_removeNth := length_eraseIdx - /-! ### tail -/ theorem length_tail_add_one (l : List α) (h : 0 < length l) : (length (tail l)) + 1 = length l := by @@ -236,8 +171,6 @@ theorem length_tail_add_one (l : List α) (h : 0 < length l) : (length (tail l)) /-! ### erase -/ -@[deprecated (since := "2024-04-22")] alias sublist.erase := Sublist.erase - theorem erase_eq_self_iff_forall_bne [BEq α] (a : α) (xs : List α) : xs.erase a = xs ↔ ∀ (x : α), x ∈ xs → ¬x == a := by rw [erase_eq_eraseP', eraseP_eq_self_iff] @@ -667,3 +600,72 @@ theorem foldlM_map [Monad m] (f : β₁ → β₂) (g : α → β₂ → m α) ( theorem foldrM_map [Monad m] [LawfulMonad m] (f : β₁ → β₂) (g : β₂ → α → m α) (l : List β₁) (init : α) : (l.map f).foldrM g init = l.foldrM (fun x y => g (f x) y) init := by induction l generalizing g init <;> simp [*] + +/-! ### deprecations -/ + +@[deprecated (since := "2024-08-15")] alias isEmpty_iff_eq_nil := isEmpty_iff +@[deprecated getElem_eq_iff (since := "2024-06-12")] +theorem get_eq_iff : List.get l n = x ↔ l.get? n.1 = some x := by + simp +@[deprecated getElem?_inj (since := "2024-06-12")] +theorem get?_inj + (h₀ : i < xs.length) (h₁ : Nodup xs) (h₂ : xs.get? i = xs.get? j) : i = j := by + apply getElem?_inj h₀ h₁ + simp_all +@[deprecated (since := "2024-10-21")] alias modifyNth_nil := modify_nil +@[deprecated (since := "2024-10-21")] alias modifyNth_zero_cons := modify_zero_cons +@[deprecated (since := "2024-10-21")] alias modifyNth_succ_cons := modify_succ_cons +@[deprecated (since := "2024-10-21")] alias modifyNthTail_id := modifyTailIdx_id +@[deprecated (since := "2024-10-21")] alias eraseIdx_eq_modifyNthTail := eraseIdx_eq_modifyTailIdx +@[deprecated (since := "2024-05-06")] alias removeNth_eq_nth_tail := eraseIdx_eq_modifyTailIdx +@[deprecated (since := "2024-10-21")] alias getElem?_modifyNth := getElem?_modify +@[deprecated getElem?_modify (since := "2024-06-12")] +theorem get?_modifyNth (f : α → α) (n) (l : List α) (m) : + (modify f n l).get? m = (fun a => if n = m then f a else a) <$> l.get? m := by + simp [getElem?_modify] +@[deprecated (since := "2024-10-21")] alias length_modifyNthTail := length_modifyTailIdx +@[deprecated (since := "2024-06-07")] alias modifyNthTail_length := length_modifyTailIdx +@[deprecated (since := "2024-10-21")] alias modifyNthTail_add := modifyTailIdx_add +@[deprecated (since := "2024-10-21")] alias exists_of_modifyNthTail := exists_of_modifyTailIdx +@[deprecated (since := "2024-10-21")] alias length_modifyNth := length_modify +@[deprecated (since := "2024-06-07")] alias modifyNth_get?_length := length_modify +@[deprecated (since := "2024-10-21")] alias getElem?_modifyNth_eq := getElem?_modify_eq +@[deprecated getElem?_modify_eq (since := "2024-06-12")] +theorem get?_modifyNth_eq (f : α → α) (n) (l : List α) : + (modify f n l).get? n = f <$> l.get? n := by + simp [getElem?_modify_eq] +@[deprecated (since := "2024-06-12")] alias getElem?_modifyNth_ne := getElem?_modify_ne +@[deprecated getElem?_modify_ne (since := "2024-06-12")] +theorem get?_modifyNth_ne (f : α → α) {m n} (l : List α) (h : m ≠ n) : + (modify f m l).get? n = l.get? n := by + simp [h] +@[deprecated (since := "2024-10-21")] alias exists_of_modifyNth := exists_of_modify +@[deprecated (since := "2024-10-21")] alias modifyNthTail_eq_take_drop := modifyTailIdx_eq_take_drop +@[deprecated (since := "2024-10-21")] alias modifyNth_eq_take_drop := modify_eq_take_drop +@[deprecated (since := "2024-10-21")] alias modifyNth_eq_take_cons_drop := modify_eq_take_cons_drop +@[deprecated (since := "2024-10-21")] alias set_eq_modifyNth := set_eq_modify +@[deprecated (since := "2024-10-21")] alias modifyNth_eq_set_get? := modify_eq_set_get? +@[deprecated (since := "2024-10-21")] alias modifyNth_eq_set_get := modify_eq_set_get +-- The naming of `exists_of_set'` and `exists_of_set` have been swapped. +-- If no one complains, we will remove this version later. +@[deprecated exists_of_set (since := "2024-07-04")] +theorem exists_of_set' {l : List α} (h : n < l.length) : + ∃ l₁ a l₂, l = l₁ ++ a :: l₂ ∧ l₁.length = n ∧ l.set n a' = l₁ ++ a' :: l₂ := by + rw [set_eq_modify]; exact exists_of_modify _ h +@[deprecated getElem?_set_self' (since := "2024-06-12")] +theorem get?_set_eq (a : α) (n) (l : List α) : (set l n a).get? n = (fun _ => a) <$> l.get? n := by + simp only [get?_eq_getElem?, getElem?_set_self', Option.map_eq_map] + rfl +@[deprecated getElem?_set_eq_of_lt (since := "2024-06-12")] +theorem get?_set_eq_of_lt (a : α) {n} {l : List α} (h : n < length l) : + (set l n a).get? n = some a := by + rw [get?_eq_getElem?, getElem?_set_self', getElem?_eq_getElem h]; rfl +@[deprecated getElem?_set_ne (since := "2024-06-12")] +theorem get?_set_ne (a : α) {m n} (l : List α) (h : m ≠ n) : (set l m a).get? n = l.get? n := by + simp [h] +@[deprecated getElem?_set (since := "2024-06-12")] +theorem get?_set (a : α) {m n} (l : List α) : + (set l m a).get? n = if m = n then (fun _ => a) <$> l.get? n else l.get? n := by + simp [getElem?_set']; rfl +@[deprecated (since := "2024-05-06")] alias length_removeNth := length_eraseIdx +@[deprecated (since := "2024-04-22")] alias sublist.erase := Sublist.erase diff --git a/Batteries/Data/List/Perm.lean b/Batteries/Data/List/Perm.lean index e72ab8f624..2196d88367 100644 --- a/Batteries/Data/List/Perm.lean +++ b/Batteries/Data/List/Perm.lean @@ -275,8 +275,8 @@ theorem Subperm.cons_left {l₁ l₂ : List α} (h : l₁ <+~ l₂) (x : α) (hx refine h y ?_ simpa [hy'] using hy -theorem perm_insertNth {α} (x : α) (l : List α) {n} (h : n ≤ l.length) : - insertNth n x l ~ x :: l := by +theorem perm_insertIdx {α} (x : α) (l : List α) {n} (h : n ≤ l.length) : + insertIdx n x l ~ x :: l := by induction l generalizing n with | nil => cases n with @@ -284,11 +284,13 @@ theorem perm_insertNth {α} (x : α) (l : List α) {n} (h : n ≤ l.length) : | succ => cases h | cons _ _ ih => cases n with - | zero => simp [insertNth] + | zero => simp [insertIdx] | succ => - simp only [insertNth, modifyNthTail] + simp only [insertIdx, modifyTailIdx] refine .trans (.cons _ (ih (Nat.le_of_succ_le_succ h))) (.swap ..) +@[deprecated (since := "2024-10-21")] alias perm_insertNth := perm_insertIdx + theorem Perm.union_right {l₁ l₂ : List α} (t₁ : List α) (h : l₁ ~ l₂) : l₁ ∪ t₁ ~ l₂ ∪ t₁ := by induction h with | nil => rfl From 7c5548eeeb1748da5c2872dbd866d3acbaea4b3f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 22 Oct 2024 10:42:29 +1100 Subject: [PATCH 128/177] chore: fix List.modify/insertIdx doc-strings (#1006) --- Batteries/Data/List/Basic.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index fa1de49516..c4c262a016 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -167,7 +167,7 @@ Split a list at every occurrence of a separator element. The separators are not Apply a function to the nth tail of `l`. Returns the input without using `f` if the index is larger than the length of the List. ``` -modifyNthTail f 2 [a, b, c] = [a, b] ++ f [c] +modifyTailIdx f 2 [a, b, c] = [a, b] ++ f [c] ``` -/ @[simp] def modifyTailIdx (f : List α → List α) : Nat → List α → List α @@ -220,9 +220,9 @@ theorem modifyTR_go_eq : ∀ l n, modifyTR.go f l n acc = acc.toList ++ modify f | x :: xs, acc => go xs (acc.push x) /-- -`insertNth n a l` inserts `a` into the list `l` after the first `n` elements of `l` +`insertIdx n a l` inserts `a` into the list `l` after the first `n` elements of `l` ``` -insertNth 2 1 [1, 2, 3, 4] = [1, 2, 1, 3, 4] +insertIdx 2 1 [1, 2, 3, 4] = [1, 2, 1, 3, 4] ``` -/ def insertIdx (n : Nat) (a : α) : List α → List α := From dc72dcdb8e97b3c56bd70f06f043ed2dee3258e6 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 22 Oct 2024 23:54:51 +1100 Subject: [PATCH 129/177] chore: switch to lakefile.toml (#1005) --- lakefile.lean | 18 ------------------ lakefile.toml | 19 +++++++++++++++++++ 2 files changed, 19 insertions(+), 18 deletions(-) delete mode 100644 lakefile.lean create mode 100644 lakefile.toml diff --git a/lakefile.lean b/lakefile.lean deleted file mode 100644 index 7c287c4c95..0000000000 --- a/lakefile.lean +++ /dev/null @@ -1,18 +0,0 @@ -import Lake - -open Lake DSL - -package batteries where - leanOptions := #[⟨`linter.missingDocs, true⟩] - -@[default_target] -lean_lib Batteries - -@[default_target, lint_driver] -lean_exe runLinter where - srcDir := "scripts" - supportInterpreter := true - -@[test_driver] -lean_exe test where - srcDir := "scripts" diff --git a/lakefile.toml b/lakefile.toml new file mode 100644 index 0000000000..a8a658ef9f --- /dev/null +++ b/lakefile.toml @@ -0,0 +1,19 @@ +name = "batteries" +testDriver = "test" +lintDriver = "runLinter" +defaultTargets = ["Batteries", "runLinter"] + +[leanOptions] +linter.missingDocs = true + +[[lean_lib]] +name = "Batteries" + +[[lean_exe]] +name = "runLinter" +srcDir = "scripts" +supportInterpreter = true + +[[lean_exe]] +name = "test" +srcDir = "scripts" From 4d2cb85d87e0e728520fe0bf531cfdddb78fde7a Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Sun, 27 Oct 2024 19:21:44 +0100 Subject: [PATCH 130/177] feat: add `#help note` command (#948) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Blizzard_inc Co-authored-by: François G. Dorais --- Batteries.lean | 2 + Batteries/Tactic/HelpCmd.lean | 49 +++++++++++++++++-- Batteries/Test/Internal/DummyLibraryNote.lean | 14 ++++++ .../Test/Internal/DummyLibraryNote2.lean | 15 ++++++ Batteries/Util/LibraryNote.lean | 3 ++ test/library_note.lean | 47 ++++++++++++++++++ 6 files changed, 126 insertions(+), 4 deletions(-) create mode 100644 Batteries/Test/Internal/DummyLibraryNote.lean create mode 100644 Batteries/Test/Internal/DummyLibraryNote2.lean create mode 100644 test/library_note.lean diff --git a/Batteries.lean b/Batteries.lean index 5c54caecd3..c9358a787a 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -96,6 +96,8 @@ import Batteries.Tactic.Trans import Batteries.Tactic.Unreachable import Batteries.Tactic.Where import Batteries.Test.Internal.DummyLabelAttr +import Batteries.Test.Internal.DummyLibraryNote +import Batteries.Test.Internal.DummyLibraryNote2 import Batteries.Util.Cache import Batteries.Util.ExtendedBinder import Batteries.Util.LibraryNote diff --git a/Batteries/Tactic/HelpCmd.lean b/Batteries/Tactic/HelpCmd.lean index cf1c0f2646..7fbe67bf3a 100644 --- a/Batteries/Tactic/HelpCmd.lean +++ b/Batteries/Tactic/HelpCmd.lean @@ -1,10 +1,11 @@ /- -Copyright (c) 2022 Mario Carneiro. All rights reserved. +Copyright (c) 2024 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro +Authors: Mario Carneiro, Edward van de Meent -/ import Lean.Elab.Syntax import Lean.DocString +import Batteries.Util.LibraryNote /-! @@ -19,9 +20,11 @@ The `#help` command can be used to list all definitions in a variety of extensib * `#help term`, `#help tactic`, `#help conv`, `#help command` are shorthand for `#help cat term` etc. * `#help cat+ C` also shows `elab` and `macro` definitions associated to the syntaxes +* `#help note "some note"` lists library notes for which "some note" is a prefix of the label -All forms take an optional identifier to narrow the search; for example `#help option pp` shows -only `pp.*` options. +Most forms take an optional identifier to narrow the search; for example `#help option pp` shows +only `pp.*` options. However, `#help cat` makes the identifier mandatory, while `#help note` takes +a mandatory string literal, rather than an identifier. -/ @@ -276,6 +279,44 @@ elab_rules : command | `(#help cat $[+%$more]? $cat $id:ident) => elabHelpCat more cat (id.getId.toString false) | `(#help cat $[+%$more]? $cat $id:str) => elabHelpCat more cat id.getString +/-- +format the string to be included in a single markdown bullet +-/ +def _root_.String.makeBullet (s:String) := "* " ++ ("\n ").intercalate (s.splitOn "\n") + +open Lean Parser Batteries.Util.LibraryNote in +/-- +`#help note "foo"` searches for all library notes whose +label starts with "foo", then displays those library notes sorted alphabetically by label, +grouped by label. +The command only displays the library notes that are declared in +imported files or in the same file above the line containing the command. +-/ +elab "#help " colGt &"note" colGt ppSpace name:strLit : command => do + let env ← getEnv + + -- get the library notes from both this and imported files + let local_entries := (libraryNoteExt.getEntries env).reverse + let imported_entries := (libraryNoteExt.toEnvExtension.getState env).importedEntries + + -- filter for the appropriate notes while casting to list + let label_prefix := name.getString + let imported_entries_filtered := imported_entries.flatten.toList.filterMap + fun x => if label_prefix.isPrefixOf x.fst then some x else none + let valid_entries := imported_entries_filtered ++ local_entries.filterMap + fun x => if label_prefix.isPrefixOf x.fst then some x else none + let grouped_valid_entries := valid_entries.mergeSort (·.fst ≤ ·.fst) + |>.groupBy (·.fst == ·.fst) + + -- display results in a readable style + if grouped_valid_entries.isEmpty then + logError "Note not found" + else + logInfo <| "\n\n".intercalate <| + grouped_valid_entries.map + fun l => "library_note \"" ++ l.head!.fst ++ "\"\n" ++ + "\n\n".intercalate (l.map (·.snd.trim.makeBullet)) + /-- The command `#help term` shows all term syntaxes that have been defined in the current environment. See `#help cat` for more information. diff --git a/Batteries/Test/Internal/DummyLibraryNote.lean b/Batteries/Test/Internal/DummyLibraryNote.lean new file mode 100644 index 0000000000..f934823474 --- /dev/null +++ b/Batteries/Test/Internal/DummyLibraryNote.lean @@ -0,0 +1,14 @@ +import Batteries.Util.LibraryNote + +library_note "test" /-- +1: This is a testnote for testing the library note feature of batteries. +The `#help note` command should be able to find this note when imported. +-/ + +library_note "test" /-- +2: This is a second testnote for testing the library note feature of batteries. +-/ + +library_note "temporary note" /-- +1: This is a testnote whose label also starts with "te", but gets sorted before "test" +-/ diff --git a/Batteries/Test/Internal/DummyLibraryNote2.lean b/Batteries/Test/Internal/DummyLibraryNote2.lean new file mode 100644 index 0000000000..3a5bc35dc0 --- /dev/null +++ b/Batteries/Test/Internal/DummyLibraryNote2.lean @@ -0,0 +1,15 @@ +import Batteries.Test.Internal.DummyLibraryNote + +library_note "test" /-- +3: this is a note in a different file importing the above testnotes, +but still imported by the actual testfile. +-/ + +library_note "Test" /-- +1: this is a testnote with a label starting with "Te" +-/ + +library_note "Other" /-- +1: this is a testnote with a label not starting with "te", +so it shouldn't appear when looking for notes with label starting with "te". +-/ diff --git a/Batteries/Util/LibraryNote.lean b/Batteries/Util/LibraryNote.lean index 8c1d45f786..3d578c342b 100644 --- a/Batteries/Util/LibraryNote.lean +++ b/Batteries/Util/LibraryNote.lean @@ -15,6 +15,7 @@ open Lean /-- A library note consists of a (short) tag and a (long) note. -/ def LibraryNoteEntry := String × String +deriving Inhabited /-- Environment extension supporting `library_note`. -/ initialize libraryNoteExt : SimplePersistentEnvExtension LibraryNoteEntry (Array LibraryNoteEntry) ← @@ -35,6 +36,8 @@ creates a new "library note", which can then be cross-referenced using -- See note [some tag] ``` in doc-comments. +Use `#help note "some tag"` to display all notes with the tag `"some tag"` in the infoview. +This command can be imported from Batteries.Tactic.HelpCmd . -/ elab "library_note " title:strLit ppSpace text:docComment : command => do modifyEnv (libraryNoteExt.addEntry · (title.getString, text.getDocString)) diff --git a/test/library_note.lean b/test/library_note.lean new file mode 100644 index 0000000000..f7ac514cdc --- /dev/null +++ b/test/library_note.lean @@ -0,0 +1,47 @@ +import Batteries.Tactic.HelpCmd +import Batteries.Test.Internal.DummyLibraryNote2 + +/-- +error: Note not found +-/ +#guard_msgs in +#help note "no note" + +/-- +info: library_note "Other" +* 1: this is a testnote with a label not starting with "te", + so it shouldn't appear when looking for notes with label starting with "te". +-/ +#guard_msgs in +#help note "Other" + +library_note "test"/-- +4: This note was not imported, and therefore appears below the imported notes. +-/ + +library_note "test"/-- +5: This note was also not imported, and therefore appears below the imported notes, +and the previously added note. +-/ + + +/-- +info: library_note "temporary note" +* 1: This is a testnote whose label also starts with "te", but gets sorted before "test" + +library_note "test" +* 1: This is a testnote for testing the library note feature of batteries. + The `#help note` command should be able to find this note when imported. + +* 2: This is a second testnote for testing the library note feature of batteries. + +* 3: this is a note in a different file importing the above testnotes, + but still imported by the actual testfile. + +* 4: This note was not imported, and therefore appears below the imported notes. + +* 5: This note was also not imported, and therefore appears below the imported notes, + and the previously added note. +-/ +#guard_msgs in +#help note "te" From cbe41b9f291422ec881530a3ec0255678623897f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sun, 27 Oct 2024 19:21:13 -0400 Subject: [PATCH 131/177] feat: use gh command to get pull request info (#1008) --- .github/workflows/test_mathlib.yml | 21 +++++++-------------- 1 file changed, 7 insertions(+), 14 deletions(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index a1af0d771d..c3904f3227 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -12,21 +12,14 @@ jobs: runs-on: ubuntu-latest if: github.event.workflow_run.conclusion == 'success' && github.event.workflow_run.event == 'pull_request' && github.repository == 'leanprover-community/batteries' steps: - - name: Retrieve PR information + - name: Checkout PR + uses: actions/checkout@v4 + + - name: Get PR info id: pr-info - uses: actions/github-script@v6 - with: - script: | - const prNumber = context.payload.workflow_run.pull_requests[0].number; - const { data: pr } = await github.rest.pulls.get({ - owner: context.repo.owner, - repo: context.repo.repo, - pull_number: prNumber - }); - core.setOutput('targetBranch', pr.base.ref); - core.setOutput('pullRequestNumber', pr.number); - core.exportVariable('HEAD_REPO', pr.head.repo.full_name); - core.exportVariable('HEAD_BRANCH', pr.head.ref); + run: | + echo "pullRequestNumber=$(gh pr --json number -q .number || echo '')" >> $GITHUB_OUTPUTS + echo "targetBranch=$(gh pr --json baseRefName -q .baseRefName || echo '')" >> $GITHUB_OUTPUTS - name: Checkout mathlib4 repository if: steps.pr-info.outputs.pullRequestNumber != '' && steps.pr-info.outputs.targetBranch == 'main' From 75fb097c5dec0fb3951322dec96602eb2cc34581 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sun, 27 Oct 2024 19:54:27 -0400 Subject: [PATCH 132/177] fix: typo in script (#1009) --- .github/workflows/test_mathlib.yml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index c3904f3227..6b5c7a4d94 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -18,8 +18,8 @@ jobs: - name: Get PR info id: pr-info run: | - echo "pullRequestNumber=$(gh pr --json number -q .number || echo '')" >> $GITHUB_OUTPUTS - echo "targetBranch=$(gh pr --json baseRefName -q .baseRefName || echo '')" >> $GITHUB_OUTPUTS + echo "pullRequestNumber=$(gh pr view --json number -q .number || echo '')" >> $GITHUB_OUTPUTS + echo "targetBranch=$(gh pr view --json baseRefName -q .baseRefName || echo '')" >> $GITHUB_OUTPUTS - name: Checkout mathlib4 repository if: steps.pr-info.outputs.pullRequestNumber != '' && steps.pr-info.outputs.targetBranch == 'main' From db2c5f8dd36bfab847ef42f5749cc4ba347be202 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sun, 27 Oct 2024 20:13:15 -0400 Subject: [PATCH 133/177] fix: missing github token (#1010) --- .github/workflows/test_mathlib.yml | 2 ++ 1 file changed, 2 insertions(+) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index 6b5c7a4d94..ac3cb13d73 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -20,6 +20,8 @@ jobs: run: | echo "pullRequestNumber=$(gh pr view --json number -q .number || echo '')" >> $GITHUB_OUTPUTS echo "targetBranch=$(gh pr view --json baseRefName -q .baseRefName || echo '')" >> $GITHUB_OUTPUTS + env: + GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} - name: Checkout mathlib4 repository if: steps.pr-info.outputs.pullRequestNumber != '' && steps.pr-info.outputs.targetBranch == 'main' From a4f1acd436c7475dad0ba262f1872ad1a99fb0d2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sun, 27 Oct 2024 20:25:23 -0400 Subject: [PATCH 134/177] fix: github script typo (#1011) --- .github/workflows/test_mathlib.yml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index ac3cb13d73..c367804aa3 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -18,8 +18,8 @@ jobs: - name: Get PR info id: pr-info run: | - echo "pullRequestNumber=$(gh pr view --json number -q .number || echo '')" >> $GITHUB_OUTPUTS - echo "targetBranch=$(gh pr view --json baseRefName -q .baseRefName || echo '')" >> $GITHUB_OUTPUTS + echo "pullRequestNumber=$(gh pr view --json number -q .number || echo '')" >> $GITHUB_OUTPUT + echo "targetBranch=$(gh pr view --json baseRefName -q .baseRefName || echo '')" >> $GITHUB_OUTPUT env: GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} From 19a1ab2cf3eafc9f4147b2ecf176f10f6d0178ca Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sun, 27 Oct 2024 21:15:00 -0400 Subject: [PATCH 135/177] fix: github script fix repo (#1012) --- .github/workflows/test_mathlib.yml | 10 +++++++++- 1 file changed, 9 insertions(+), 1 deletion(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index c367804aa3..3a6ece425e 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -14,12 +14,19 @@ jobs: steps: - name: Checkout PR uses: actions/checkout@v4 + with: + repository: ${{ github.event.workflow_run.head_repository }} + ref: ${{github.event.workflow_run.head_branch }} + fetch-depth: 0 - name: Get PR info id: pr-info run: | echo "pullRequestNumber=$(gh pr view --json number -q .number || echo '')" >> $GITHUB_OUTPUT echo "targetBranch=$(gh pr view --json baseRefName -q .baseRefName || echo '')" >> $GITHUB_OUTPUT + echo "HEAD_BRANCH=$(gh pr view --json headRefName -q .headRefName || echo '')" >> $GITHUB_ENV + echo "HEAD_REPO=$(gh pr view --json headRepository -q .headRepository.name || echo '')" >> $GITHUB_ENV + echo "HEAD_OWNER=$(gh pr view --json headRepositoryOwner-q .headRepositoryOwner.login || echo '')" >> $GITHUB_ENV env: GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} @@ -45,13 +52,14 @@ jobs: id: check_mathlib_tag env: PR_NUMBER: ${{ steps.pr-info.outputs.pullRequestNumber }} + HEAD_OWNER: ${{ env.HEAD_OWNER }} HEAD_REPO: ${{ env.HEAD_REPO }} HEAD_BRANCH: ${{ env.HEAD_BRANCH }} run: | git config user.name "leanprover-community-mathlib4-bot" git config user.email "leanprover-community-mathlib4-bot@users.noreply.github.com" - echo "PR info: $HEAD_REPO $HEAD_BRANCH" + echo "PR info: $HEAD_OWNER $HEAD_REPO $HEAD_BRANCH" BASE=master echo "Using base tag: $BASE" From 04a3779a32afba98df5c012e9cc250df58d37d08 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sun, 27 Oct 2024 21:23:38 -0400 Subject: [PATCH 136/177] fix: github script fix repo again (#1013) --- .github/workflows/test_mathlib.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index 3a6ece425e..6296ae38c1 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -15,7 +15,7 @@ jobs: - name: Checkout PR uses: actions/checkout@v4 with: - repository: ${{ github.event.workflow_run.head_repository }} + repository: ${{ github.event.workflow_run.head_repository.full_name }} ref: ${{github.event.workflow_run.head_branch }} fetch-depth: 0 From 85f6511d93f9e6a8188ec9985c82f08f65c26cae Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sun, 27 Oct 2024 21:31:30 -0400 Subject: [PATCH 137/177] fix: github script another typo (#1014) --- .github/workflows/test_mathlib.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index 6296ae38c1..e3bf84d3ca 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -26,7 +26,7 @@ jobs: echo "targetBranch=$(gh pr view --json baseRefName -q .baseRefName || echo '')" >> $GITHUB_OUTPUT echo "HEAD_BRANCH=$(gh pr view --json headRefName -q .headRefName || echo '')" >> $GITHUB_ENV echo "HEAD_REPO=$(gh pr view --json headRepository -q .headRepository.name || echo '')" >> $GITHUB_ENV - echo "HEAD_OWNER=$(gh pr view --json headRepositoryOwner-q .headRepositoryOwner.login || echo '')" >> $GITHUB_ENV + echo "HEAD_OWNER=$(gh pr view --json headRepositoryOwner -q .headRepositoryOwner.login || echo '')" >> $GITHUB_ENV env: GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} From 6806fd0ab4c002acc24c00d8e573d33bfefee440 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 28 Oct 2024 13:11:31 -0400 Subject: [PATCH 138/177] fix: github script use list instead of view (#1015) --- .github/workflows/test_mathlib.yml | 13 ++++++------- 1 file changed, 6 insertions(+), 7 deletions(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index e3bf84d3ca..5347db224b 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -15,20 +15,19 @@ jobs: - name: Checkout PR uses: actions/checkout@v4 with: - repository: ${{ github.event.workflow_run.head_repository.full_name }} - ref: ${{github.event.workflow_run.head_branch }} fetch-depth: 0 - name: Get PR info id: pr-info run: | - echo "pullRequestNumber=$(gh pr view --json number -q .number || echo '')" >> $GITHUB_OUTPUT - echo "targetBranch=$(gh pr view --json baseRefName -q .baseRefName || echo '')" >> $GITHUB_OUTPUT - echo "HEAD_BRANCH=$(gh pr view --json headRefName -q .headRefName || echo '')" >> $GITHUB_ENV - echo "HEAD_REPO=$(gh pr view --json headRepository -q .headRepository.name || echo '')" >> $GITHUB_ENV - echo "HEAD_OWNER=$(gh pr view --json headRepositoryOwner -q .headRepositoryOwner.login || echo '')" >> $GITHUB_ENV + echo "pullRequestNumber=$(gh pr list $SHA --json number -q '.[0].number' || echo '')" >> $GITHUB_OUTPUT + echo "targetBranch=$(gh pr list $SHA --json baseRefName -q '.[0].baseRefName' || echo '')" >> $GITHUB_OUTPUT + echo "HEAD_BRANCH=$(gh pr list $SHA --json headRefName -q '.[0].headRefName' || echo '')" >> $GITHUB_ENV + echo "HEAD_REPO=$(gh pr list $SHA --json headRepository -q '.[0].headRepository.name' || echo '')" >> $GITHUB_ENV + echo "HEAD_OWNER=$(gh pr list $SHA --json headRepositoryOwner -q '.[0].headRepositoryOwner.login' || echo '')" >> $GITHUB_ENV env: GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} + SHA: ${{ github.event.workflow_run.head_sha }} - name: Checkout mathlib4 repository if: steps.pr-info.outputs.pullRequestNumber != '' && steps.pr-info.outputs.targetBranch == 'main' From 589385a68ce55a8de18431984010cfe2f6af7878 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 28 Oct 2024 13:19:34 -0400 Subject: [PATCH 139/177] fix: github script typo (#1016) --- .github/workflows/test_mathlib.yml | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index 5347db224b..57b7aebe09 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -20,11 +20,11 @@ jobs: - name: Get PR info id: pr-info run: | - echo "pullRequestNumber=$(gh pr list $SHA --json number -q '.[0].number' || echo '')" >> $GITHUB_OUTPUT - echo "targetBranch=$(gh pr list $SHA --json baseRefName -q '.[0].baseRefName' || echo '')" >> $GITHUB_OUTPUT - echo "HEAD_BRANCH=$(gh pr list $SHA --json headRefName -q '.[0].headRefName' || echo '')" >> $GITHUB_ENV - echo "HEAD_REPO=$(gh pr list $SHA --json headRepository -q '.[0].headRepository.name' || echo '')" >> $GITHUB_ENV - echo "HEAD_OWNER=$(gh pr list $SHA --json headRepositoryOwner -q '.[0].headRepositoryOwner.login' || echo '')" >> $GITHUB_ENV + echo "pullRequestNumber=$(gh pr list --search $SHA --json number -q '.[0].number' || echo '')" >> $GITHUB_OUTPUT + echo "targetBranch=$(gh pr list --search $SHA --json baseRefName -q '.[0].baseRefName' || echo '')" >> $GITHUB_OUTPUT + echo "HEAD_BRANCH=$(gh pr list --search $SHA --json headRefName -q '.[0].headRefName' || echo '')" >> $GITHUB_ENV + echo "HEAD_REPO=$(gh pr list --search $SHA --json headRepository -q '.[0].headRepository.name' || echo '')" >> $GITHUB_ENV + echo "HEAD_OWNER=$(gh pr list --search $SHA --json headRepositoryOwner -q '.[0].headRepositoryOwner.login' || echo '')" >> $GITHUB_ENV env: GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} SHA: ${{ github.event.workflow_run.head_sha }} From eb6c831c05bc6ca6d37454ca97531a9b780dfcb5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 28 Oct 2024 13:39:35 -0400 Subject: [PATCH 140/177] fix: github script fix head repo (#1017) --- .github/workflows/test_mathlib.yml | 10 +++------- 1 file changed, 3 insertions(+), 7 deletions(-) diff --git a/.github/workflows/test_mathlib.yml b/.github/workflows/test_mathlib.yml index 57b7aebe09..3bb9c6dd71 100644 --- a/.github/workflows/test_mathlib.yml +++ b/.github/workflows/test_mathlib.yml @@ -22,9 +22,6 @@ jobs: run: | echo "pullRequestNumber=$(gh pr list --search $SHA --json number -q '.[0].number' || echo '')" >> $GITHUB_OUTPUT echo "targetBranch=$(gh pr list --search $SHA --json baseRefName -q '.[0].baseRefName' || echo '')" >> $GITHUB_OUTPUT - echo "HEAD_BRANCH=$(gh pr list --search $SHA --json headRefName -q '.[0].headRefName' || echo '')" >> $GITHUB_ENV - echo "HEAD_REPO=$(gh pr list --search $SHA --json headRepository -q '.[0].headRepository.name' || echo '')" >> $GITHUB_ENV - echo "HEAD_OWNER=$(gh pr list --search $SHA --json headRepositoryOwner -q '.[0].headRepositoryOwner.login' || echo '')" >> $GITHUB_ENV env: GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} SHA: ${{ github.event.workflow_run.head_sha }} @@ -51,14 +48,13 @@ jobs: id: check_mathlib_tag env: PR_NUMBER: ${{ steps.pr-info.outputs.pullRequestNumber }} - HEAD_OWNER: ${{ env.HEAD_OWNER }} - HEAD_REPO: ${{ env.HEAD_REPO }} - HEAD_BRANCH: ${{ env.HEAD_BRANCH }} + HEAD_REPO: ${{ github.event.workflow_run.head_repository.full_name }} + HEAD_BRANCH: ${{ github.event.workflow_run.head_branch }} run: | git config user.name "leanprover-community-mathlib4-bot" git config user.email "leanprover-community-mathlib4-bot@users.noreply.github.com" - echo "PR info: $HEAD_OWNER $HEAD_REPO $HEAD_BRANCH" + echo "PR info: $HEAD_REPO $HEAD_BRANCH" BASE=master echo "Using base tag: $BASE" From 500a529408399c44d7e1649577e3c98697f95aa4 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 31 Oct 2024 13:44:36 +1100 Subject: [PATCH 141/177] feat: fill some gaps in Vector lemmas (#1018) --- Batteries/Data/Vector/Lemmas.lean | 23 ++++++++++++++++++++--- 1 file changed, 20 insertions(+), 3 deletions(-) diff --git a/Batteries/Data/Vector/Lemmas.lean b/Batteries/Data/Vector/Lemmas.lean index 0f8caa5c45..260025344a 100644 --- a/Batteries/Data/Vector/Lemmas.lean +++ b/Batteries/Data/Vector/Lemmas.lean @@ -19,6 +19,26 @@ namespace Batteries namespace Vector +attribute [simp] Vector.size_eq + +@[simp] theorem length_toList {α n} (xs : Vector α n) : xs.toList.length = n := + xs.size_eq + +@[simp] theorem getElem_mk {data : Array α} {size : data.size = n} {i : Nat} (h : i < n) : + (Vector.mk data size)[i] = data[i] := rfl + +@[simp] theorem getElem_toArray {α n} (xs : Vector α n) (i : Nat) (h : i < xs.toArray.size) : + xs.toArray[i] = xs[i]'(by simpa using h) := by + cases xs + simp + +@[simp] theorem getElem_toList {α n} (xs : Vector α n) (i : Nat) (h : i < xs.toList.length) : + xs.toList[i] = xs[i]'(by simpa using h) := by + simp [toList] + +@[simp] theorem getElem_ofFn {α n} (f : Fin n → α) (i : Nat) (h : i < n) : + (Vector.ofFn f)[i] = f ⟨i, by simpa using h⟩ := by + simp [ofFn] /-- An `empty` vector maps to a `empty` vector. -/ @[simp] @@ -47,9 +67,6 @@ protected theorem ext {a b : Vector α n} (h : (i : Nat) → (_ : i < n) → a[i rw [a.size_eq] at hi exact h i hi -@[simp] theorem getElem_mk {data : Array α} {size : data.size = n} {i : Nat} (h : i < n) : - (Vector.mk data size)[i] = data[i] := rfl - @[simp] theorem push_mk {data : Array α} {size : data.size = n} {x : α} : (Vector.mk data size).push x = Vector.mk (data.push x) (by simp [size, Nat.succ_eq_add_one]) := rfl From 31a10a332858d6981dbcf55d54ee51680dd75f18 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 1 Nov 2024 13:51:27 +1100 Subject: [PATCH 142/177] chore: bump toolchain to v4.13.0 (#1021) --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index eff86fd63d..4f86f953fb 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.13.0-rc3 +leanprover/lean4:v4.13.0 From 76e9ebe4176d29cb9cc89c669ab9f1ce32b33c3d Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 4 Nov 2024 11:48:28 +1100 Subject: [PATCH 143/177] chore: move toolchain to v4.14.0-rc1 (#1023) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: leanprover-community-mathlib4-bot Co-authored-by: Jeremy Tan Jie Rui Co-authored-by: Joachim Breitner Co-authored-by: Kyle Miller Co-authored-by: Matthew Ballard Co-authored-by: Henrik Böving Co-authored-by: Mario Carneiro --- Batteries.lean | 2 - Batteries/Data/Array/Basic.lean | 11 +- Batteries/Data/Array/Lemmas.lean | 65 ++----- Batteries/Data/Array/Match.lean | 2 +- Batteries/Data/Array/Monadic.lean | 23 ++- Batteries/Data/Array/Pairwise.lean | 2 +- Batteries/Data/AssocList.lean | 4 +- Batteries/Data/ByteArray.lean | 4 +- Batteries/Data/HashMap/Basic.lean | 6 +- Batteries/Data/HashMap/WF.lean | 22 +-- Batteries/Data/List/Basic.lean | 70 +------- Batteries/Data/List/Lemmas.lean | 97 +---------- Batteries/Data/List/Perm.lean | 8 +- Batteries/Data/Nat/Lemmas.lean | 3 +- Batteries/Data/Range/Lemmas.lean | 13 +- Batteries/Data/String/Basic.lean | 93 ---------- Batteries/Data/String/Lemmas.lean | 6 +- Batteries/Data/Sum.lean | 2 - Batteries/Data/Sum/Basic.lean | 164 ------------------ Batteries/Data/Sum/Lemmas.lean | 250 --------------------------- Batteries/Data/UInt.lean | 18 +- Batteries/Data/UnionFind/Basic.lean | 4 +- Batteries/Data/UnionFind/Lemmas.lean | 4 +- Batteries/Data/Vector/Basic.lean | 14 +- Batteries/Data/Vector/Lemmas.lean | 2 +- Batteries/Lean/Expr.lean | 10 -- Batteries/Logic.lean | 14 -- Batteries/Tactic/Alias.lean | 15 +- Batteries/Tactic/Classical.lean | 41 ----- Batteries/Tactic/HelpCmd.lean | 2 +- Batteries/Tactic/PrintPrefix.lean | 17 +- Batteries/Util/LibraryNote.lean | 2 +- lean-toolchain | 2 +- test/array.lean | 2 +- test/case.lean | 10 -- test/classical.lean | 23 --- test/congr.lean | 6 +- test/help_cmd.lean | 96 +++++++++- test/lint_simpNF.lean | 3 - test/print_prefix.lean | 42 ++--- test/simpa.lean | 2 +- test/vector.lean | 2 +- 42 files changed, 210 insertions(+), 968 deletions(-) delete mode 100644 Batteries/Data/Sum.lean delete mode 100644 Batteries/Data/Sum/Basic.lean delete mode 100644 Batteries/Data/Sum/Lemmas.lean delete mode 100644 Batteries/Tactic/Classical.lean delete mode 100644 test/classical.lean diff --git a/Batteries.lean b/Batteries.lean index c9358a787a..b66a913263 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -32,7 +32,6 @@ import Batteries.Data.RBMap import Batteries.Data.Range import Batteries.Data.Rat import Batteries.Data.String -import Batteries.Data.Sum import Batteries.Data.UInt import Batteries.Data.UnionFind import Batteries.Data.Vector @@ -71,7 +70,6 @@ import Batteries.StdDeprecations import Batteries.Tactic.Alias import Batteries.Tactic.Basic import Batteries.Tactic.Case -import Batteries.Tactic.Classical import Batteries.Tactic.Congr import Batteries.Tactic.Exact import Batteries.Tactic.HelpCmd diff --git a/Batteries/Data/Array/Basic.lean b/Batteries/Data/Array/Basic.lean index c74b8f4750..403c42d9e2 100644 --- a/Batteries/Data/Array/Basic.lean +++ b/Batteries/Data/Array/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Arthur Paulino, Floris van Doorn, Jannis Limperg -/ import Batteries.Data.Array.Init.Lemmas +import Batteries.Tactic.Alias /-! ## Definitions on Arrays @@ -14,10 +15,6 @@ proofs about these definitions, those are contained in other files in `Batteries namespace Array -/-- Drop `none`s from a Array, and replace each remaining `some a` with `a`. -/ -def reduceOption (l : Array (Option α)) : Array α := - l.filterMap id - /-- Check whether `xs` and `ys` are equal as sets, i.e. they contain the same elements when disregarding order and duplicates. `O(n*m)`! If your element type @@ -122,11 +119,7 @@ protected def maxI [ord : Ord α] [Inhabited α] (xs : Array α) (start := 0) (stop := xs.size) : α := xs.minI (ord := ord.opposite) start stop -/-- -`O(|join L|)`. `join L` concatenates all the arrays in `L` into one array. -* `join #[#[a], #[], #[b, c], #[d, e, f]] = #[a, b, c, d, e, f]` --/ -@[inline] def join (l : Array (Array α)) : Array α := l.foldl (· ++ ·) #[] +@[deprecated (since := "2024-10-15")] alias join := flatten /-! ### Safe Nat Indexed Array functions diff --git a/Batteries/Data/Array/Lemmas.lean b/Batteries/Data/Array/Lemmas.lean index 124f8d412c..8a5544089d 100644 --- a/Batteries/Data/Array/Lemmas.lean +++ b/Batteries/Data/Array/Lemmas.lean @@ -14,19 +14,8 @@ namespace Array theorem forIn_eq_forIn_toList [Monad m] (as : Array α) (b : β) (f : α → β → m (ForInStep β)) : forIn as b f = forIn as.toList b f := by - let rec loop : ∀ {i h b j}, j + i = as.size → - Array.forIn.loop as f i h b = forIn (as.toList.drop j) b f - | 0, _, _, _, rfl => by rw [List.drop_length]; rfl - | i+1, _, _, j, ij => by - simp only [forIn.loop, Nat.add] - have j_eq : j = size as - 1 - i := by simp [← ij, ← Nat.add_assoc] - have : as.size - 1 - i < as.size := j_eq ▸ ij ▸ Nat.lt_succ_of_le (Nat.le_add_right ..) - have : as[size as - 1 - i] :: as.toList.drop (j + 1) = as.toList.drop j := by - rw [j_eq]; exact List.getElem_cons_drop _ _ this - simp only [← this, List.forIn_cons]; congr; funext x; congr; funext b - rw [loop (i := i)]; rw [← ij, Nat.succ_add]; rfl - conv => lhs; simp only [forIn, Array.forIn] - rw [loop (Nat.zero_add _)]; rfl + cases as + simp @[deprecated (since := "2024-09-09")] alias forIn_eq_forIn_data := forIn_eq_forIn_toList @[deprecated (since := "2024-08-13")] alias forIn_eq_data_forIn := forIn_eq_forIn_data @@ -99,28 +88,11 @@ theorem size_filter_le (p : α → Bool) (l : Array α) : simp only [← length_toList, toList_filter] apply List.length_filter_le -/-! ### join -/ - -@[simp] theorem toList_join {l : Array (Array α)} : l.join.toList = (l.toList.map toList).join := by - dsimp [join] - simp only [foldl_eq_foldl_toList] - generalize l.toList = l - have : ∀ a : Array α, (List.foldl ?_ a l).toList = a.toList ++ ?_ := ?_ - exact this #[] - induction l with - | nil => simp - | cons h => induction h.toList <;> simp [*] -@[deprecated (since := "2024-09-09")] alias data_join := toList_join -@[deprecated (since := "2024-08-13")] alias join_data := data_join - -theorem mem_join : ∀ {L : Array (Array α)}, a ∈ L.join ↔ ∃ l, l ∈ L ∧ a ∈ l := by - simp only [mem_def, toList_join, List.mem_join, List.mem_map] - intro l - constructor - · rintro ⟨_, ⟨s, m, rfl⟩, h⟩ - exact ⟨s, m, h⟩ - · rintro ⟨s, h₁, h₂⟩ - refine ⟨s.toList, ⟨⟨s, h₁, rfl⟩, h₂⟩⟩ +/-! ### flatten -/ + +@[deprecated (since := "2024-09-09")] alias data_join := toList_flatten +@[deprecated (since := "2024-08-13")] alias join_data := toList_flatten +@[deprecated (since := "2024-10-15")] alias mem_join := mem_flatten /-! ### indexOf? -/ @@ -154,27 +126,10 @@ where (a.eraseIdx i).size = if i < a.size then a.size-1 else a.size := by simp only [eraseIdx]; split; simp; rfl -/-! ### shrink -/ - -theorem size_shrink_loop (a : Array α) (n) : (shrink.loop n a).size = a.size - n := by - induction n generalizing a with simp only [shrink.loop, Nat.sub_zero] - | succ n ih => simp only [ih, size_pop]; omega - -@[simp] theorem size_shrink (a : Array α) (n) : (a.shrink n).size = min a.size n := by - simp [shrink, size_shrink_loop] - omega - /-! ### set -/ theorem size_set! (a : Array α) (i v) : (a.set! i v).size = a.size := by simp -/-! ### swapAt -/ - -theorem size_swapAt (a : Array α) (x i) : (a.swapAt i x).snd.size = a.size := by simp - -@[simp] theorem size_swapAt! (a : Array α) (x) (h : i < a.size) : - (a.swapAt! i x).snd.size = a.size := by simp [h] - /-! ### map -/ theorem mapM_empty [Monad m] (f : α → m β) : mapM f #[] = pure #[] := by @@ -266,14 +221,14 @@ theorem getElem_insertAt_lt (as : Array α) (i : Fin (as.size+1)) (v : α) (k) (hlt : k < i.val) {hk : k < (as.insertAt i v).size} {hk' : k < as.size} : (as.insertAt i v)[k] = as[k] := by simp only [insertAt] - rw [get_insertAt_loop_lt, get_push, dif_pos hk'] + rw [get_insertAt_loop_lt, getElem_push, dif_pos hk'] exact hlt theorem getElem_insertAt_gt (as : Array α) (i : Fin (as.size+1)) (v : α) (k) (hgt : k > i.val) {hk : k < (as.insertAt i v).size} {hk' : k - 1 < as.size} : (as.insertAt i v)[k] = as[k - 1] := by simp only [insertAt] - rw [get_insertAt_loop_gt_le, get_push, dif_pos hk'] + rw [get_insertAt_loop_gt_le, getElem_push, dif_pos hk'] exact hgt rw [size_insertAt] at hk exact Nat.le_of_lt_succ hk @@ -282,6 +237,6 @@ theorem getElem_insertAt_eq (as : Array α) (i : Fin (as.size+1)) (v : α) (k) (heq : i.val = k) {hk : k < (as.insertAt i v).size} : (as.insertAt i v)[k] = v := by simp only [insertAt] - rw [get_insertAt_loop_eq, Fin.getElem_fin, get_push_eq] + rw [get_insertAt_loop_eq, Fin.getElem_fin, getElem_push_eq] exact heq exact Nat.le_of_lt_succ i.is_lt diff --git a/Batteries/Data/Array/Match.lean b/Batteries/Data/Array/Match.lean index 46b2239f6f..b47ee90adc 100644 --- a/Batteries/Data/Array/Match.lean +++ b/Batteries/Data/Array/Match.lean @@ -52,7 +52,7 @@ termination_by k => k.val def PrefixTable.extend [BEq α] (t : PrefixTable α) (x : α) : PrefixTable α where toArray := t.toArray.push (x, t.step x ⟨t.size, Nat.lt_succ_self _⟩) valid _ := by - rw [Array.get_push] + rw [Array.getElem_push] split · exact t.valid .. · next h => exact Nat.le_trans (Nat.lt_succ.1 <| Fin.isLt ..) (Nat.not_lt.1 h) diff --git a/Batteries/Data/Array/Monadic.lean b/Batteries/Data/Array/Monadic.lean index 316813972a..934111e131 100644 --- a/Batteries/Data/Array/Monadic.lean +++ b/Batteries/Data/Array/Monadic.lean @@ -44,7 +44,7 @@ theorem SatisfiesM_mapM [Monad m] [LawfulMonad m] (as : Array α) (f : α → m · case s => intro ⟨i, hi⟩ arr ⟨ih₁, eq, ih₂⟩ refine (hs _ ih₁).map fun ⟨h₁, h₂⟩ => ⟨h₂, by simp [eq], fun j hj => ?_⟩ - simp [get_push] at hj ⊢; split; {apply ih₂} + simp [getElem_push] at hj ⊢; split; {apply ih₂} cases j; cases (Nat.le_or_eq_of_le_succ hj).resolve_left ‹_›; cases eq; exact h₁ theorem SatisfiesM_mapM' [Monad m] [LawfulMonad m] (as : Array α) (f : α → m β) @@ -125,27 +125,36 @@ theorem SatisfiesM_foldrM [Monad m] [LawfulMonad m] simp [foldrM]; split; {exact go _ h0} · next h => exact .pure (Nat.eq_zero_of_not_pos h ▸ h0) -theorem SatisfiesM_mapIdxM [Monad m] [LawfulMonad m] (as : Array α) (f : Fin as.size → α → m β) +theorem SatisfiesM_mapFinIdxM [Monad m] [LawfulMonad m] (as : Array α) (f : Fin as.size → α → m β) (motive : Nat → Prop) (h0 : motive 0) (p : Fin as.size → β → Prop) (hs : ∀ i, motive i.1 → SatisfiesM (p i · ∧ motive (i + 1)) (f i as[i])) : SatisfiesM (fun arr => motive as.size ∧ ∃ eq : arr.size = as.size, ∀ i h, p ⟨i, h⟩ arr[i]) - (Array.mapIdxM as f) := by + (Array.mapFinIdxM as f) := by let rec go {bs i j h} (h₁ : j = bs.size) (h₂ : ∀ i h h', p ⟨i, h⟩ bs[i]) (hm : motive j) : SatisfiesM (fun arr => motive as.size ∧ ∃ eq : arr.size = as.size, ∀ i h, p ⟨i, h⟩ arr[i]) - (Array.mapIdxM.map as f i j h bs) := by - induction i generalizing j bs with simp [mapIdxM.map] + (Array.mapFinIdxM.map as f i j h bs) := by + induction i generalizing j bs with simp [mapFinIdxM.map] | zero => have := (Nat.zero_add _).symm.trans h exact .pure ⟨this ▸ hm, h₁ ▸ this, fun _ _ => h₂ ..⟩ | succ i ih => refine (hs _ (by exact hm)).bind fun b hb => ih (by simp [h₁]) (fun i hi hi' => ?_) hb.2 - simp at hi'; simp [get_push]; split + simp at hi'; simp [getElem_push]; split · next h => exact h₂ _ _ h · next h => cases h₁.symm ▸ (Nat.le_or_eq_of_le_succ hi').resolve_left h; exact hb.1 - simp [mapIdxM]; exact go rfl nofun h0 + simp [mapFinIdxM]; exact go rfl nofun h0 + +theorem SatisfiesM_mapIdxM [Monad m] [LawfulMonad m] (as : Array α) (f : Nat → α → m β) + (motive : Nat → Prop) (h0 : motive 0) + (p : Fin as.size → β → Prop) + (hs : ∀ i, motive i.1 → SatisfiesM (p i · ∧ motive (i + 1)) (f i as[i])) : + SatisfiesM + (fun arr => motive as.size ∧ ∃ eq : arr.size = as.size, ∀ i h, p ⟨i, h⟩ arr[i]) + (Array.mapIdxM as f) := + SatisfiesM_mapFinIdxM as (fun i => f i) motive h0 p hs theorem size_modifyM [Monad m] [LawfulMonad m] (a : Array α) (i : Nat) (f : α → m α) : SatisfiesM (·.size = a.size) (a.modifyM i f) := by diff --git a/Batteries/Data/Array/Pairwise.lean b/Batteries/Data/Array/Pairwise.lean index 9e61e642ed..82ace9609f 100644 --- a/Batteries/Data/Array/Pairwise.lean +++ b/Batteries/Data/Array/Pairwise.lean @@ -21,7 +21,7 @@ def Pairwise (R : α → α → Prop) (as : Array α) : Prop := as.toList.Pairwi theorem pairwise_iff_get {as : Array α} : as.Pairwise R ↔ ∀ (i j : Fin as.size), i < j → R (as.get i) (as.get j) := by - unfold Pairwise; simp [List.pairwise_iff_get, getElem_fin_eq_toList_get] + unfold Pairwise; simp [List.pairwise_iff_get, getElem_fin_eq_getElem_toList] theorem pairwise_iff_getElem {as : Array α} : as.Pairwise R ↔ ∀ (i j : Nat) (_ : i < as.size) (_ : j < as.size), i < j → R as[i] as[j] := by diff --git a/Batteries/Data/AssocList.lean b/Batteries/Data/AssocList.lean index 1e94b1d2f6..eba8927cef 100644 --- a/Batteries/Data/AssocList.lean +++ b/Batteries/Data/AssocList.lean @@ -244,8 +244,8 @@ instance : ForIn m (AssocList α β) (α × β) where @[simp] theorem forIn_eq [Monad m] (l : AssocList α β) (init : δ) (f : (α × β) → δ → m (ForInStep δ)) : forIn l init f = forIn l.toList init f := by - simp [forIn, List.forIn] - induction l generalizing init <;> simp [AssocList.forIn, List.forIn.loop] + simp only [forIn] + induction l generalizing init <;> simp [AssocList.forIn] congr; funext a; split <;> simp [*] /-- Split the list into head and tail, if possible. -/ diff --git a/Batteries/Data/ByteArray.lean b/Batteries/Data/ByteArray.lean index c10cf6ba65..6a95a241aa 100644 --- a/Batteries/Data/ByteArray.lean +++ b/Batteries/Data/ByteArray.lean @@ -36,11 +36,11 @@ theorem getElem_eq_data_getElem (a : ByteArray) (h : i < a.size) : a[i] = a.data Array.size_push .. @[simp] theorem get_push_eq (a : ByteArray) (x : UInt8) : (a.push x)[a.size] = x := - Array.get_push_eq .. + Array.getElem_push_eq .. theorem get_push_lt (a : ByteArray) (x : UInt8) (i : Nat) (h : i < a.size) : (a.push x)[i]'(size_push .. ▸ Nat.lt_succ_of_lt h) = a[i] := - Array.get_push_lt .. + Array.getElem_push_lt .. /-! ### set -/ diff --git a/Batteries/Data/HashMap/Basic.lean b/Batteries/Data/HashMap/Basic.lean index 98905c3343..3fd1d3350d 100644 --- a/Batteries/Data/HashMap/Basic.lean +++ b/Batteries/Data/HashMap/Basic.lean @@ -37,7 +37,7 @@ def update (data : Buckets α β) (i : USize) The number of elements in the bucket array. Note: this is marked `noncomputable` because it is only intended for specification. -/ -noncomputable def size (data : Buckets α β) : Nat := .sum (data.1.toList.map (·.toList.length)) +noncomputable def size (data : Buckets α β) : Nat := (data.1.toList.map (·.toList.length)).sum @[simp] theorem update_size (self : Buckets α β) (i d h) : (self.update i d h).1.size = self.1.size := Array.size_uset .. @@ -56,7 +56,7 @@ structure WF [BEq α] [Hashable α] (buckets : Buckets α β) : Prop where bucket.toList.Pairwise fun a b => ¬(a.1 == b.1) /-- Every element in a bucket should hash to its location. -/ hash_self (i : Nat) (h : i < buckets.1.size) : - buckets.1[i].All fun k _ => ((hash k).toUSize % buckets.1.size).toNat = i + buckets.1[i].All fun k _ => ((hash k).toUSize % USize.ofNat buckets.1.size).toNat = i end Buckets end Imp @@ -93,7 +93,7 @@ def empty (capacity := 0) : Imp α β := /-- Calculates the bucket index from a hash value `u`. -/ def mkIdx {n : Nat} (h : 0 < n) (u : USize) : {u : USize // u.toNat < n} := - ⟨u % n, USize.modn_lt _ h⟩ + ⟨u % USize.ofNat n, USize.toNat_mod_lt _ h⟩ /-- Inserts a key-value pair into the bucket array. This function assumes that the data is not diff --git a/Batteries/Data/HashMap/WF.lean b/Batteries/Data/HashMap/WF.lean index c95e3ad133..aa33711633 100644 --- a/Batteries/Data/HashMap/WF.lean +++ b/Batteries/Data/HashMap/WF.lean @@ -36,7 +36,7 @@ theorem update_update (self : Buckets α β) (i d d' h h') : rw [Array.set_set] theorem size_eq (data : Buckets α β) : - size data = .sum (data.1.toList.map (·.toList.length)) := rfl + size data = (data.1.toList.map (·.toList.length)).sum := rfl theorem mk_size (h) : (mk n h : Buckets α β).size = 0 := by simp only [mk, mkArray, size_eq]; clear h @@ -52,8 +52,8 @@ theorem WF.update [BEq α] [Hashable α] {buckets : Buckets α β} {i d h} (H : (h₁ : ∀ [PartialEquivBEq α] [LawfulHashable α], (buckets.1[i].toList.Pairwise fun a b => ¬(a.1 == b.1)) → d.toList.Pairwise fun a b => ¬(a.1 == b.1)) - (h₂ : (buckets.1[i].All fun k _ => ((hash k).toUSize % buckets.1.size).toNat = i.toNat) → - d.All fun k _ => ((hash k).toUSize % buckets.1.size).toNat = i.toNat) : + (h₂ : (buckets.1[i].All fun k _ => ((hash k).toUSize % .ofNat buckets.1.size).toNat = i.toNat) → + d.All fun k _ => ((hash k).toUSize % USize.ofNat buckets.1.size).toNat = i.toNat) : (buckets.update i d h).WF := by refine ⟨fun l hl => ?_, fun i hi p hp => ?_⟩ · exact match List.mem_or_eq_of_mem_set hl with @@ -93,7 +93,7 @@ theorem expand_size [Hashable α] {buckets : Buckets α β} : where go (i source) (target : Buckets α β) (hs : ∀ j < i, source.toList[j]?.getD .nil = .nil) : (expand.go i source target).size = - .sum (source.toList.map (·.toList.length)) + target.size := by + (source.toList.map (·.toList.length)).sum + target.size := by unfold expand.go; split · next H => refine (go (i+1) _ _ fun j hj => ?a).trans ?b @@ -107,14 +107,14 @@ where refine have ⟨l₁, l₂, h₁, _, eq⟩ := List.exists_of_set H; eq ▸ ?_ rw [h₁] simp only [Buckets.size_eq, List.map_append, List.map_cons, AssocList.toList, - List.length_nil, Nat.sum_append, Nat.sum_cons, Nat.zero_add, Array.length_toList] + List.length_nil, Nat.sum_append, List.sum_cons, Nat.zero_add, Array.length_toList] rw [Nat.add_assoc, Nat.add_assoc, Nat.add_assoc]; congr 1 (conv => rhs; rw [Nat.add_left_comm]); congr 1 rw [← Array.getElem_eq_getElem_toList] have := @reinsertAux_size α β _; simp [Buckets.size] at this induction source[i].toList generalizing target <;> simp [*, Nat.succ_add]; rfl · next H => - rw [(_ : Nat.sum _ = 0), Nat.zero_add] + rw [(_ : List.sum _ = 0), Nat.zero_add] rw [← (_ : source.toList.map (fun _ => .nil) = source.toList)] · simp only [List.map_map] induction source.toList <;> simp [*] @@ -163,9 +163,9 @@ where (hs₁ : ∀ [LawfulHashable α] [PartialEquivBEq α], ∀ bucket ∈ source.toList, bucket.toList.Pairwise fun a b => ¬(a.1 == b.1)) (hs₂ : ∀ (j : Nat) (h : j < source.size), - source[j].All fun k _ => ((hash k).toUSize % source.size).toNat = j) + source[j].All fun k _ => ((hash k).toUSize % USize.ofNat source.size).toNat = j) {target : Buckets α β} (ht : target.WF ∧ ∀ bucket ∈ target.1.toList, - bucket.All fun k _ => ((hash k).toUSize % source.size).toNat < i) : + bucket.All fun k _ => ((hash k).toUSize % USize.ofNat source.size).toNat < i) : (expand.go i source target).WF := by unfold expand.go; split · next H => @@ -178,7 +178,7 @@ where split · nofun · exact hs₂ _ (by simp_all) - · let rank (k : α) := ((hash k).toUSize % source.size).toNat + · let rank (k : α) := ((hash k).toUSize % USize.ofNat source.size).toNat have := expand_WF.foldl rank ?_ (hs₂ _ H) ht.1 (fun _ h₁ _ h₂ => ?_) · simp only [Array.get_eq_getElem, AssocList.foldl_eq, Array.size_set] exact ⟨this.1, fun _ h₁ _ h₂ => Nat.lt_succ_of_le (this.2 _ h₁ _ h₂)⟩ @@ -267,7 +267,7 @@ theorem erase_size [BEq α] [Hashable α] {m : Imp α β} {k} simp only [h, Buckets.size] refine have ⟨_, _, h₁, _, eq⟩ := Buckets.exists_of_update ..; eq ▸ ?_ simp only [h₁, Array.length_toList, Array.ugetElem_eq_getElem, List.map_append, List.map_cons, - Nat.sum_append, Nat.sum_cons, AssocList.toList_erase] + Nat.sum_append, List.sum_cons, AssocList.toList_erase] rw [(_ : List.length _ = _ + 1), Nat.add_right_comm]; {rfl} clear h₁ eq simp only [AssocList.contains_eq, List.any_eq_true] at H @@ -347,7 +347,7 @@ theorem WF.filterMap {α β γ} {f : α → β → Option γ} [BEq α] [Hashable let M := StateT (ULift Nat) Id have H2 (l : List (AssocList α β)) n : l.mapM (m := M) (filterMap.go f .nil) n = - (l.map g, ⟨n.1 + .sum ((l.map g).map (·.toList.length))⟩) := by + (l.map g, ⟨n.1 + ((l.map g).map (·.toList.length)).sum⟩) := by induction l generalizing n with | nil => rfl | cons l L IH => simp [bind, StateT.bind, IH, H1, Nat.add_assoc, g]; rfl diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index c4c262a016..3e2460f18c 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -32,17 +32,6 @@ open Option Nat | [] => none | a :: l => some (a, l) -/-- -Given a function `f : Nat → α → β` and `as : list α`, `as = [a₀, a₁, ...]`, returns the list -`[f 0 a₀, f 1 a₁, ...]`. --/ -@[inline] def mapIdx (f : Nat → α → β) (as : List α) : List β := go as #[] where - /-- Auxiliary for `mapIdx`: - `mapIdx.go [a₀, a₁, ...] acc = acc.toList ++ [f acc.size a₀, f (acc.size + 1) a₁, ...]` -/ - @[specialize] go : List α → Array β → List β - | [], acc => acc.toList - | a :: as, acc => go as (acc.push (f acc.size a)) - /-- Monadic variant of `mapIdx`. -/ @[inline] def mapIdxM {m : Type v → Type w} [Monad m] (as : List α) (f : Nat → α → m β) : m (List β) := go as #[] where @@ -163,54 +152,9 @@ Split a list at every occurrence of a separator element. The separators are not -/ @[inline] def splitOn [BEq α] (a : α) (as : List α) : List (List α) := as.splitOnP (· == a) -/-- -Apply a function to the nth tail of `l`. Returns the input without -using `f` if the index is larger than the length of the List. -``` -modifyTailIdx f 2 [a, b, c] = [a, b] ++ f [c] -``` --/ -@[simp] def modifyTailIdx (f : List α → List α) : Nat → List α → List α - | 0, l => f l - | _+1, [] => [] - | n+1, a :: l => a :: modifyTailIdx f n l - @[deprecated (since := "2024-10-21")] alias modifyNthTail := modifyTailIdx - -/-- Apply `f` to the head of the list, if it exists. -/ -@[inline] def modifyHead (f : α → α) : List α → List α - | [] => [] - | a :: l => f a :: l - -@[simp] theorem modifyHead_nil (f : α → α) : [].modifyHead f = [] := by rw [modifyHead] - -@[simp] theorem modifyHead_cons (a : α) (l : List α) (f : α → α) : - (a :: l).modifyHead f = f a :: l := by rw [modifyHead] - -/-- -Apply `f` to the nth element of the list, if it exists, replacing that element with the result. --/ -def modify (f : α → α) : Nat → List α → List α := - modifyTailIdx (modifyHead f) - @[deprecated (since := "2024-10-21")] alias modifyNth := modify -/-- Tail-recursive version of `modify`. -/ -def modifyTR (f : α → α) (n : Nat) (l : List α) : List α := go l n #[] where - /-- Auxiliary for `modifyTR`: `modifyTR.go f l n acc = acc.toList ++ modify f n l`. -/ - go : List α → Nat → Array α → List α - | [], _, acc => acc.toList - | a :: l, 0, acc => acc.toListAppend (f a :: l) - | a :: l, n+1, acc => go l n (acc.push a) - -theorem modifyTR_go_eq : ∀ l n, modifyTR.go f l n acc = acc.toList ++ modify f n l - | [], n => by cases n <;> simp [modifyTR.go, modify] - | a :: l, 0 => by simp [modifyTR.go, modify] - | a :: l, n+1 => by simp [modifyTR.go, modify, modifyTR_go_eq l] - -@[csimp] theorem modify_eq_modifyTR : @modify = @modifyTR := by - funext α f n l; simp [modifyTR, modifyTR_go_eq] - /-- Apply `f` to the last element of `l`, if it exists. -/ @[inline] def modifyLast (f : α → α) (l : List α) : List α := go l #[] where /-- Auxiliary for `modifyLast`: `modifyLast.go f l acc = acc.toList ++ modifyLast f l`. -/ @@ -424,7 +368,7 @@ sublists [1, 2, 3] = [[], [1], [2], [1, 2], [3], [1, 3], [2, 3], [1, 2, 3]] ``` -/ def sublists (l : List α) : List (List α) := - l.foldr (fun a acc => acc.bind fun x => [x, a :: x]) [[]] + l.foldr (fun a acc => acc.flatMap fun x => [x, a :: x]) [[]] /-- A version of `List.sublists` that has faster runtime performance but worse kernel performance -/ def sublistsFast (l : List α) : List (List α) := @@ -522,7 +466,7 @@ of `[L₁, L₂, ..., Lₙ]` is a list whose first element comes from -/ @[simp] def sections : List (List α) → List (List α) | [] => [[]] - | l :: L => (sections L).bind fun s => l.map fun a => a :: s + | l :: L => (sections L).flatMap fun s => l.map fun a => a :: s /-- Optimized version of `sections`. -/ def sectionsTR (L : List (List α)) : List (List α) := @@ -547,7 +491,7 @@ theorem sections_eq_nil_of_isEmpty : ∀ {L}, L.any isEmpty → @sections α L = cases e : L.any isEmpty <;> simp [sections_eq_nil_of_isEmpty, *] clear e; induction L with | nil => rfl | cons l L IH => ?_ simp [IH, sectionsTR.go] - rw [Array.foldl_eq_foldl_toList, Array.foldl_toList_eq_bind]; rfl + rw [Array.foldl_eq_foldl_toList, Array.foldl_toList_eq_flatMap]; rfl intros; apply Array.foldl_toList_eq_map /-- @@ -577,7 +521,7 @@ def revzip (l : List α) : List (α × α) := zip l l.reverse product [1, 2] [5, 6] = [(1, 5), (1, 6), (2, 5), (2, 6)] ``` -/ -def product (l₁ : List α) (l₂ : List β) : List (α × β) := l₁.bind fun a => l₂.map (Prod.mk a) +def product (l₁ : List α) (l₂ : List β) : List (α × β) := l₁.flatMap fun a => l₂.map (Prod.mk a) /-- Optimized version of `product`. -/ def productTR (l₁ : List α) (l₂ : List β) : List (α × β) := @@ -585,7 +529,7 @@ def productTR (l₁ : List α) (l₂ : List β) : List (α × β) := @[csimp] theorem product_eq_productTR : @product = @productTR := by funext α β l₁ l₂; simp [product, productTR] - rw [Array.foldl_toList_eq_bind]; rfl + rw [Array.foldl_toList_eq_flatMap]; rfl intros; apply Array.foldl_toList_eq_map /-- `sigma l₁ l₂` is the list of dependent pairs `(a, b)` where `a ∈ l₁` and `b ∈ l₂ a`. @@ -593,7 +537,7 @@ def productTR (l₁ : List α) (l₂ : List β) : List (α × β) := sigma [1, 2] (λ_, [(5 : Nat), 6]) = [(1, 5), (1, 6), (2, 5), (2, 6)] ``` -/ protected def sigma {σ : α → Type _} (l₁ : List α) (l₂ : ∀ a, List (σ a)) : List (Σ a, σ a) := - l₁.bind fun a => (l₂ a).map (Sigma.mk a) + l₁.flatMap fun a => (l₂ a).map (Sigma.mk a) /-- Optimized version of `sigma`. -/ def sigmaTR {σ : α → Type _} (l₁ : List α) (l₂ : ∀ a, List (σ a)) : List (Σ a, σ a) := @@ -601,7 +545,7 @@ def sigmaTR {σ : α → Type _} (l₁ : List α) (l₂ : ∀ a, List (σ a)) : @[csimp] theorem sigma_eq_sigmaTR : @List.sigma = @sigmaTR := by funext α β l₁ l₂; simp [List.sigma, sigmaTR] - rw [Array.foldl_toList_eq_bind]; rfl + rw [Array.foldl_toList_eq_flatMap]; rfl intros; apply Array.foldl_toList_eq_map /-- diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 5179c81ba0..89f1de368d 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -36,89 +36,6 @@ theorem dropLast_eq_eraseIdx {xs : List α} {i : Nat} (last_idx : i + 1 = xs.len exact ih last_idx exact fun _ => nomatch xs -/-! ### modifyHead -/ - -@[simp] theorem modifyHead_modifyHead (l : List α) (f g : α → α) : - (l.modifyHead f).modifyHead g = l.modifyHead (g ∘ f) := by cases l <;> simp [modifyHead] - -/-! ### modify -/ - -@[simp] theorem modify_nil (f : α → α) (n) : [].modify f n = [] := by cases n <;> rfl - -@[simp] theorem modify_zero_cons (f : α → α) (a : α) (l : List α) : - (a :: l).modify f 0 = f a :: l := rfl - -@[simp] theorem modify_succ_cons (f : α → α) (a : α) (l : List α) (n) : - (a :: l).modify f (n + 1) = a :: l.modify f n := by rfl - -theorem modifyTailIdx_id : ∀ n (l : List α), l.modifyTailIdx id n = l - | 0, _ => rfl - | _+1, [] => rfl - | n+1, a :: l => congrArg (cons a) (modifyTailIdx_id n l) - -theorem eraseIdx_eq_modifyTailIdx : ∀ n (l : List α), eraseIdx l n = modifyTailIdx tail n l - | 0, l => by cases l <;> rfl - | _+1, [] => rfl - | _+1, _ :: _ => congrArg (cons _) (eraseIdx_eq_modifyTailIdx _ _) - -theorem getElem?_modify (f : α → α) : - ∀ n (l : List α) m, (modify f n l)[m]? = (fun a => if n = m then f a else a) <$> l[m]? - | n, l, 0 => by cases l <;> cases n <;> simp - | n, [], _+1 => by cases n <;> rfl - | 0, _ :: l, m+1 => by cases h : l[m]? <;> simp [h, modify, m.succ_ne_zero.symm] - | n+1, a :: l, m+1 => by - simp only [modify_succ_cons, getElem?_cons_succ, Nat.reduceEqDiff, Option.map_eq_map] - refine (getElem?_modify f n l m).trans ?_ - cases h' : l[m]? <;> by_cases h : n = m <;> - simp [h, if_pos, if_neg, Option.map, mt Nat.succ.inj, not_false_iff, h'] - -theorem length_modifyTailIdx (f : List α → List α) (H : ∀ l, length (f l) = length l) : - ∀ n l, length (modifyTailIdx f n l) = length l - | 0, _ => H _ - | _+1, [] => rfl - | _+1, _ :: _ => congrArg (·+1) (length_modifyTailIdx _ H _ _) - -theorem modifyTailIdx_add (f : List α → List α) (n) (l₁ l₂ : List α) : - modifyTailIdx f (l₁.length + n) (l₁ ++ l₂) = l₁ ++ modifyTailIdx f n l₂ := by - induction l₁ <;> simp [*, Nat.succ_add] - -theorem exists_of_modifyTailIdx (f : List α → List α) {n} {l : List α} (h : n ≤ l.length) : - ∃ l₁ l₂, l = l₁ ++ l₂ ∧ l₁.length = n ∧ modifyTailIdx f n l = l₁ ++ f l₂ := - have ⟨_, _, eq, hl⟩ : ∃ l₁ l₂, l = l₁ ++ l₂ ∧ l₁.length = n := - ⟨_, _, (take_append_drop n l).symm, length_take_of_le h⟩ - ⟨_, _, eq, hl, hl ▸ eq ▸ modifyTailIdx_add (n := 0) ..⟩ - -@[simp] theorem length_modify (f : α → α) : ∀ n l, length (modify f n l) = length l := - length_modifyTailIdx _ fun l => by cases l <;> rfl - -@[simp] theorem getElem?_modify_eq (f : α → α) (n) (l : List α) : - (modify f n l)[n]? = f <$> l[n]? := by - simp only [getElem?_modify, if_pos] - -@[simp] theorem getElem?_modify_ne (f : α → α) {m n} (l : List α) (h : m ≠ n) : - (modify f m l)[n]? = l[n]? := by - simp only [getElem?_modify, if_neg h, id_map'] - -theorem exists_of_modify (f : α → α) {n} {l : List α} (h : n < l.length) : - ∃ l₁ a l₂, l = l₁ ++ a :: l₂ ∧ l₁.length = n ∧ modify f n l = l₁ ++ f a :: l₂ := - match exists_of_modifyTailIdx _ (Nat.le_of_lt h) with - | ⟨_, _::_, eq, hl, H⟩ => ⟨_, _, _, eq, hl, H⟩ - | ⟨_, [], eq, hl, _⟩ => nomatch Nat.ne_of_gt h (eq ▸ append_nil _ ▸ hl) - -theorem modifyTailIdx_eq_take_drop (f : List α → List α) (H : f [] = []) : - ∀ n l, modifyTailIdx f n l = take n l ++ f (drop n l) - | 0, _ => rfl - | _ + 1, [] => H.symm - | n + 1, b :: l => congrArg (cons b) (modifyTailIdx_eq_take_drop f H n l) - -theorem modify_eq_take_drop (f : α → α) : - ∀ n l, modify f n l = take n l ++ modifyHead f (drop n l) := - modifyTailIdx_eq_take_drop _ rfl - -theorem modify_eq_take_cons_drop (f : α → α) {n l} (h : n < length l) : - modify f n l = take n l ++ f l[n] :: drop (n + 1) l := by - rw [modify_eq_take_drop, drop_eq_getElem_cons h]; rfl - /-! ### set -/ theorem set_eq_modify (a : α) : ∀ n (l : List α), set l n a = modify (fun _ => a) n l @@ -128,7 +45,7 @@ theorem set_eq_modify (a : α) : ∀ n (l : List α), set l n a = modify (fun _ theorem set_eq_take_cons_drop (a : α) {n l} (h : n < length l) : set l n a = take n l ++ a :: drop (n + 1) l := by - rw [set_eq_modify, modify_eq_take_cons_drop _ h] + rw [set_eq_modify, modify_eq_take_cons_drop h] theorem modify_eq_set_get? (f : α → α) : ∀ n (l : List α), l.modify f n = ((fun a => l.set n (f a)) <$> l.get? n).getD l @@ -328,7 +245,7 @@ theorem inter_def [BEq α] (l₁ l₂ : List α) : l₁ ∩ l₂ = filter (elem theorem pair_mem_product {xs : List α} {ys : List β} {x : α} {y : β} : (x, y) ∈ product xs ys ↔ x ∈ xs ∧ y ∈ ys := by simp only [product, and_imp, mem_map, Prod.mk.injEq, - exists_eq_right_right, mem_bind, iff_self] + exists_eq_right_right, mem_flatMap, iff_self] /-! ### monadic operations -/ @@ -591,16 +508,6 @@ theorem insertP_loop (a : α) (l r : List α) : induction l with simp [insertP, insertP.loop, cond] | cons _ _ ih => split <;> simp [insertP_loop, ih] -/-! ### foldlM and foldrM -/ - -theorem foldlM_map [Monad m] (f : β₁ → β₂) (g : α → β₂ → m α) (l : List β₁) (init : α) : - (l.map f).foldlM g init = l.foldlM (fun x y => g x (f y)) init := by - induction l generalizing g init <;> simp [*] - -theorem foldrM_map [Monad m] [LawfulMonad m] (f : β₁ → β₂) (g : β₂ → α → m α) (l : List β₁) - (init : α) : (l.map f).foldrM g init = l.foldrM (fun x y => g (f x) y) init := by - induction l generalizing g init <;> simp [*] - /-! ### deprecations -/ @[deprecated (since := "2024-08-15")] alias isEmpty_iff_eq_nil := isEmpty_iff diff --git a/Batteries/Data/List/Perm.lean b/Batteries/Data/List/Perm.lean index 2196d88367..6eb06d1d9e 100644 --- a/Batteries/Data/List/Perm.lean +++ b/Batteries/Data/List/Perm.lean @@ -316,10 +316,12 @@ theorem Perm.inter {l₁ l₂ t₁ t₂ : List α} (p₁ : l₁ ~ l₂) (p₂ : end DecidableEq -theorem Perm.join_congr : - ∀ {l₁ l₂ : List (List α)} (_ : List.Forall₂ (· ~ ·) l₁ l₂), l₁.join ~ l₂.join +theorem Perm.flatten_congr : + ∀ {l₁ l₂ : List (List α)} (_ : List.Forall₂ (· ~ ·) l₁ l₂), l₁.flatten ~ l₂.flatten | _, _, .nil => .rfl - | _ :: _, _ :: _, .cons h₁ h₂ => h₁.append (Perm.join_congr h₂) + | _ :: _, _ :: _, .cons h₁ h₂ => h₁.append (Perm.flatten_congr h₂) + +@[deprecated (since := "2024-10-15")] alias Perm.join_congr := Perm.flatten_congr theorem perm_insertP (p : α → Bool) (a l) : insertP p a l ~ a :: l := by induction l with simp [insertP, insertP.loop, cond] diff --git a/Batteries/Data/Nat/Lemmas.lean b/Batteries/Data/Nat/Lemmas.lean index 7fa93705f7..a97c558c44 100644 --- a/Batteries/Data/Nat/Lemmas.lean +++ b/Batteries/Data/Nat/Lemmas.lean @@ -159,5 +159,6 @@ protected def sum_trichotomy (a b : Nat) : a < b ⊕' a = b ⊕' b < a := /-! ### sum -/ -@[simp] theorem sum_append : Nat.sum (l₁ ++ l₂) = Nat.sum l₁ + Nat.sum l₂ := by + +@[simp] theorem sum_append {l₁ l₂ : List Nat}: (l₁ ++ l₂).sum = l₁.sum + l₂.sum := by induction l₁ <;> simp [*, Nat.add_assoc] diff --git a/Batteries/Data/Range/Lemmas.lean b/Batteries/Data/Range/Lemmas.lean index d71082848f..a075e0b6b4 100644 --- a/Batteries/Data/Range/Lemmas.lean +++ b/Batteries/Data/Range/Lemmas.lean @@ -69,7 +69,7 @@ theorem forIn'_eq_forIn_range' [Monad m] (r : Std.Range) suffices ∀ fuel l i hle H, l ≤ fuel → (∀ j, stop ≤ i + step * j ↔ l ≤ j) → ∀ init, forIn'.loop start stop step f fuel i hle init = - List.forIn ((List.range' i l step).pmap Subtype.mk H) init f' by + forIn ((List.range' i l step).pmap Subtype.mk H) init f' by refine this _ _ _ _ _ ((numElems_le_iff hstep).2 (Nat.le_trans ?_ (Nat.le_add_left ..))) (fun _ => (numElems_le_iff hstep).symm) _ @@ -85,7 +85,7 @@ theorem forIn'_eq_forIn_range' [Monad m] (r : Std.Range) (List.forall_mem_cons.1 H).2 (Nat.le_of_succ_le_succ h1) fun i => by rw [Nat.add_right_comm, Nat.add_assoc, ← Nat.mul_succ, h2, Nat.succ_le_succ_iff] have := h2 0; simp at this - rw [forIn'.loop]; simp [List.forIn, this, ih]; rfl + rw [forIn'.loop]; simp [this, ih]; rfl else simp [List.range', h, numElems_stop_le_start ⟨start, stop, step⟩ (Nat.not_lt.1 h), L] cases stop <;> unfold forIn'.loop <;> simp [List.forIn', h] @@ -94,13 +94,6 @@ theorem forIn_eq_forIn_range' [Monad m] (r : Std.Range) (init : β) (f : Nat → β → m (ForInStep β)) : forIn r init f = forIn (List.range' r.start r.numElems r.step) init f := by refine Eq.trans ?_ <| (forIn'_eq_forIn_range' r init (fun x _ => f x)).trans ?_ - · simp [forIn, forIn', Range.forIn, Range.forIn'] - suffices ∀ fuel i hl b, forIn'.loop r.start r.stop r.step (fun x _ => f x) fuel i hl b = - forIn.loop f fuel i r.stop r.step b from (this _ ..).symm - intro fuel; induction fuel <;> intro i hl b <;> - unfold forIn.loop forIn'.loop <;> simp [*] - split - · simp [if_neg (Nat.not_le.2 ‹_›)] - · simp [if_pos (Nat.not_lt.1 ‹_›)] + · simp [forIn, forIn'] · suffices ∀ L H, forIn (List.pmap Subtype.mk L H) init (f ·.1) = forIn L init f from this _ .. intro L; induction L generalizing init <;> intro H <;> simp [*] diff --git a/Batteries/Data/String/Basic.lean b/Batteries/Data/String/Basic.lean index 6ff6ac98bf..aad3484c05 100644 --- a/Batteries/Data/String/Basic.lean +++ b/Batteries/Data/String/Basic.lean @@ -8,99 +8,6 @@ instance : Coe String Substring := ⟨String.toSubstring⟩ namespace String -protected theorem Pos.ne_zero_of_lt : {a b : Pos} → a < b → b ≠ 0 - | _, _, hlt, rfl => Nat.not_lt_zero _ hlt - -end String - -namespace Substring - -/-- -Returns the longest common prefix of two substrings. -The returned substring will use the same underlying string as `s`. --/ -def commonPrefix (s t : Substring) : Substring := - { s with stopPos := loop s.startPos t.startPos } -where - /-- Returns the ending position of the common prefix, working up from `spos, tpos`. -/ - loop spos tpos := - if h : spos < s.stopPos ∧ tpos < t.stopPos then - if s.str.get spos == t.str.get tpos then - have := Nat.sub_lt_sub_left h.1 (s.str.lt_next spos) - loop (s.str.next spos) (t.str.next tpos) - else - spos - else - spos - termination_by s.stopPos.byteIdx - spos.byteIdx - -/-- -Returns the longest common suffix of two substrings. -The returned substring will use the same underlying string as `s`. --/ -def commonSuffix (s t : Substring) : Substring := - { s with startPos := loop s.stopPos t.stopPos } -where - /-- Returns the starting position of the common prefix, working down from `spos, tpos`. -/ - loop spos tpos := - if h : s.startPos < spos ∧ t.startPos < tpos then - let spos' := s.str.prev spos - let tpos' := t.str.prev tpos - if s.str.get spos' == t.str.get tpos' then - have : spos' < spos := s.str.prev_lt_of_pos spos (String.Pos.ne_zero_of_lt h.1) - loop spos' tpos' - else - spos - else - spos - termination_by spos.byteIdx - -/-- -If `pre` is a prefix of `s`, i.e. `s = pre ++ t`, returns the remainder `t`. --/ -def dropPrefix? (s : Substring) (pre : Substring) : Option Substring := - let t := s.commonPrefix pre - if t.bsize = pre.bsize then - some { s with startPos := t.stopPos } - else - none - -/-- -If `suff` is a suffix of `s`, i.e. `s = t ++ suff`, returns the remainder `t`. --/ -def dropSuffix? (s : Substring) (suff : Substring) : Option Substring := - let t := s.commonSuffix suff - if t.bsize = suff.bsize then - some { s with stopPos := t.startPos } - else - none - -end Substring - -namespace String - -/-- -If `pre` is a prefix of `s`, i.e. `s = pre ++ t`, returns the remainder `t`. --/ -def dropPrefix? (s : String) (pre : Substring) : Option Substring := - Substring.dropPrefix? s pre - -/-- -If `suff` is a suffix of `s`, i.e. `s = t ++ suff`, returns the remainder `t`. --/ -def dropSuffix? (s : String) (suff : Substring) : Option Substring := - Substring.dropSuffix? s suff - -/-- `s.stripPrefix pre` will remove `pre` from the beginning of `s` if it occurs there, -or otherwise return `s`. -/ -def stripPrefix (s : String) (pre : Substring) : String := - s.dropPrefix? pre |>.map Substring.toString |>.getD s - -/-- `s.stripSuffix suff` will remove `suff` from the end of `s` if it occurs there, -or otherwise return `s`. -/ -def stripSuffix (s : String) (suff : Substring) : String := - s.dropSuffix? suff |>.map Substring.toString |>.getD s - /-- Count the occurrences of a character in a string. -/ def count (s : String) (c : Char) : Nat := s.foldl (fun n d => if d = c then n + 1 else n) 0 diff --git a/Batteries/Data/String/Lemmas.lean b/Batteries/Data/String/Lemmas.lean index 490c34908b..5854890f25 100644 --- a/Batteries/Data/String/Lemmas.lean +++ b/Batteries/Data/String/Lemmas.lean @@ -447,12 +447,12 @@ theorem split_of_valid (s p) : split s p = (List.splitOnP p s.1).map mk := by attribute [simp] toSubstring' -theorem join_eq (ss : List String) : join ss = ⟨(ss.map data).join⟩ := go ss [] where - go : ∀ (ss : List String) cs, ss.foldl (· ++ ·) (mk cs) = ⟨cs ++ (ss.map data).join⟩ +theorem join_eq (ss : List String) : join ss = ⟨(ss.map data).flatten⟩ := go ss [] where + go : ∀ (ss : List String) cs, ss.foldl (· ++ ·) (mk cs) = ⟨cs ++ (ss.map data).flatten⟩ | [], _ => by simp | ⟨s⟩::ss, _ => (go ss _).trans (by simp) -@[simp] theorem data_join (ss : List String) : (join ss).data = (ss.map data).join := by +@[simp] theorem data_join (ss : List String) : (join ss).data = (ss.map data).flatten := by rw [join_eq] @[deprecated (since := "2024-06-06")] alias append_nil := append_empty diff --git a/Batteries/Data/Sum.lean b/Batteries/Data/Sum.lean deleted file mode 100644 index 51d4518942..0000000000 --- a/Batteries/Data/Sum.lean +++ /dev/null @@ -1,2 +0,0 @@ -import Batteries.Data.Sum.Basic -import Batteries.Data.Sum.Lemmas diff --git a/Batteries/Data/Sum/Basic.lean b/Batteries/Data/Sum/Basic.lean deleted file mode 100644 index 6740f19228..0000000000 --- a/Batteries/Data/Sum/Basic.lean +++ /dev/null @@ -1,164 +0,0 @@ -/- -Copyright (c) 2017 Mario Carneiro. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro, Yury G. Kudryashov --/ - -/-! -# Disjoint union of types - -This file defines basic operations on the the sum type `α ⊕ β`. - -`α ⊕ β` is the type made of a copy of `α` and a copy of `β`. It is also called *disjoint union*. - -## Main declarations - -* `Sum.isLeft`: Returns whether `x : α ⊕ β` comes from the left component or not. -* `Sum.isRight`: Returns whether `x : α ⊕ β` comes from the right component or not. -* `Sum.getLeft`: Retrieves the left content of a `x : α ⊕ β` that is known to come from the left. -* `Sum.getRight`: Retrieves the right content of `x : α ⊕ β` that is known to come from the right. -* `Sum.getLeft?`: Retrieves the left content of `x : α ⊕ β` as an option type or returns `none` - if it's coming from the right. -* `Sum.getRight?`: Retrieves the right content of `x : α ⊕ β` as an option type or returns `none` - if it's coming from the left. -* `Sum.map`: Maps `α ⊕ β` to `γ ⊕ δ` component-wise. -* `Sum.elim`: Nondependent eliminator/induction principle for `α ⊕ β`. -* `Sum.swap`: Maps `α ⊕ β` to `β ⊕ α` by swapping components. -* `Sum.LiftRel`: The disjoint union of two relations. -* `Sum.Lex`: Lexicographic order on `α ⊕ β` induced by a relation on `α` and a relation on `β`. - -## Further material - -See `Batteries.Data.Sum.Lemmas` for theorems about these definitions. - -## Notes - -The definition of `Sum` takes values in `Type _`. This effectively forbids `Prop`- valued sum types. -To this effect, we have `PSum`, which takes value in `Sort _` and carries a more complicated -universe signature in consequence. The `Prop` version is `Or`. --/ - -namespace Sum - -deriving instance DecidableEq for Sum -deriving instance BEq for Sum - -section get - -/-- Check if a sum is `inl`. -/ -def isLeft : α ⊕ β → Bool - | inl _ => true - | inr _ => false - -/-- Check if a sum is `inr`. -/ -def isRight : α ⊕ β → Bool - | inl _ => false - | inr _ => true - -/-- Retrieve the contents from a sum known to be `inl`.-/ -def getLeft : (ab : α ⊕ β) → ab.isLeft → α - | inl a, _ => a - -/-- Retrieve the contents from a sum known to be `inr`.-/ -def getRight : (ab : α ⊕ β) → ab.isRight → β - | inr b, _ => b - -@[simp] theorem isLeft_inl : (inl x : α ⊕ β).isLeft = true := rfl -@[simp] theorem isLeft_inr : (inr x : α ⊕ β).isLeft = false := rfl -@[simp] theorem isRight_inl : (inl x : α ⊕ β).isRight = false := rfl -@[simp] theorem isRight_inr : (inr x : α ⊕ β).isRight = true := rfl - -@[simp] theorem getLeft_inl (h : (inl x : α ⊕ β).isLeft) : (inl x).getLeft h = x := rfl -@[simp] theorem getRight_inr (h : (inr x : α ⊕ β).isRight) : (inr x).getRight h = x := rfl - -@[simp] theorem getLeft?_inl : (inl x : α ⊕ β).getLeft? = some x := rfl -@[simp] theorem getLeft?_inr : (inr x : α ⊕ β).getLeft? = none := rfl -@[simp] theorem getRight?_inl : (inl x : α ⊕ β).getRight? = none := rfl -@[simp] theorem getRight?_inr : (inr x : α ⊕ β).getRight? = some x := rfl - -end get - -/-- Define a function on `α ⊕ β` by giving separate definitions on `α` and `β`. -/ -protected def elim {α β γ} (f : α → γ) (g : β → γ) : α ⊕ β → γ := - fun x => Sum.casesOn x f g - -@[simp] theorem elim_inl (f : α → γ) (g : β → γ) (x : α) : - Sum.elim f g (inl x) = f x := rfl - -@[simp] theorem elim_inr (f : α → γ) (g : β → γ) (x : β) : - Sum.elim f g (inr x) = g x := rfl - -/-- Map `α ⊕ β` to `α' ⊕ β'` sending `α` to `α'` and `β` to `β'`. -/ -protected def map (f : α → α') (g : β → β') : α ⊕ β → α' ⊕ β' := - Sum.elim (inl ∘ f) (inr ∘ g) - -@[simp] theorem map_inl (f : α → α') (g : β → β') (x : α) : (inl x).map f g = inl (f x) := rfl - -@[simp] theorem map_inr (f : α → α') (g : β → β') (x : β) : (inr x).map f g = inr (g x) := rfl - -/-- Swap the factors of a sum type -/ -def swap : α ⊕ β → β ⊕ α := Sum.elim inr inl - -@[simp] theorem swap_inl : swap (inl x : α ⊕ β) = inr x := rfl - -@[simp] theorem swap_inr : swap (inr x : α ⊕ β) = inl x := rfl - -section LiftRel - -/-- Lifts pointwise two relations between `α` and `γ` and between `β` and `δ` to a relation between -`α ⊕ β` and `γ ⊕ δ`. -/ -inductive LiftRel (r : α → γ → Prop) (s : β → δ → Prop) : α ⊕ β → γ ⊕ δ → Prop - /-- `inl a` and `inl c` are related via `LiftRel r s` if `a` and `c` are related via `r`. -/ - | protected inl {a c} : r a c → LiftRel r s (inl a) (inl c) - /-- `inr b` and `inr d` are related via `LiftRel r s` if `b` and `d` are related via `s`. -/ - | protected inr {b d} : s b d → LiftRel r s (inr b) (inr d) - -@[simp] theorem liftRel_inl_inl : LiftRel r s (inl a) (inl c) ↔ r a c := - ⟨fun h => by cases h; assumption, LiftRel.inl⟩ - -@[simp] theorem not_liftRel_inl_inr : ¬LiftRel r s (inl a) (inr d) := nofun - -@[simp] theorem not_liftRel_inr_inl : ¬LiftRel r s (inr b) (inl c) := nofun - -@[simp] theorem liftRel_inr_inr : LiftRel r s (inr b) (inr d) ↔ s b d := - ⟨fun h => by cases h; assumption, LiftRel.inr⟩ - -instance {r : α → γ → Prop} {s : β → δ → Prop} - [∀ a c, Decidable (r a c)] [∀ b d, Decidable (s b d)] : - ∀ (ab : α ⊕ β) (cd : γ ⊕ δ), Decidable (LiftRel r s ab cd) - | inl _, inl _ => decidable_of_iff' _ liftRel_inl_inl - | inl _, inr _ => Decidable.isFalse not_liftRel_inl_inr - | inr _, inl _ => Decidable.isFalse not_liftRel_inr_inl - | inr _, inr _ => decidable_of_iff' _ liftRel_inr_inr - -end LiftRel - -section Lex - -/-- Lexicographic order for sum. Sort all the `inl a` before the `inr b`, otherwise use the -respective order on `α` or `β`. -/ -inductive Lex (r : α → α → Prop) (s : β → β → Prop) : α ⊕ β → α ⊕ β → Prop - /-- `inl a₁` and `inl a₂` are related via `Lex r s` if `a₁` and `a₂` are related via `r`. -/ - | protected inl {a₁ a₂} (h : r a₁ a₂) : Lex r s (inl a₁) (inl a₂) - /-- `inr b₁` and `inr b₂` are related via `Lex r s` if `b₁` and `b₂` are related via `s`. -/ - | protected inr {b₁ b₂} (h : s b₁ b₂) : Lex r s (inr b₁) (inr b₂) - /-- `inl a` and `inr b` are always related via `Lex r s`. -/ - | sep (a b) : Lex r s (inl a) (inr b) - -attribute [simp] Lex.sep - -@[simp] theorem lex_inl_inl : Lex r s (inl a₁) (inl a₂) ↔ r a₁ a₂ := - ⟨fun h => by cases h; assumption, Lex.inl⟩ - -@[simp] theorem lex_inr_inr : Lex r s (inr b₁) (inr b₂) ↔ s b₁ b₂ := - ⟨fun h => by cases h; assumption, Lex.inr⟩ - -@[simp] theorem lex_inr_inl : ¬Lex r s (inr b) (inl a) := nofun - -instance instDecidableRelSumLex [DecidableRel r] [DecidableRel s] : DecidableRel (Lex r s) - | inl _, inl _ => decidable_of_iff' _ lex_inl_inl - | inl _, inr _ => Decidable.isTrue (Lex.sep _ _) - | inr _, inl _ => Decidable.isFalse lex_inr_inl - | inr _, inr _ => decidable_of_iff' _ lex_inr_inr - -end Lex diff --git a/Batteries/Data/Sum/Lemmas.lean b/Batteries/Data/Sum/Lemmas.lean deleted file mode 100644 index ffc24dcd76..0000000000 --- a/Batteries/Data/Sum/Lemmas.lean +++ /dev/null @@ -1,250 +0,0 @@ -/- -Copyright (c) 2017 Mario Carneiro. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro, Yury G. Kudryashov --/ - -import Batteries.Data.Sum.Basic -import Batteries.Logic - -/-! -# Disjoint union of types - -Theorems about the definitions introduced in `Batteries.Data.Sum.Basic`. --/ - -open Function - -namespace Sum - -@[simp] protected theorem «forall» {p : α ⊕ β → Prop} : - (∀ x, p x) ↔ (∀ a, p (inl a)) ∧ ∀ b, p (inr b) := - ⟨fun h => ⟨fun _ => h _, fun _ => h _⟩, fun ⟨h₁, h₂⟩ => Sum.rec h₁ h₂⟩ - -@[simp] protected theorem «exists» {p : α ⊕ β → Prop} : - (∃ x, p x) ↔ (∃ a, p (inl a)) ∨ ∃ b, p (inr b) := - ⟨ fun - | ⟨inl a, h⟩ => Or.inl ⟨a, h⟩ - | ⟨inr b, h⟩ => Or.inr ⟨b, h⟩, - fun - | Or.inl ⟨a, h⟩ => ⟨inl a, h⟩ - | Or.inr ⟨b, h⟩ => ⟨inr b, h⟩⟩ - -theorem forall_sum {γ : α ⊕ β → Sort _} (p : (∀ ab, γ ab) → Prop) : - (∀ fab, p fab) ↔ (∀ fa fb, p (Sum.rec fa fb)) := by - refine ⟨fun h fa fb => h _, fun h fab => ?_⟩ - have h1 : fab = Sum.rec (fun a => fab (Sum.inl a)) (fun b => fab (Sum.inr b)) := by - ext ab; cases ab <;> rfl - rw [h1]; exact h _ _ - -section get - -@[simp] theorem inl_getLeft : ∀ (x : α ⊕ β) (h : x.isLeft), inl (x.getLeft h) = x - | inl _, _ => rfl -@[simp] theorem inr_getRight : ∀ (x : α ⊕ β) (h : x.isRight), inr (x.getRight h) = x - | inr _, _ => rfl - -@[simp] theorem getLeft?_eq_none_iff {x : α ⊕ β} : x.getLeft? = none ↔ x.isRight := by - cases x <;> simp only [getLeft?, isRight, eq_self_iff_true, reduceCtorEq] - -@[simp] theorem getRight?_eq_none_iff {x : α ⊕ β} : x.getRight? = none ↔ x.isLeft := by - cases x <;> simp only [getRight?, isLeft, eq_self_iff_true, reduceCtorEq] - -theorem eq_left_getLeft_of_isLeft : ∀ {x : α ⊕ β} (h : x.isLeft), x = inl (x.getLeft h) - | inl _, _ => rfl - -@[simp] theorem getLeft_eq_iff (h : x.isLeft) : x.getLeft h = a ↔ x = inl a := by - cases x <;> simp at h ⊢ - -theorem eq_right_getRight_of_isRight : ∀ {x : α ⊕ β} (h : x.isRight), x = inr (x.getRight h) - | inr _, _ => rfl - -@[simp] theorem getRight_eq_iff (h : x.isRight) : x.getRight h = b ↔ x = inr b := by - cases x <;> simp at h ⊢ - -@[simp] theorem getLeft?_eq_some_iff : x.getLeft? = some a ↔ x = inl a := by - cases x <;> simp only [getLeft?, Option.some.injEq, inl.injEq, reduceCtorEq] - -@[simp] theorem getRight?_eq_some_iff : x.getRight? = some b ↔ x = inr b := by - cases x <;> simp only [getRight?, Option.some.injEq, inr.injEq, reduceCtorEq] - -@[simp] theorem bnot_isLeft (x : α ⊕ β) : !x.isLeft = x.isRight := by cases x <;> rfl - -@[simp] theorem isLeft_eq_false {x : α ⊕ β} : x.isLeft = false ↔ x.isRight := by cases x <;> simp - -theorem not_isLeft {x : α ⊕ β} : ¬x.isLeft ↔ x.isRight := by simp - -@[simp] theorem bnot_isRight (x : α ⊕ β) : !x.isRight = x.isLeft := by cases x <;> rfl - -@[simp] theorem isRight_eq_false {x : α ⊕ β} : x.isRight = false ↔ x.isLeft := by cases x <;> simp - -theorem not_isRight {x : α ⊕ β} : ¬x.isRight ↔ x.isLeft := by simp - -theorem isLeft_iff : x.isLeft ↔ ∃ y, x = Sum.inl y := by cases x <;> simp - -theorem isRight_iff : x.isRight ↔ ∃ y, x = Sum.inr y := by cases x <;> simp - -end get - -theorem inl.inj_iff : (inl a : α ⊕ β) = inl b ↔ a = b := ⟨inl.inj, congrArg _⟩ - -theorem inr.inj_iff : (inr a : α ⊕ β) = inr b ↔ a = b := ⟨inr.inj, congrArg _⟩ - -theorem inl_ne_inr : inl a ≠ inr b := nofun - -theorem inr_ne_inl : inr b ≠ inl a := nofun - -/-! ### `Sum.elim` -/ - -@[simp] theorem elim_comp_inl (f : α → γ) (g : β → γ) : Sum.elim f g ∘ inl = f := - rfl - -@[simp] theorem elim_comp_inr (f : α → γ) (g : β → γ) : Sum.elim f g ∘ inr = g := - rfl - -@[simp] theorem elim_inl_inr : @Sum.elim α β _ inl inr = id := - funext fun x => Sum.casesOn x (fun _ => rfl) fun _ => rfl - -theorem comp_elim (f : γ → δ) (g : α → γ) (h : β → γ) : - f ∘ Sum.elim g h = Sum.elim (f ∘ g) (f ∘ h) := - funext fun x => Sum.casesOn x (fun _ => rfl) fun _ => rfl - -@[simp] theorem elim_comp_inl_inr (f : α ⊕ β → γ) : - Sum.elim (f ∘ inl) (f ∘ inr) = f := - funext fun x => Sum.casesOn x (fun _ => rfl) fun _ => rfl - -theorem elim_eq_iff {u u' : α → γ} {v v' : β → γ} : - Sum.elim u v = Sum.elim u' v' ↔ u = u' ∧ v = v' := by - simp [funext_iff] - -/-! ### `Sum.map` -/ - -@[simp] theorem map_map (f' : α' → α'') (g' : β' → β'') (f : α → α') (g : β → β') : - ∀ x : Sum α β, (x.map f g).map f' g' = x.map (f' ∘ f) (g' ∘ g) - | inl _ => rfl - | inr _ => rfl - -@[simp] theorem map_comp_map (f' : α' → α'') (g' : β' → β'') (f : α → α') (g : β → β') : - Sum.map f' g' ∘ Sum.map f g = Sum.map (f' ∘ f) (g' ∘ g) := - funext <| map_map f' g' f g - -@[simp] theorem map_id_id : Sum.map (@id α) (@id β) = id := - funext fun x => Sum.recOn x (fun _ => rfl) fun _ => rfl - -theorem elim_map {f₁ : α → β} {f₂ : β → ε} {g₁ : γ → δ} {g₂ : δ → ε} {x} : - Sum.elim f₂ g₂ (Sum.map f₁ g₁ x) = Sum.elim (f₂ ∘ f₁) (g₂ ∘ g₁) x := by - cases x <;> rfl - -theorem elim_comp_map {f₁ : α → β} {f₂ : β → ε} {g₁ : γ → δ} {g₂ : δ → ε} : - Sum.elim f₂ g₂ ∘ Sum.map f₁ g₁ = Sum.elim (f₂ ∘ f₁) (g₂ ∘ g₁) := - funext fun _ => elim_map - -@[simp] theorem isLeft_map (f : α → β) (g : γ → δ) (x : α ⊕ γ) : - isLeft (x.map f g) = isLeft x := by - cases x <;> rfl - -@[simp] theorem isRight_map (f : α → β) (g : γ → δ) (x : α ⊕ γ) : - isRight (x.map f g) = isRight x := by - cases x <;> rfl - -@[simp] theorem getLeft?_map (f : α → β) (g : γ → δ) (x : α ⊕ γ) : - (x.map f g).getLeft? = x.getLeft?.map f := by - cases x <;> rfl - -@[simp] theorem getRight?_map (f : α → β) (g : γ → δ) (x : α ⊕ γ) : - (x.map f g).getRight? = x.getRight?.map g := by cases x <;> rfl - -/-! ### `Sum.swap` -/ - -@[simp] theorem swap_swap (x : α ⊕ β) : swap (swap x) = x := by cases x <;> rfl - -@[simp] theorem swap_swap_eq : swap ∘ swap = @id (α ⊕ β) := funext <| swap_swap - -@[simp] theorem isLeft_swap (x : α ⊕ β) : x.swap.isLeft = x.isRight := by cases x <;> rfl - -@[simp] theorem isRight_swap (x : α ⊕ β) : x.swap.isRight = x.isLeft := by cases x <;> rfl - -@[simp] theorem getLeft?_swap (x : α ⊕ β) : x.swap.getLeft? = x.getRight? := by cases x <;> rfl - -@[simp] theorem getRight?_swap (x : α ⊕ β) : x.swap.getRight? = x.getLeft? := by cases x <;> rfl - -section LiftRel - -theorem LiftRel.mono (hr : ∀ a b, r₁ a b → r₂ a b) (hs : ∀ a b, s₁ a b → s₂ a b) - (h : LiftRel r₁ s₁ x y) : LiftRel r₂ s₂ x y := by - cases h - · exact LiftRel.inl (hr _ _ ‹_›) - · exact LiftRel.inr (hs _ _ ‹_›) - -theorem LiftRel.mono_left (hr : ∀ a b, r₁ a b → r₂ a b) (h : LiftRel r₁ s x y) : - LiftRel r₂ s x y := - (h.mono hr) fun _ _ => id - -theorem LiftRel.mono_right (hs : ∀ a b, s₁ a b → s₂ a b) (h : LiftRel r s₁ x y) : - LiftRel r s₂ x y := - h.mono (fun _ _ => id) hs - -protected theorem LiftRel.swap (h : LiftRel r s x y) : LiftRel s r x.swap y.swap := by - cases h - · exact LiftRel.inr ‹_› - · exact LiftRel.inl ‹_› - -@[simp] theorem liftRel_swap_iff : LiftRel s r x.swap y.swap ↔ LiftRel r s x y := - ⟨fun h => by rw [← swap_swap x, ← swap_swap y]; exact h.swap, LiftRel.swap⟩ - -end LiftRel - -section Lex - -protected theorem LiftRel.lex {a b : α ⊕ β} (h : LiftRel r s a b) : Lex r s a b := by - cases h - · exact Lex.inl ‹_› - · exact Lex.inr ‹_› - -theorem liftRel_subrelation_lex : Subrelation (LiftRel r s) (Lex r s) := LiftRel.lex - -theorem Lex.mono (hr : ∀ a b, r₁ a b → r₂ a b) (hs : ∀ a b, s₁ a b → s₂ a b) (h : Lex r₁ s₁ x y) : - Lex r₂ s₂ x y := by - cases h - · exact Lex.inl (hr _ _ ‹_›) - · exact Lex.inr (hs _ _ ‹_›) - · exact Lex.sep _ _ - -theorem Lex.mono_left (hr : ∀ a b, r₁ a b → r₂ a b) (h : Lex r₁ s x y) : Lex r₂ s x y := - (h.mono hr) fun _ _ => id - -theorem Lex.mono_right (hs : ∀ a b, s₁ a b → s₂ a b) (h : Lex r s₁ x y) : Lex r s₂ x y := - h.mono (fun _ _ => id) hs - -theorem lex_acc_inl (aca : Acc r a) : Acc (Lex r s) (inl a) := by - induction aca with - | intro _ _ IH => - constructor - intro y h - cases h with - | inl h' => exact IH _ h' - -theorem lex_acc_inr (aca : ∀ a, Acc (Lex r s) (inl a)) {b} (acb : Acc s b) : - Acc (Lex r s) (inr b) := by - induction acb with - | intro _ _ IH => - constructor - intro y h - cases h with - | inr h' => exact IH _ h' - | sep => exact aca _ - -theorem lex_wf (ha : WellFounded r) (hb : WellFounded s) : WellFounded (Lex r s) := - have aca : ∀ a, Acc (Lex r s) (inl a) := fun a => lex_acc_inl (ha.apply a) - ⟨fun x => Sum.recOn x aca fun b => lex_acc_inr aca (hb.apply b)⟩ - -end Lex - -theorem elim_const_const (c : γ) : - Sum.elim (const _ c : α → γ) (const _ c : β → γ) = const _ c := by - ext x - cases x <;> rfl - -@[simp] theorem elim_lam_const_lam_const (c : γ) : - Sum.elim (fun _ : α => c) (fun _ : β => c) = fun _ => c := - Sum.elim_const_const c diff --git a/Batteries/Data/UInt.lean b/Batteries/Data/UInt.lean index a7c2a4cbaa..3f9b495629 100644 --- a/Batteries/Data/UInt.lean +++ b/Batteries/Data/UInt.lean @@ -12,9 +12,6 @@ import Batteries.Classes.Order @[simp] theorem UInt8.val_val_eq_toNat (x : UInt8) : x.val.val = x.toNat := rfl -@[simp] theorem UInt8.val_ofNat (n) : - (no_index (OfNat.ofNat n) : UInt8).val = OfNat.ofNat n := rfl - @[simp] theorem UInt8.toNat_ofNat (n) : (no_index (OfNat.ofNat n) : UInt8).toNat = n % UInt8.size := rfl @@ -49,9 +46,6 @@ instance : Batteries.LawfulOrd UInt8 := .compareOfLessAndEq @[simp] theorem UInt16.val_val_eq_toNat (x : UInt16) : x.val.val = x.toNat := rfl -@[simp] theorem UInt16.val_ofNat (n) : - (no_index (OfNat.ofNat n) : UInt16).val = OfNat.ofNat n := rfl - @[simp] theorem UInt16.toNat_ofNat (n) : (no_index (OfNat.ofNat n) : UInt16).toNat = n % UInt16.size := rfl @@ -86,9 +80,6 @@ instance : Batteries.LawfulOrd UInt16 := .compareOfLessAndEq @[simp] theorem UInt32.val_val_eq_toNat (x : UInt32) : x.val.val = x.toNat := rfl -@[simp] theorem UInt32.val_ofNat (n) : - (no_index (OfNat.ofNat n) : UInt32).val = OfNat.ofNat n := rfl - @[simp] theorem UInt32.toNat_ofNat (n) : (no_index (OfNat.ofNat n) : UInt32).toNat = n % UInt32.size := rfl @@ -123,9 +114,6 @@ instance : Batteries.LawfulOrd UInt32 := .compareOfLessAndEq @[simp] theorem UInt64.val_val_eq_toNat (x : UInt64) : x.val.val = x.toNat := rfl -@[simp] theorem UInt64.val_ofNat (n) : - (no_index (OfNat.ofNat n) : UInt64).val = OfNat.ofNat n := rfl - @[simp] theorem UInt64.toNat_ofNat (n) : (no_index (OfNat.ofNat n) : UInt64).toNat = n % UInt64.size := rfl @@ -160,15 +148,11 @@ instance : Batteries.LawfulOrd UInt64 := .compareOfLessAndEq @[simp] theorem USize.val_val_eq_toNat (x : USize) : x.val.val = x.toNat := rfl -@[simp] theorem USize.val_ofNat (n) : - (no_index (OfNat.ofNat n) : USize).val = OfNat.ofNat n := rfl - @[simp] theorem USize.toNat_ofNat (n) : (no_index (OfNat.ofNat n) : USize).toNat = n % USize.size := rfl theorem USize.size_eq : USize.size = 2 ^ System.Platform.numBits := by - have : 1 ≤ 2 ^ System.Platform.numBits := Nat.succ_le_of_lt (Nat.two_pow_pos _) - rw [USize.size, Nat.sub_add_cancel this] + rw [USize.size] theorem USize.le_size : 2 ^ 32 ≤ USize.size := by rw [size_eq] diff --git a/Batteries/Data/UnionFind/Basic.lean b/Batteries/Data/UnionFind/Basic.lean index ce9c99a77e..c48bb5bda1 100644 --- a/Batteries/Data/UnionFind/Basic.lean +++ b/Batteries/Data/UnionFind/Basic.lean @@ -161,11 +161,11 @@ theorem rankD_lt_rankMax (self : UnionFind) (i : Nat) : theorem lt_rankMax (self : UnionFind) (i : Nat) : self.rank i < self.rankMax := rankD_lt_rankMax .. theorem push_rankD (arr : Array UFNode) : rankD (arr.push ⟨arr.size, 0⟩) i = rankD arr i := by - simp only [rankD, Array.size_push, Array.get_eq_getElem, Array.get_push, dite_eq_ite] + simp only [rankD, Array.size_push, Array.get_eq_getElem, Array.getElem_push, dite_eq_ite] split <;> split <;> first | simp | cases ‹¬_› (Nat.lt_succ_of_lt ‹_›) theorem push_parentD (arr : Array UFNode) : parentD (arr.push ⟨arr.size, 0⟩) i = parentD arr i := by - simp only [parentD, Array.size_push, Array.get_eq_getElem, Array.get_push, dite_eq_ite] + simp only [parentD, Array.size_push, Array.get_eq_getElem, Array.getElem_push, dite_eq_ite] split <;> split <;> try simp · exact Nat.le_antisymm (Nat.ge_of_not_lt ‹_›) (Nat.le_of_lt_succ ‹_›) · cases ‹¬_› (Nat.lt_succ_of_lt ‹_›) diff --git a/Batteries/Data/UnionFind/Lemmas.lean b/Batteries/Data/UnionFind/Lemmas.lean index a42ece4508..9a7b0dac58 100644 --- a/Batteries/Data/UnionFind/Lemmas.lean +++ b/Batteries/Data/UnionFind/Lemmas.lean @@ -16,7 +16,7 @@ namespace Batteries.UnionFind @[simp] theorem parentD_push {arr : Array UFNode} : parentD (arr.push ⟨arr.size, 0⟩) a = parentD arr a := by - simp [parentD]; split <;> split <;> try simp [Array.get_push, *] + simp [parentD]; split <;> split <;> try simp [Array.getElem_push, *] · next h1 h2 => simp [Nat.lt_succ] at h1 h2 exact Nat.le_antisymm h2 h1 @@ -26,7 +26,7 @@ namespace Batteries.UnionFind @[simp] theorem rankD_push {arr : Array UFNode} : rankD (arr.push ⟨arr.size, 0⟩) a = rankD arr a := by - simp [rankD]; split <;> split <;> try simp [Array.get_push, *] + simp [rankD]; split <;> split <;> try simp [Array.getElem_push, *] next h1 h2 => cases h1 (Nat.lt_succ_of_lt h2) @[simp] theorem rank_push {m : UnionFind} : m.push.rank a = m.rank a := by simp [rank] diff --git a/Batteries/Data/Vector/Basic.lean b/Batteries/Data/Vector/Basic.lean index 534f94030a..2d4708d5b2 100644 --- a/Batteries/Data/Vector/Basic.lean +++ b/Batteries/Data/Vector/Basic.lean @@ -238,11 +238,13 @@ Otherwise it panics The old entry is returned with the modified vector. def range (n : Nat) : Vector Nat n := ⟨Array.range n, Array.size_range ..⟩ /-- -`shrink v m` shrinks the vector to the first `m` elements if `m < n`. +`take v m` shrinks the vector to the first `m` elements if `m < n`. Returns `v` unchanged if `m ≥ n`. -/ -def shrink (v : Vector α n) (m : Nat) : Vector α (min n m) := - ⟨v.toArray.shrink m, by simp [Array.size_shrink, v.size_eq]⟩ +def take (v : Vector α n) (m : Nat) : Vector α (min n m) := + ⟨v.toArray.take m, by simp [Array.size_take, v.size_eq, Nat.min_comm]⟩ + +@[deprecated (since := "2024-10-22")] alias shrink := take /-- Drops the first (up to) `i` elements from a vector of length `n`. @@ -252,12 +254,6 @@ def drop (i : Nat) (v : Vector α n) : Vector α (n - i) := have : min n n - i = n - i := by rw [Nat.min_self] Vector.cast this (extract v i n) -/-- -Takes the first (up to) `i` elements from a vector of length `n`. - --/ -alias take := shrink - /-- `isEqv` takes a given boolean property `p`. It returns `true` if and only if `p a[i] b[i]` holds true for all valid indices `i`. diff --git a/Batteries/Data/Vector/Lemmas.lean b/Batteries/Data/Vector/Lemmas.lean index 260025344a..d40c7b79cb 100644 --- a/Batteries/Data/Vector/Lemmas.lean +++ b/Batteries/Data/Vector/Lemmas.lean @@ -82,7 +82,7 @@ protected theorem ext {a b : Vector α n} (h : (i : Nat) → (_ : i < n) → a[i @[simp, nolint simpNF] theorem getElem_push_lt {v : Vector α n} {x : α} {i : Nat} (h : i < n) : (v.push x)[i] = v[i] := by rcases v with ⟨data, rfl⟩ - simp [Array.get_push_lt, h] + simp [Array.getElem_push_lt, h] @[simp] theorem getElem_pop {v : Vector α n} {i : Nat} (h : i < n - 1) : (v.pop)[i] = v[i] := by rcases v with ⟨data, rfl⟩ diff --git a/Batteries/Lean/Expr.lean b/Batteries/Lean/Expr.lean index d1c568f3e4..9ac9a09238 100644 --- a/Batteries/Lean/Expr.lean +++ b/Batteries/Lean/Expr.lean @@ -39,16 +39,6 @@ def forallArity : Expr → Nat -- @[deprecated (since := "2024-10-16"), inherit_doc getNumHeadForalls] -- abbrev forallArity := @getNumHeadForalls -/-- Like `getAppNumArgs` but ignores metadata. -/ -def getAppNumArgs' (e : Expr) : Nat := - go e 0 -where - /-- Auxiliary definition for `getAppNumArgs'`. -/ - go : Expr → Nat → Nat - | mdata _ b, n => go b n - | app f _ , n => go f (n + 1) - | _ , n => n - /-- Like `withApp` but ignores metadata. -/ @[inline] def withApp' (e : Expr) (k : Expr → Array Expr → α) : α := diff --git a/Batteries/Logic.lean b/Batteries/Logic.lean index 70091ae973..4bc5f9f102 100644 --- a/Batteries/Logic.lean +++ b/Batteries/Logic.lean @@ -57,9 +57,6 @@ theorem funext₃ {β : α → Sort _} {γ : ∀ a, β a → Sort _} {δ : ∀ a @[deprecated (since := "2024-10-17")] protected alias Function.funext_iff := funext_iff -theorem ne_of_apply_ne {α β : Sort _} (f : α → β) {x y : α} : f x ≠ f y → x ≠ y := - mt <| congrArg _ - protected theorem Eq.congr (h₁ : x₁ = y₁) (h₂ : x₂ = y₂) : x₁ = x₂ ↔ y₁ = y₂ := by subst h₁; subst h₂; rfl @@ -102,17 +99,6 @@ theorem heq_eqRec_iff_heq {α : Sort _} {a : α} {motive : (a' : α) → a = a' HEq y (@Eq.rec α a motive x a' e) ↔ HEq y x := by subst e; rfl -/-! ## membership -/ - -section Mem -variable [Membership α β] {s t : β} {a b : α} - -theorem ne_of_mem_of_not_mem (h : a ∈ s) : b ∉ s → a ≠ b := mt fun e => e ▸ h - -theorem ne_of_mem_of_not_mem' (h : a ∈ s) : a ∉ t → s ≠ t := mt fun e => e ▸ h - -end Mem - /-! ## miscellaneous -/ @[simp] theorem not_nonempty_empty : ¬Nonempty Empty := fun ⟨h⟩ => h.elim diff --git a/Batteries/Tactic/Alias.lean b/Batteries/Tactic/Alias.lean index 471dfd883a..1f22940b0a 100644 --- a/Batteries/Tactic/Alias.lean +++ b/Batteries/Tactic/Alias.lean @@ -109,10 +109,7 @@ elab (name := alias) mods:declModifiers "alias " alias:ident " := " name:ident : addDecl decl else addAndCompile decl - Lean.addDeclarationRanges declName { - range := ← getDeclarationRange (← getRef) - selectionRange := ← getDeclarationRange alias - } + addDeclarationRangesFromSyntax declName (← getRef) alias Term.addTermInfo' alias (← mkConstWithLevelParams declName) (isBinder := true) addDocString' declName declMods.docString? Term.applyAttributes declName declMods.attrs @@ -174,16 +171,10 @@ elab (name := aliasLR) mods:declModifiers "alias " if let `(binderIdent| $idFwd:ident) := aliasFwd then let (declName, _) ← mkDeclName (← getCurrNamespace) declMods idFwd.getId addSide true declName declMods thm - Lean.addDeclarationRanges declName { - range := ← getDeclarationRange (← getRef) - selectionRange := ← getDeclarationRange idFwd - } + addDeclarationRangesFromSyntax declName (← getRef) idFwd Term.addTermInfo' idFwd (← mkConstWithLevelParams declName) (isBinder := true) if let `(binderIdent| $idRev:ident) := aliasRev then let (declName, _) ← mkDeclName (← getCurrNamespace) declMods idRev.getId addSide false declName declMods thm - Lean.addDeclarationRanges declName { - range := ← getDeclarationRange (← getRef) - selectionRange := ← getDeclarationRange idRev - } + addDeclarationRangesFromSyntax declName (← getRef) idRev Term.addTermInfo' idRev (← mkConstWithLevelParams declName) (isBinder := true) diff --git a/Batteries/Tactic/Classical.lean b/Batteries/Tactic/Classical.lean deleted file mode 100644 index a3bc78b0a4..0000000000 --- a/Batteries/Tactic/Classical.lean +++ /dev/null @@ -1,41 +0,0 @@ -/- -Copyright (c) 2021 Mario Carneiro. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro --/ -import Lean.Elab.Tactic.Basic - -/-! # `classical` tactic -/ - -namespace Batteries.Tactic -open Lean Meta Elab.Tactic - -/-- -`classical t` runs `t` in a scope where `Classical.propDecidable` is a low priority -local instance. --/ -def classical [Monad m] [MonadEnv m] [MonadFinally m] [MonadLiftT MetaM m] (t : m α) : - m α := do - modifyEnv Meta.instanceExtension.pushScope - Meta.addInstance ``Classical.propDecidable .local 10 - try - t - finally - modifyEnv Meta.instanceExtension.popScope - -/-- `classical!` has been removed; use `classical` instead -/ --- Deprecated 2024-04-19 -elab "classical!" : tactic => do - throwError "`classical!` has been removed; use `classical` instead" - -/-- -`classical tacs` runs `tacs` in a scope where `Classical.propDecidable` is a low priority -local instance. - -Note that (unlike lean 3) `classical` is a scoping tactic - it adds the instance only within the -scope of the tactic. --/ --- FIXME: using ppDedent looks good in the common case, but produces the incorrect result when --- the `classical` does not scope over the rest of the block. -elab "classical" tacs:ppDedent(tacticSeq) : tactic => do - classical <| Elab.Tactic.evalTactic tacs diff --git a/Batteries/Tactic/HelpCmd.lean b/Batteries/Tactic/HelpCmd.lean index 7fbe67bf3a..1d955310a0 100644 --- a/Batteries/Tactic/HelpCmd.lean +++ b/Batteries/Tactic/HelpCmd.lean @@ -306,7 +306,7 @@ elab "#help " colGt &"note" colGt ppSpace name:strLit : command => do let valid_entries := imported_entries_filtered ++ local_entries.filterMap fun x => if label_prefix.isPrefixOf x.fst then some x else none let grouped_valid_entries := valid_entries.mergeSort (·.fst ≤ ·.fst) - |>.groupBy (·.fst == ·.fst) + |>.splitBy (·.fst == ·.fst) -- display results in a readable style if grouped_valid_entries.isEmpty then diff --git a/Batteries/Tactic/PrintPrefix.lean b/Batteries/Tactic/PrintPrefix.lean index dbedad072f..42b879b9b7 100644 --- a/Batteries/Tactic/PrintPrefix.lean +++ b/Batteries/Tactic/PrintPrefix.lean @@ -27,7 +27,7 @@ structure PrintPrefixConfig where internals : Bool := false /-- Function elaborating `Config`. -/ -declare_config_elab elabPrintPrefixConfig PrintPrefixConfig +declare_command_config_elab elabPrintPrefixConfig PrintPrefixConfig /-- `reverseName name` reverses the components of a name. @@ -109,11 +109,12 @@ The complete set of flags can be seen in the documentation for `Lean.Elab.Command.PrintPrefixConfig`. -/ elab (name := printPrefix) tk:"#print " colGt "prefix" - cfg:(Lean.Parser.Tactic.config)? name:(ident)? : command => liftTermElabM do + cfg:Lean.Parser.Tactic.optConfig name:(ident)? : command => do if let some name := name then - let opts ← elabPrintPrefixConfig (mkOptionalNode cfg) - let mut msgs ← matchingConstants opts name.getId - if msgs.isEmpty then - if let [name] ← resolveGlobalConst name then - msgs ← matchingConstants opts name - logInfoAt tk (.joinSep msgs.toList "") + let opts ← elabPrintPrefixConfig cfg + liftTermElabM do + let mut msgs ← matchingConstants opts name.getId + if msgs.isEmpty then + if let [name] ← resolveGlobalConst name then + msgs ← matchingConstants opts name + logInfoAt tk (.joinSep msgs.toList "") diff --git a/Batteries/Util/LibraryNote.lean b/Batteries/Util/LibraryNote.lean index 3d578c342b..8268c757ca 100644 --- a/Batteries/Util/LibraryNote.lean +++ b/Batteries/Util/LibraryNote.lean @@ -21,7 +21,7 @@ deriving Inhabited initialize libraryNoteExt : SimplePersistentEnvExtension LibraryNoteEntry (Array LibraryNoteEntry) ← registerSimplePersistentEnvExtension { addEntryFn := Array.push - addImportedFn := Array.concatMap id + addImportedFn := Array.flatMap id } open Lean Parser Command in diff --git a/lean-toolchain b/lean-toolchain index 4f86f953fb..0bef727630 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.13.0 +leanprover/lean4:v4.14.0-rc1 diff --git a/test/array.lean b/test/array.lean index 2ba3a29577..89f784293e 100644 --- a/test/array.lean +++ b/test/array.lean @@ -10,7 +10,7 @@ variable (g : i < (a.set! i v).size) variable (j_lt : j < (a.set! i v).size) #check_simp (a.set! i v).get ⟨i, g⟩ ~> v -#check_simp (a.set! i v).get! i ~> if i < a.size then v else default +#check_simp (a.set! i v).get! i ~> (a.setD i v)[i]! #check_simp (a.set! i v).getD i d ~> if i < a.size then v else d #check_simp (a.set! i v)[i] ~> v diff --git a/test/case.lean b/test/case.lean index 4e7892d964..b8e3444d1d 100644 --- a/test/case.lean +++ b/test/case.lean @@ -182,16 +182,6 @@ example : True ∧ ∀ x : Nat, x = x := by rfl -- Test focusing by full match, suffix match, and prefix match -/-- -warning: unused variable `x` -note: this linter can be disabled with `set_option linter.unusedVariables false` ---- -warning: unused variable `y` -note: this linter can be disabled with `set_option linter.unusedVariables false` ---- -warning: unused variable `z` -note: this linter can be disabled with `set_option linter.unusedVariables false` --/ #guard_msgs in example : True := by have x : Bool := ?a diff --git a/test/classical.lean b/test/classical.lean deleted file mode 100644 index d2d460253c..0000000000 --- a/test/classical.lean +++ /dev/null @@ -1,23 +0,0 @@ -import Batteries.Tactic.Classical -import Batteries.Tactic.PermuteGoals - -example : Bool := by - fail_if_success have := ∀ p, decide p -- no classical in scope - classical - have := ∀ p, decide p -- uses the classical instance - guard_expr decide (0 < 1) = @decide (0 < 1) (Nat.decLt 0 1) - exact decide (0 < 1) -- will use the decidable instance - --- double check no leakage -example : Bool := by - fail_if_success have := ∀ p, decide p -- no classical in scope - exact decide (0 < 1) -- uses the decidable instance - --- check that classical respects tactic blocks -example : Bool := by - fail_if_success have := ∀ p, decide p -- no classical in scope - on_goal 1 => - classical - have := ∀ p, decide p -- uses the classical instance - fail_if_success have := ∀ p, decide p -- no classical in scope again - exact decide (0 < 1) -- will use the decidable instance diff --git a/test/congr.lean b/test/congr.lean index 065a899f99..eee4c6dbac 100644 --- a/test/congr.lean +++ b/test/congr.lean @@ -38,11 +38,9 @@ section -- In order to preserve the test behaviour we locally remove the `ext` attribute. attribute [-ext] List.ext_getElem? -private opaque List.sum : List Nat → Nat - example {ls : List Nat} : - (ls.map fun x => (ls.map fun y => 1 + y).sum + 1) = - (ls.map fun x => (ls.map fun y => Nat.succ y).sum + 1) := by + (ls.map fun _ => (ls.map fun y => 1 + y).sum + 1) = + (ls.map fun _ => (ls.map fun y => Nat.succ y).sum + 1) := by rcongr (_x y) guard_target =ₐ 1 + y = y.succ rw [Nat.add_comm] diff --git a/test/help_cmd.lean b/test/help_cmd.lean index 57e214cafd..23e3698b6d 100644 --- a/test/help_cmd.lean +++ b/test/help_cmd.lean @@ -184,22 +184,106 @@ error: no command declarations start with foobarbaz #help command foobarbaz /-- -info: -syntax "#eval"... [Lean.Parser.Command.eval] +info: syntax "#eval"... [Lean.Parser.Command.eval] + `#eval e` evaluates the expression `e` by compiling and evaluating it. + ⏎ + * The command attempts to use `ToExpr`, `Repr`, or `ToString` instances to print the result. + * If `e` is a monadic value of type `m ty`, then the command tries to adapt the monad `m` + to one of the monads that `#eval` supports, which include `IO`, `CoreM`, `MetaM`, `TermElabM`, and `CommandElabM`. + Users can define `MonadEval` instances to extend the list of supported monads. + ⏎ + The `#eval` command gracefully degrades in capability depending on what is imported. + Importing the `Lean.Elab.Command` module provides full capabilities. + ⏎ + Due to unsoundness, `#eval` refuses to evaluate expressions that depend on `sorry`, even indirectly, + since the presence of `sorry` can lead to runtime instability and crashes. + This check can be overridden with the `#eval! e` command. + ⏎ + Options: + * If `eval.pp` is true (default: true) then tries to use `ToExpr` instances to make use of the + usual pretty printer. Otherwise, only tries using `Repr` and `ToString` instances. + * If `eval.type` is true (default: false) then pretty prints the type of the evaluated value. + * If `eval.derive.repr` is true (default: true) then attempts to auto-derive a `Repr` instance + when there is no other way to print the result. + ⏎ + See also: `#reduce e` for evaluation by term reduction. syntax "#eval!"... [Lean.Parser.Command.evalBang] - -syntax "#exit"... [Lean.Parser.Command.exit] + `#eval e` evaluates the expression `e` by compiling and evaluating it. + ⏎ + * The command attempts to use `ToExpr`, `Repr`, or `ToString` instances to print the result. + * If `e` is a monadic value of type `m ty`, then the command tries to adapt the monad `m` + to one of the monads that `#eval` supports, which include `IO`, `CoreM`, `MetaM`, `TermElabM`, and `CommandElabM`. + Users can define `MonadEval` instances to extend the list of supported monads. + ⏎ + The `#eval` command gracefully degrades in capability depending on what is imported. + Importing the `Lean.Elab.Command` module provides full capabilities. + ⏎ + Due to unsoundness, `#eval` refuses to evaluate expressions that depend on `sorry`, even indirectly, + since the presence of `sorry` can lead to runtime instability and crashes. + This check can be overridden with the `#eval! e` command. + ⏎ + Options: + * If `eval.pp` is true (default: true) then tries to use `ToExpr` instances to make use of the + usual pretty printer. Otherwise, only tries using `Repr` and `ToString` instances. + * If `eval.type` is true (default: false) then pretty prints the type of the evaluated value. + * If `eval.derive.repr` is true (default: true) then attempts to auto-derive a `Repr` instance + when there is no other way to print the result. + ⏎ + See also: `#reduce e` for evaluation by term reduction. syntax "#exit"... [Lean.Parser.Command.exit] -/ #guard_msgs in #help command "#e" /-- -info: -syntax "#eval"... [Lean.Parser.Command.eval] +info: syntax "#eval"... [Lean.Parser.Command.eval] + `#eval e` evaluates the expression `e` by compiling and evaluating it. + ⏎ + * The command attempts to use `ToExpr`, `Repr`, or `ToString` instances to print the result. + * If `e` is a monadic value of type `m ty`, then the command tries to adapt the monad `m` + to one of the monads that `#eval` supports, which include `IO`, `CoreM`, `MetaM`, `TermElabM`, and `CommandElabM`. + Users can define `MonadEval` instances to extend the list of supported monads. + ⏎ + The `#eval` command gracefully degrades in capability depending on what is imported. + Importing the `Lean.Elab.Command` module provides full capabilities. + ⏎ + Due to unsoundness, `#eval` refuses to evaluate expressions that depend on `sorry`, even indirectly, + since the presence of `sorry` can lead to runtime instability and crashes. + This check can be overridden with the `#eval! e` command. + ⏎ + Options: + * If `eval.pp` is true (default: true) then tries to use `ToExpr` instances to make use of the + usual pretty printer. Otherwise, only tries using `Repr` and `ToString` instances. + * If `eval.type` is true (default: false) then pretty prints the type of the evaluated value. + * If `eval.derive.repr` is true (default: true) then attempts to auto-derive a `Repr` instance + when there is no other way to print the result. + ⏎ + See also: `#reduce e` for evaluation by term reduction. + command elab Lean.Elab.Command.elabEval syntax "#eval!"... [Lean.Parser.Command.evalBang] + `#eval e` evaluates the expression `e` by compiling and evaluating it. + ⏎ + * The command attempts to use `ToExpr`, `Repr`, or `ToString` instances to print the result. + * If `e` is a monadic value of type `m ty`, then the command tries to adapt the monad `m` + to one of the monads that `#eval` supports, which include `IO`, `CoreM`, `MetaM`, `TermElabM`, and `CommandElabM`. + Users can define `MonadEval` instances to extend the list of supported monads. + ⏎ + The `#eval` command gracefully degrades in capability depending on what is imported. + Importing the `Lean.Elab.Command` module provides full capabilities. + ⏎ + Due to unsoundness, `#eval` refuses to evaluate expressions that depend on `sorry`, even indirectly, + since the presence of `sorry` can lead to runtime instability and crashes. + This check can be overridden with the `#eval! e` command. + ⏎ + Options: + * If `eval.pp` is true (default: true) then tries to use `ToExpr` instances to make use of the + usual pretty printer. Otherwise, only tries using `Repr` and `ToString` instances. + * If `eval.type` is true (default: false) then pretty prints the type of the evaluated value. + * If `eval.derive.repr` is true (default: true) then attempts to auto-derive a `Repr` instance + when there is no other way to print the result. + ⏎ + See also: `#reduce e` for evaluation by term reduction. + command elab Lean.Elab.Command.elabEvalBang syntax "#exit"... [Lean.Parser.Command.exit] diff --git a/test/lint_simpNF.lean b/test/lint_simpNF.lean index e06e067d18..c9c8f9d307 100644 --- a/test/lint_simpNF.lean +++ b/test/lint_simpNF.lean @@ -2,9 +2,6 @@ import Batteries.Tactic.Lint set_option linter.missingDocs false -protected def Sum.elim {α β γ : Sort _} (f : α → γ) (g : β → γ) : Sum α β → γ := - fun x => Sum.casesOn x f g - structure Equiv (α : Sort _) (β : Sort _) where toFun : α → β invFun : β → α diff --git a/test/print_prefix.lean b/test/print_prefix.lean index 0d14e95aa3..32467e4b37 100644 --- a/test/print_prefix.lean +++ b/test/print_prefix.lean @@ -14,7 +14,7 @@ TEmpty.recOn.{u} (motive : TEmpty → Sort u) (t : TEmpty) : motive t /-- info: -/ #guard_msgs in -#print prefix (config := {imported := false}) Empty +#print prefix -imported Empty namespace EmptyPrefixTest @@ -52,10 +52,10 @@ TestStruct.casesOn.{u} {motive : TestStruct → Sort u} (t : TestStruct) (mk : (foo bar : Int) → motive { foo := foo, bar := bar }) : motive t TestStruct.foo (self : TestStruct) : Int TestStruct.mk (foo bar : Int) : TestStruct -TestStruct.mk.inj {foo bar : Int} : - ∀ {foo_1 bar_1 : Int}, { foo := foo, bar := bar } = { foo := foo_1, bar := bar_1 } → foo = foo_1 ∧ bar = bar_1 -TestStruct.mk.injEq (foo bar : Int) : - ∀ (foo_1 bar_1 : Int), ({ foo := foo, bar := bar } = { foo := foo_1, bar := bar_1 }) = (foo = foo_1 ∧ bar = bar_1) +TestStruct.mk.inj {foo bar foo✝ bar✝ : Int} : + { foo := foo, bar := bar } = { foo := foo✝, bar := bar✝ } → foo = foo✝ ∧ bar = bar✝ +TestStruct.mk.injEq (foo bar foo✝ bar✝ : Int) : + ({ foo := foo, bar := bar } = { foo := foo✝, bar := bar✝ }) = (foo = foo✝ ∧ bar = bar✝) TestStruct.mk.sizeOf_spec (foo bar : Int) : sizeOf { foo := foo, bar := bar } = 1 + sizeOf foo + sizeOf bar TestStruct.noConfusion.{u} {P : Sort u} {v1 v2 : TestStruct} (h12 : v1 = v2) : TestStruct.noConfusionType P v1 v2 TestStruct.noConfusionType.{u} (P : Sort u) (v1 v2 : TestStruct) : Sort u @@ -82,17 +82,17 @@ TestStruct.recOn.{u} {motive : TestStruct → Sort u} (t : TestStruct) (mk : (foo bar : Int) → motive { foo := foo, bar := bar }) : motive t -/ #guard_msgs in -#print prefix (config := {propositions := false}) TestStruct +#print prefix -propositions TestStruct /-- -info: TestStruct.mk.inj {foo bar : Int} : - ∀ {foo_1 bar_1 : Int}, { foo := foo, bar := bar } = { foo := foo_1, bar := bar_1 } → foo = foo_1 ∧ bar = bar_1 -TestStruct.mk.injEq (foo bar : Int) : - ∀ (foo_1 bar_1 : Int), ({ foo := foo, bar := bar } = { foo := foo_1, bar := bar_1 }) = (foo = foo_1 ∧ bar = bar_1) +info: TestStruct.mk.inj {foo bar foo✝ bar✝ : Int} : + { foo := foo, bar := bar } = { foo := foo✝, bar := bar✝ } → foo = foo✝ ∧ bar = bar✝ +TestStruct.mk.injEq (foo bar foo✝ bar✝ : Int) : + ({ foo := foo, bar := bar } = { foo := foo✝, bar := bar✝ }) = (foo = foo✝ ∧ bar = bar✝) TestStruct.mk.sizeOf_spec (foo bar : Int) : sizeOf { foo := foo, bar := bar } = 1 + sizeOf foo + sizeOf bar -/ #guard_msgs in -#print prefix (config := {propositionsOnly := true}) TestStruct +#print prefix +propositionsOnly TestStruct /-- info: TestStruct @@ -109,7 +109,7 @@ TestStruct.rec TestStruct.recOn -/ #guard_msgs in -#print prefix (config := {showTypes := false}) TestStruct +#print prefix -showTypes TestStruct /-- Artificial test function to show #print prefix filters out internals @@ -133,21 +133,17 @@ testMatchProof._cstage1 (n : Nat) : Fin n → Unit testMatchProof._cstage2 : _obj → _obj → _obj testMatchProof._sunfold (n : Nat) : Fin n → Unit testMatchProof._unsafe_rec (n : Nat) : Fin n → Unit -testMatchProof.match_1.{u_1} (motive : (x : Nat) → Fin x → Sort u_1) : - (x : Nat) → - (x_1 : Fin x) → - ((n : Nat) → (isLt : 0 < n) → motive n ⟨0, isLt⟩) → - ((as i : Nat) → (h : i.succ < as.succ) → motive as.succ ⟨i.succ, h⟩) → motive x x_1 -testMatchProof.match_1._cstage1.{u_1} (motive : (x : Nat) → Fin x → Sort u_1) : - (x : Nat) → - (x_1 : Fin x) → - ((n : Nat) → (isLt : 0 < n) → motive n ⟨0, isLt⟩) → - ((as i : Nat) → (h : i.succ < as.succ) → motive as.succ ⟨i.succ, h⟩) → motive x x_1 +testMatchProof.match_1.{u_1} (motive : (x : Nat) → Fin x → Sort u_1) (x✝ : Nat) (x✝¹ : Fin x✝) + (h_1 : (n : Nat) → (isLt : 0 < n) → motive n ⟨0, isLt⟩) + (h_2 : (as i : Nat) → (h : i.succ < as.succ) → motive as.succ ⟨i.succ, h⟩) : motive x✝ x✝¹ +testMatchProof.match_1._cstage1.{u_1} (motive : (x : Nat) → Fin x → Sort u_1) (x✝ : Nat) (x✝¹ : Fin x✝) + (h_1 : (n : Nat) → (isLt : 0 < n) → motive n ⟨0, isLt⟩) + (h_2 : (as i : Nat) → (h : i.succ < as.succ) → motive as.succ ⟨i.succ, h⟩) : motive x✝ x✝¹ testMatchProof.proof_1 (as i : Nat) (h : i.succ < as.succ) : i.succ ≤ as testMatchProof.proof_2 (as i : Nat) (h : i.succ < as.succ) : i.succ ≤ as -/ #guard_msgs in -#print prefix (config := {internals := true}) testMatchProof +#print prefix +internals testMatchProof private inductive TestInd where | foo : TestInd diff --git a/test/simpa.lean b/test/simpa.lean index 9528b26382..edbd58c161 100644 --- a/test/simpa.lean +++ b/test/simpa.lean @@ -70,7 +70,7 @@ end Prod theorem implicit_lambda (h : ∀ {x : Nat}, a = x) : a = 2 := by simpa using h -theorem implicit_lambda2 (h : a = 2) : ∀ {x : Nat}, a = 2 := by +theorem implicit_lambda2 (h : a = 2) : ∀ {_ : Nat}, a = 2 := by simpa using h theorem no_implicit_lambda (h : ∀ {x : Nat}, a = x) : ∀ {x : Nat}, a = x := by diff --git a/test/vector.lean b/test/vector.lean index 5a8b9e8587..7c93888528 100644 --- a/test/vector.lean +++ b/test/vector.lean @@ -17,7 +17,7 @@ def and : Gate 2 := .if (.if (.const true) (.const false)) (.if (.const false) ( def eval (g : Gate n) (v : Vector Bool n) : Bool := match g, v with | .const b, _ => b - | .if g₁ g₂, v => if v.1.back then eval g₁ v.pop else eval g₂ v.pop + | .if g₁ g₂, v => if v.1.back! then eval g₁ v.pop else eval g₂ v.pop example : ∀ v, and.eval v = (v[0] && v[1]) := by decide example : ∃ v, and.eval v = false := by decide From 401646a0c432318bd679969e8870c9331bf7e37b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 4 Nov 2024 15:42:54 -0500 Subject: [PATCH 144/177] feat: add `drop` and `take` for streams (#985) --- Batteries.lean | 1 + Batteries/Data/Stream.lean | 93 ++++++++++++++++++++++++++++++++++++++ 2 files changed, 94 insertions(+) create mode 100644 Batteries/Data/Stream.lean diff --git a/Batteries.lean b/Batteries.lean index b66a913263..8f88944fea 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -31,6 +31,7 @@ import Batteries.Data.PairingHeap import Batteries.Data.RBMap import Batteries.Data.Range import Batteries.Data.Rat +import Batteries.Data.Stream import Batteries.Data.String import Batteries.Data.UInt import Batteries.Data.UnionFind diff --git a/Batteries/Data/Stream.lean b/Batteries/Data/Stream.lean new file mode 100644 index 0000000000..fb9df787be --- /dev/null +++ b/Batteries/Data/Stream.lean @@ -0,0 +1,93 @@ +/- +Copyright (c) 2024 François G. Dorais. All rights reserved. +Released under Apache 2. license as described in the file LICENSE. +Authors: François G. Dorais +-/ + +namespace Stream + +/-- Drop up to `n` values from the stream `s`. -/ +def drop [Stream σ α] (s : σ) : Nat → σ + | 0 => s + | n+1 => + match next? s with + | none => s + | some (_, s) => drop s n + +/-- Read up to `n` values from the stream `s` as a list from first to last. -/ +def take [Stream σ α] (s : σ) : Nat → List α × σ + | 0 => ([], s) + | n+1 => + match next? s with + | none => ([], s) + | some (a, s) => + match take s n with + | (as, s) => (a :: as, s) + +@[simp] theorem fst_take_zero [Stream σ α] (s : σ) : + (take s 0).fst = [] := rfl + +theorem fst_take_succ [Stream σ α] (s : σ) : + (take s (n+1)).fst = match next? s with + | none => [] + | some (a, s) => a :: (take s n).fst := by + simp only [take]; split <;> rfl + +@[simp] theorem snd_take_eq_drop [Stream σ α] (s : σ) (n : Nat) : + (take s n).snd = drop s n := by + induction n generalizing s with + | zero => rfl + | succ n ih => + simp only [take, drop] + split <;> simp [ih] + +/-- Tail recursive version of `Stream.take`. -/ +def takeTR [Stream σ α] (s : σ) (n : Nat) : List α × σ := + loop s [] n +where + /-- Inner loop for `Stream.takeTR`. -/ + loop (s : σ) (acc : List α) + | 0 => (acc.reverse, s) + | n+1 => match next? s with + | none => (acc.reverse, s) + | some (a, s) => loop s (a :: acc) n + +theorem fst_takeTR_loop [Stream σ α] (s : σ) (acc : List α) (n : Nat) : + (takeTR.loop s acc n).fst = acc.reverseAux (take s n).fst := by + induction n generalizing acc s with + | zero => rfl + | succ n ih => simp only [take, takeTR.loop]; split; rfl; simp [ih] + +theorem fst_takeTR [Stream σ α] (s : σ) (n : Nat) : (takeTR s n).fst = (take s n).fst := + fst_takeTR_loop .. + +theorem snd_takeTR_loop [Stream σ α] (s : σ) (acc : List α) (n : Nat) : + (takeTR.loop s acc n).snd = drop s n := by + induction n generalizing acc s with + | zero => rfl + | succ n ih => simp only [takeTR.loop, drop]; split; rfl; simp [ih] + +theorem snd_takeTR [Stream σ α] (s : σ) (n : Nat) : + (takeTR s n).snd = drop s n := snd_takeTR_loop .. + +@[csimp] theorem take_eq_takeTR : @take = @takeTR := by + funext; ext : 1; rw [fst_takeTR]; rw [snd_takeTR, snd_take_eq_drop] + +end Stream + +@[simp] theorem List.stream_drop_eq_drop (l : List α) : Stream.drop l n = l.drop n := by + induction n generalizing l with + | zero => rfl + | succ n ih => + match l with + | [] => rfl + | _::_ => simp [Stream.drop, List.drop_succ_cons, ih] + +@[simp] theorem List.stream_take_eq_take_drop (l : List α) : + Stream.take l n = (l.take n, l.drop n) := by + induction n generalizing l with + | zero => rfl + | succ n ih => + match l with + | [] => rfl + | _::_ => simp [Stream.take, ih] From af731107d531b39cd7278c73f91c573f40340612 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Tue, 5 Nov 2024 01:25:47 +0100 Subject: [PATCH 145/177] chore: turn batteries tests into a library (#1024) Co-authored-by: Blizzard_inc --- .github/workflows/build.yml | 2 +- Batteries.lean | 3 --- .../Test => BatteriesTest}/Internal/DummyLabelAttr.lean | 0 .../Test => BatteriesTest}/Internal/DummyLibraryNote.lean | 0 .../Test => BatteriesTest}/Internal/DummyLibraryNote2.lean | 2 +- {test => BatteriesTest}/MLList.lean | 0 {test => BatteriesTest}/absurd.lean | 0 {test => BatteriesTest}/alias.lean | 0 {test => BatteriesTest}/array.lean | 0 {test => BatteriesTest}/by_contra.lean | 0 {test => BatteriesTest}/case.lean | 0 {test => BatteriesTest}/congr.lean | 0 {test => BatteriesTest}/conv_equals.lean | 0 {test => BatteriesTest}/exfalso.lean | 0 {test => BatteriesTest}/float.lean | 0 {test => BatteriesTest}/help_cmd.lean | 0 {test => BatteriesTest}/import_lean.lean | 0 {test => BatteriesTest}/instances.lean | 0 {test => BatteriesTest}/isIndependentOf.lean | 0 {test => BatteriesTest}/kmp_matcher.lean | 0 {test => BatteriesTest}/lemma_cmd.lean | 0 {test => BatteriesTest}/library_note.lean | 2 +- {test => BatteriesTest}/lintTC.lean | 0 {test => BatteriesTest}/lint_docBlame.lean | 0 {test => BatteriesTest}/lint_docBlameThm.lean | 0 {test => BatteriesTest}/lint_dupNamespace.lean | 0 {test => BatteriesTest}/lint_empty.lean | 0 {test => BatteriesTest}/lint_lean.lean | 0 {test => BatteriesTest}/lint_simpNF.lean | 0 {test => BatteriesTest}/lint_unreachableTactic.lean | 0 {test => BatteriesTest}/lintsimp.lean | 0 {test => BatteriesTest}/lintunused.lean | 0 {test => BatteriesTest}/list_sublists.lean | 0 {test => BatteriesTest}/nondet.lean | 0 {test => BatteriesTest}/norm_cast.lean | 0 {test => BatteriesTest}/omega/benchmark.lean | 0 {test => BatteriesTest}/on_goal.lean | 0 {test => BatteriesTest}/print_prefix.lean | 0 {test => BatteriesTest}/proof_wanted.lean | 0 {test => BatteriesTest}/register_label_attr.lean | 2 +- {test => BatteriesTest}/rfl.lean | 0 {test => BatteriesTest}/seq_focus.lean | 0 {test => BatteriesTest}/show_term.lean | 0 {test => BatteriesTest}/show_unused.lean | 0 {test => BatteriesTest}/simp_trace.lean | 0 {test => BatteriesTest}/simpa.lean | 0 {test => BatteriesTest}/solve_by_elim.lean | 2 +- {test => BatteriesTest}/trans.lean | 0 {test => BatteriesTest}/tryThis.lean | 0 {test => BatteriesTest}/vector.lean | 0 {test => BatteriesTest}/where.lean | 0 lakefile.toml | 7 ++++++- 52 files changed, 11 insertions(+), 9 deletions(-) rename {Batteries/Test => BatteriesTest}/Internal/DummyLabelAttr.lean (100%) rename {Batteries/Test => BatteriesTest}/Internal/DummyLibraryNote.lean (100%) rename {Batteries/Test => BatteriesTest}/Internal/DummyLibraryNote2.lean (88%) rename {test => BatteriesTest}/MLList.lean (100%) rename {test => BatteriesTest}/absurd.lean (100%) rename {test => BatteriesTest}/alias.lean (100%) rename {test => BatteriesTest}/array.lean (100%) rename {test => BatteriesTest}/by_contra.lean (100%) rename {test => BatteriesTest}/case.lean (100%) rename {test => BatteriesTest}/congr.lean (100%) rename {test => BatteriesTest}/conv_equals.lean (100%) rename {test => BatteriesTest}/exfalso.lean (100%) rename {test => BatteriesTest}/float.lean (100%) rename {test => BatteriesTest}/help_cmd.lean (100%) rename {test => BatteriesTest}/import_lean.lean (100%) rename {test => BatteriesTest}/instances.lean (100%) rename {test => BatteriesTest}/isIndependentOf.lean (100%) rename {test => BatteriesTest}/kmp_matcher.lean (100%) rename {test => BatteriesTest}/lemma_cmd.lean (100%) rename {test => BatteriesTest}/library_note.lean (96%) rename {test => BatteriesTest}/lintTC.lean (100%) rename {test => BatteriesTest}/lint_docBlame.lean (100%) rename {test => BatteriesTest}/lint_docBlameThm.lean (100%) rename {test => BatteriesTest}/lint_dupNamespace.lean (100%) rename {test => BatteriesTest}/lint_empty.lean (100%) rename {test => BatteriesTest}/lint_lean.lean (100%) rename {test => BatteriesTest}/lint_simpNF.lean (100%) rename {test => BatteriesTest}/lint_unreachableTactic.lean (100%) rename {test => BatteriesTest}/lintsimp.lean (100%) rename {test => BatteriesTest}/lintunused.lean (100%) rename {test => BatteriesTest}/list_sublists.lean (100%) rename {test => BatteriesTest}/nondet.lean (100%) rename {test => BatteriesTest}/norm_cast.lean (100%) rename {test => BatteriesTest}/omega/benchmark.lean (100%) rename {test => BatteriesTest}/on_goal.lean (100%) rename {test => BatteriesTest}/print_prefix.lean (100%) rename {test => BatteriesTest}/proof_wanted.lean (100%) rename {test => BatteriesTest}/register_label_attr.lean (92%) rename {test => BatteriesTest}/rfl.lean (100%) rename {test => BatteriesTest}/seq_focus.lean (100%) rename {test => BatteriesTest}/show_term.lean (100%) rename {test => BatteriesTest}/show_unused.lean (100%) rename {test => BatteriesTest}/simp_trace.lean (100%) rename {test => BatteriesTest}/simpa.lean (100%) rename {test => BatteriesTest}/solve_by_elim.lean (97%) rename {test => BatteriesTest}/trans.lean (100%) rename {test => BatteriesTest}/tryThis.lean (100%) rename {test => BatteriesTest}/vector.lean (100%) rename {test => BatteriesTest}/where.lean (100%) diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 61af30e96e..8f8349df89 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -54,4 +54,4 @@ jobs: - name: Don't 'import Lean', use precise imports if: always() run: | - ! (find . -name "*.lean" ! -path "./test/import_lean.lean" -type f -print0 | xargs -0 grep -E -n '^import Lean$') + ! (find . -name "*.lean" ! -path "./BatteriesTest/import_lean.lean" -type f -print0 | xargs -0 grep -E -n '^import Lean$') diff --git a/Batteries.lean b/Batteries.lean index 8f88944fea..7e5c73f5fd 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -94,9 +94,6 @@ import Batteries.Tactic.SqueezeScope import Batteries.Tactic.Trans import Batteries.Tactic.Unreachable import Batteries.Tactic.Where -import Batteries.Test.Internal.DummyLabelAttr -import Batteries.Test.Internal.DummyLibraryNote -import Batteries.Test.Internal.DummyLibraryNote2 import Batteries.Util.Cache import Batteries.Util.ExtendedBinder import Batteries.Util.LibraryNote diff --git a/Batteries/Test/Internal/DummyLabelAttr.lean b/BatteriesTest/Internal/DummyLabelAttr.lean similarity index 100% rename from Batteries/Test/Internal/DummyLabelAttr.lean rename to BatteriesTest/Internal/DummyLabelAttr.lean diff --git a/Batteries/Test/Internal/DummyLibraryNote.lean b/BatteriesTest/Internal/DummyLibraryNote.lean similarity index 100% rename from Batteries/Test/Internal/DummyLibraryNote.lean rename to BatteriesTest/Internal/DummyLibraryNote.lean diff --git a/Batteries/Test/Internal/DummyLibraryNote2.lean b/BatteriesTest/Internal/DummyLibraryNote2.lean similarity index 88% rename from Batteries/Test/Internal/DummyLibraryNote2.lean rename to BatteriesTest/Internal/DummyLibraryNote2.lean index 3a5bc35dc0..0676060219 100644 --- a/Batteries/Test/Internal/DummyLibraryNote2.lean +++ b/BatteriesTest/Internal/DummyLibraryNote2.lean @@ -1,4 +1,4 @@ -import Batteries.Test.Internal.DummyLibraryNote +import BatteriesTest.Internal.DummyLibraryNote library_note "test" /-- 3: this is a note in a different file importing the above testnotes, diff --git a/test/MLList.lean b/BatteriesTest/MLList.lean similarity index 100% rename from test/MLList.lean rename to BatteriesTest/MLList.lean diff --git a/test/absurd.lean b/BatteriesTest/absurd.lean similarity index 100% rename from test/absurd.lean rename to BatteriesTest/absurd.lean diff --git a/test/alias.lean b/BatteriesTest/alias.lean similarity index 100% rename from test/alias.lean rename to BatteriesTest/alias.lean diff --git a/test/array.lean b/BatteriesTest/array.lean similarity index 100% rename from test/array.lean rename to BatteriesTest/array.lean diff --git a/test/by_contra.lean b/BatteriesTest/by_contra.lean similarity index 100% rename from test/by_contra.lean rename to BatteriesTest/by_contra.lean diff --git a/test/case.lean b/BatteriesTest/case.lean similarity index 100% rename from test/case.lean rename to BatteriesTest/case.lean diff --git a/test/congr.lean b/BatteriesTest/congr.lean similarity index 100% rename from test/congr.lean rename to BatteriesTest/congr.lean diff --git a/test/conv_equals.lean b/BatteriesTest/conv_equals.lean similarity index 100% rename from test/conv_equals.lean rename to BatteriesTest/conv_equals.lean diff --git a/test/exfalso.lean b/BatteriesTest/exfalso.lean similarity index 100% rename from test/exfalso.lean rename to BatteriesTest/exfalso.lean diff --git a/test/float.lean b/BatteriesTest/float.lean similarity index 100% rename from test/float.lean rename to BatteriesTest/float.lean diff --git a/test/help_cmd.lean b/BatteriesTest/help_cmd.lean similarity index 100% rename from test/help_cmd.lean rename to BatteriesTest/help_cmd.lean diff --git a/test/import_lean.lean b/BatteriesTest/import_lean.lean similarity index 100% rename from test/import_lean.lean rename to BatteriesTest/import_lean.lean diff --git a/test/instances.lean b/BatteriesTest/instances.lean similarity index 100% rename from test/instances.lean rename to BatteriesTest/instances.lean diff --git a/test/isIndependentOf.lean b/BatteriesTest/isIndependentOf.lean similarity index 100% rename from test/isIndependentOf.lean rename to BatteriesTest/isIndependentOf.lean diff --git a/test/kmp_matcher.lean b/BatteriesTest/kmp_matcher.lean similarity index 100% rename from test/kmp_matcher.lean rename to BatteriesTest/kmp_matcher.lean diff --git a/test/lemma_cmd.lean b/BatteriesTest/lemma_cmd.lean similarity index 100% rename from test/lemma_cmd.lean rename to BatteriesTest/lemma_cmd.lean diff --git a/test/library_note.lean b/BatteriesTest/library_note.lean similarity index 96% rename from test/library_note.lean rename to BatteriesTest/library_note.lean index f7ac514cdc..7eadf44a44 100644 --- a/test/library_note.lean +++ b/BatteriesTest/library_note.lean @@ -1,5 +1,5 @@ import Batteries.Tactic.HelpCmd -import Batteries.Test.Internal.DummyLibraryNote2 +import BatteriesTest.Internal.DummyLibraryNote2 /-- error: Note not found diff --git a/test/lintTC.lean b/BatteriesTest/lintTC.lean similarity index 100% rename from test/lintTC.lean rename to BatteriesTest/lintTC.lean diff --git a/test/lint_docBlame.lean b/BatteriesTest/lint_docBlame.lean similarity index 100% rename from test/lint_docBlame.lean rename to BatteriesTest/lint_docBlame.lean diff --git a/test/lint_docBlameThm.lean b/BatteriesTest/lint_docBlameThm.lean similarity index 100% rename from test/lint_docBlameThm.lean rename to BatteriesTest/lint_docBlameThm.lean diff --git a/test/lint_dupNamespace.lean b/BatteriesTest/lint_dupNamespace.lean similarity index 100% rename from test/lint_dupNamespace.lean rename to BatteriesTest/lint_dupNamespace.lean diff --git a/test/lint_empty.lean b/BatteriesTest/lint_empty.lean similarity index 100% rename from test/lint_empty.lean rename to BatteriesTest/lint_empty.lean diff --git a/test/lint_lean.lean b/BatteriesTest/lint_lean.lean similarity index 100% rename from test/lint_lean.lean rename to BatteriesTest/lint_lean.lean diff --git a/test/lint_simpNF.lean b/BatteriesTest/lint_simpNF.lean similarity index 100% rename from test/lint_simpNF.lean rename to BatteriesTest/lint_simpNF.lean diff --git a/test/lint_unreachableTactic.lean b/BatteriesTest/lint_unreachableTactic.lean similarity index 100% rename from test/lint_unreachableTactic.lean rename to BatteriesTest/lint_unreachableTactic.lean diff --git a/test/lintsimp.lean b/BatteriesTest/lintsimp.lean similarity index 100% rename from test/lintsimp.lean rename to BatteriesTest/lintsimp.lean diff --git a/test/lintunused.lean b/BatteriesTest/lintunused.lean similarity index 100% rename from test/lintunused.lean rename to BatteriesTest/lintunused.lean diff --git a/test/list_sublists.lean b/BatteriesTest/list_sublists.lean similarity index 100% rename from test/list_sublists.lean rename to BatteriesTest/list_sublists.lean diff --git a/test/nondet.lean b/BatteriesTest/nondet.lean similarity index 100% rename from test/nondet.lean rename to BatteriesTest/nondet.lean diff --git a/test/norm_cast.lean b/BatteriesTest/norm_cast.lean similarity index 100% rename from test/norm_cast.lean rename to BatteriesTest/norm_cast.lean diff --git a/test/omega/benchmark.lean b/BatteriesTest/omega/benchmark.lean similarity index 100% rename from test/omega/benchmark.lean rename to BatteriesTest/omega/benchmark.lean diff --git a/test/on_goal.lean b/BatteriesTest/on_goal.lean similarity index 100% rename from test/on_goal.lean rename to BatteriesTest/on_goal.lean diff --git a/test/print_prefix.lean b/BatteriesTest/print_prefix.lean similarity index 100% rename from test/print_prefix.lean rename to BatteriesTest/print_prefix.lean diff --git a/test/proof_wanted.lean b/BatteriesTest/proof_wanted.lean similarity index 100% rename from test/proof_wanted.lean rename to BatteriesTest/proof_wanted.lean diff --git a/test/register_label_attr.lean b/BatteriesTest/register_label_attr.lean similarity index 92% rename from test/register_label_attr.lean rename to BatteriesTest/register_label_attr.lean index 9add2305c6..d1928052e8 100644 --- a/test/register_label_attr.lean +++ b/BatteriesTest/register_label_attr.lean @@ -1,4 +1,4 @@ -import Batteries.Test.Internal.DummyLabelAttr +import BatteriesTest.Internal.DummyLabelAttr import Lean.LabelAttribute set_option linter.missingDocs false diff --git a/test/rfl.lean b/BatteriesTest/rfl.lean similarity index 100% rename from test/rfl.lean rename to BatteriesTest/rfl.lean diff --git a/test/seq_focus.lean b/BatteriesTest/seq_focus.lean similarity index 100% rename from test/seq_focus.lean rename to BatteriesTest/seq_focus.lean diff --git a/test/show_term.lean b/BatteriesTest/show_term.lean similarity index 100% rename from test/show_term.lean rename to BatteriesTest/show_term.lean diff --git a/test/show_unused.lean b/BatteriesTest/show_unused.lean similarity index 100% rename from test/show_unused.lean rename to BatteriesTest/show_unused.lean diff --git a/test/simp_trace.lean b/BatteriesTest/simp_trace.lean similarity index 100% rename from test/simp_trace.lean rename to BatteriesTest/simp_trace.lean diff --git a/test/simpa.lean b/BatteriesTest/simpa.lean similarity index 100% rename from test/simpa.lean rename to BatteriesTest/simpa.lean diff --git a/test/solve_by_elim.lean b/BatteriesTest/solve_by_elim.lean similarity index 97% rename from test/solve_by_elim.lean rename to BatteriesTest/solve_by_elim.lean index 94baf49ea7..6bfe72bf02 100644 --- a/test/solve_by_elim.lean +++ b/BatteriesTest/solve_by_elim.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import Batteries.Tactic.PermuteGoals -import Batteries.Test.Internal.DummyLabelAttr +import BatteriesTest.Internal.DummyLabelAttr import Lean.Meta.Tactic.Constructor import Lean.Elab.SyntheticMVars import Lean.Elab.Tactic.SolveByElim -- FIXME we need to make SolveByElimConfig builtin diff --git a/test/trans.lean b/BatteriesTest/trans.lean similarity index 100% rename from test/trans.lean rename to BatteriesTest/trans.lean diff --git a/test/tryThis.lean b/BatteriesTest/tryThis.lean similarity index 100% rename from test/tryThis.lean rename to BatteriesTest/tryThis.lean diff --git a/test/vector.lean b/BatteriesTest/vector.lean similarity index 100% rename from test/vector.lean rename to BatteriesTest/vector.lean diff --git a/test/where.lean b/BatteriesTest/where.lean similarity index 100% rename from test/where.lean rename to BatteriesTest/where.lean diff --git a/lakefile.toml b/lakefile.toml index a8a658ef9f..18abb8a007 100644 --- a/lakefile.toml +++ b/lakefile.toml @@ -1,5 +1,5 @@ name = "batteries" -testDriver = "test" +testDriver = "BatteriesTest" lintDriver = "runLinter" defaultTargets = ["Batteries", "runLinter"] @@ -9,6 +9,11 @@ linter.missingDocs = true [[lean_lib]] name = "Batteries" +[[lean_lib]] +name = "BatteriesTest" +globs = ["BatteriesTest.+"] +leanOptions = {linter.missingDocs = false} + [[lean_exe]] name = "runLinter" srcDir = "scripts" From 430036e4a850a6423bba53054c1f07cf31a08437 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sun, 10 Nov 2024 18:44:41 -0500 Subject: [PATCH 146/177] refactor: cleanup `Vector` API (#1025) --- Batteries/Data/Vector/Basic.lean | 365 +++++++++++++++--------------- Batteries/Data/Vector/Lemmas.lean | 18 +- 2 files changed, 184 insertions(+), 199 deletions(-) diff --git a/Batteries/Data/Vector/Basic.lean b/Batteries/Data/Vector/Basic.lean index 2d4708d5b2..0ae5871356 100644 --- a/Batteries/Data/Vector/Basic.lean +++ b/Batteries/Data/Vector/Basic.lean @@ -7,28 +7,29 @@ Authors: Shreyas Srinivas, François G. Dorais import Batteries.Data.Array import Batteries.Data.List.Basic import Batteries.Data.List.Lemmas +import Batteries.Tactic.Alias import Batteries.Tactic.Lint.Misc /-! -## Vectors -`Vector α n` is an array with a statically fixed size `n`. -It combines the benefits of Lean's special support for `Arrays` -that offer `O(1)` accesses and in-place mutations for arrays with no more than one reference, -with static guarantees about the size of the underlying array. +# Vectors + +`Vector α n` is a thin wrapper around `Array α` for arrays of fixed size `n`. -/ namespace Batteries -/-- `Vector α n` is an `Array α` whose size is statically fixed to `n` -/ -structure Vector (α : Type u) (n : Nat) where - /-- Internally, a vector is stored as an array for fast access -/ - toArray : Array α - /-- `size_eq` fixes the size of `toArray` statically -/ - size_eq : toArray.size = n +/-- `Vector α n` is an `Array α` with size `n`. -/ +structure Vector (α : Type u) (n : Nat) extends Array α where + /-- Array size. -/ + size_toArray : toArray.size = n deriving Repr, DecidableEq +attribute [simp] Vector.size_toArray + namespace Vector +@[deprecated (since := "2024-10-15")] alias size_eq := size_toArray + /-- Syntax for `Vector α n` -/ syntax "#v[" withoutPosition(sepBy(term, ", ")) "]" : term @@ -38,283 +39,271 @@ macro_rules /-- Custom eliminator for `Vector α n` through `Array α` -/ @[elab_as_elim] -def elimAsArray {motive : ∀ {n}, Vector α n → Sort u} (mk : ∀ a : Array α, motive ⟨a, rfl⟩) : - {n : Nat} → (v : Vector α n) → motive v - | _, ⟨a, rfl⟩ => mk a +def elimAsArray {motive : Vector α n → Sort u} + (mk : ∀ (a : Array α) (ha : a.size = n), motive ⟨a, ha⟩) : + (v : Vector α n) → motive v + | ⟨a, ha⟩ => mk a ha /-- Custom eliminator for `Vector α n` through `List α` -/ @[elab_as_elim] -def elimAsList {motive : ∀ {n}, Vector α n → Sort u} (mk : ∀ a : List α, motive ⟨⟨a⟩, rfl⟩) : - {n : Nat} → (v : Vector α n) → motive v - | _, ⟨⟨a⟩, rfl⟩ => mk a +def elimAsList {motive : Vector α n → Sort u} + (mk : ∀ (a : List α) (ha : a.length = n), motive ⟨⟨a⟩, ha⟩) : + (v : Vector α n) → motive v + | ⟨⟨a⟩, ha⟩ => mk a ha -/-- `Vector.size` gives the size of a vector. -/ -@[nolint unusedArguments] -def size (_ : Vector α n) : Nat := n +/-- The empty vector. -/ +@[inline] def empty : Vector α 0 := ⟨.empty, rfl⟩ -/-- `Vector.empty` produces an empty vector -/ -def empty : Vector α 0 := ⟨Array.empty, rfl⟩ +/-- Make an empty vector with pre-allocated capacity. -/ +@[inline] def mkEmpty (capacity : Nat) : Vector α 0 := ⟨.mkEmpty capacity, rfl⟩ -/-- Make an empty vector with pre-allocated capacity-/ -def mkEmpty (capacity : Nat) : Vector α 0 := ⟨Array.mkEmpty capacity, rfl⟩ +/-- Makes a vector of size `n` with all cells containing `v`. -/ +@[inline] def mkVector (n) (v : α) : Vector α n := ⟨mkArray n v, by simp⟩ -/-- Makes a vector of size `n` with all cells containing `v` -/ -def mkVector (n : Nat) (v : α) : Vector α n := ⟨mkArray n v, Array.size_mkArray ..⟩ +/-- Returns a vector of size `1` with element `v`. -/ +@[inline] def singleton (v : α) : Vector α 1 := ⟨#[v], rfl⟩ -/-- Returns a vector of size `1` with a single element `v` -/ -def singleton (v : α) : Vector α 1 := mkVector 1 v - -/-- -The Inhabited instance for `Vector α n` with `[Inhabited α]` produces a vector of size `n` -with all its elements equal to the `default` element of type `α` --/ instance [Inhabited α] : Inhabited (Vector α n) where default := mkVector n default -/-- The list obtained from a vector. -/ -def toList (v : Vector α n) : List α := v.toArray.toList +/-- Get an element of a vector using a `Fin` index. -/ +@[inline] def get (v : Vector α n) (i : Fin n) : α := + v.toArray.get (i.cast v.size_toArray.symm) -/-- nth element of a vector, indexed by a `Fin` type. -/ -def get (v : Vector α n) (i : Fin n) : α := v.toArray.get <| i.cast v.size_eq.symm +/-- Get an element of a vector using a `USize` index and a proof that the index is within bounds. -/ +@[inline] def uget (v : Vector α n) (i : USize) (h : i.toNat < n) : α := + v.toArray.uget i (v.size_toArray.symm ▸ h) -/-- Vector lookup function that takes an index `i` of type `USize` -/ -def uget (v : Vector α n) (i : USize) (h : i.toNat < n) : α := v.toArray.uget i (v.size_eq.symm ▸ h) - -/-- `Vector α n` instance for the `GetElem` typeclass. -/ instance : GetElem (Vector α n) Nat α fun _ i => i < n where - getElem := fun x i h => get x ⟨i, h⟩ + getElem x i h := get x ⟨i, h⟩ /-- -`getD v i v₀` gets the `iᵗʰ` element of v if valid. -Otherwise it returns `v₀` by default +Get an element of a vector using a `Nat` index. Returns the given default value if the index is out +of bounds. -/ -def getD (v : Vector α n) (i : Nat) (v₀ : α) : α := Array.getD v.toArray i v₀ +@[inline] def getD (v : Vector α n) (i : Nat) (default : α) : α := v.toArray.getD i default -/-- -`v.back! v` gets the last element of the vector. -panics if `v` is empty. --/ -abbrev back! [Inhabited α] (v : Vector α n) : α := v[n - 1]! +/-- The last element of a vector. Panics if the vector is empty. -/ +@[inline] def back! [Inhabited α] (v : Vector α n) : α := v.toArray.back! -/-- -`v.back?` gets the last element `x` of the array as `some x` -if it exists. Else the vector is empty and it returns `none` --/ -abbrev back? (v : Vector α n) : Option α := v[n - 1]? +/-- The last element of a vector, or `none` if the array is empty. -/ +@[inline] def back? (v : Vector α n) : Option α := v.toArray.back? -/-- Abbreviation for the last element of a non-empty `Vector`.-/ -abbrev back (v : Vector α (n + 1)) : α := v[n] +/-- The last element of a non-empty vector. -/ +@[inline] def back [NeZero n] (v : Vector α n) : α := + -- TODO: change to just `v[n]` + have : Inhabited α := ⟨v[0]'(Nat.pos_of_neZero n)⟩ + v.back! -/-- `Vector.head` produces the head of a vector -/ -abbrev head (v : Vector α (n+1)) := v[0] +/-- The first element of a non-empty vector. -/ +@[inline] def head [NeZero n] (v : Vector α n) := v[0]'(Nat.pos_of_neZero n) -/-- `push v x` pushes `x` to the end of vector `v` in O(1) time -/ -def push (x : α) (v : Vector α n) : Vector α (n + 1) := - ⟨v.toArray.push x, by simp [v.size_eq]⟩ +/-- Push an element `x` to the end of a vector. -/ +@[inline] def push (x : α) (v : Vector α n) : Vector α (n + 1) := + ⟨v.toArray.push x, by simp⟩ -/-- `pop v` returns the vector with the last element removed -/ -def pop (v : Vector α n) : Vector α (n - 1) := - ⟨Array.pop v.toArray, by simp [v.size_eq]⟩ +/-- Remove the last element of a vector. -/ +@[inline] def pop (v : Vector α n) : Vector α (n - 1) := + ⟨Array.pop v.toArray, by simp⟩ /-- -Sets an element in a vector using a Fin index. +Set an element in a vector using a `Fin` index. -This will perform the update destructively provided that a has a reference count of 1 when called. +This will perform the update destructively provided that the vector has a reference count of 1. -/ -def set (v : Vector α n) (i : Fin n) (x : α) : Vector α n := - ⟨v.toArray.set (Fin.cast v.size_eq.symm i) x, by simp [v.size_eq]⟩ +@[inline] def set (v : Vector α n) (i : Fin n) (x : α) : Vector α n := + ⟨v.toArray.set (i.cast v.size_toArray.symm) x, by simp⟩ /-- -`setN v i h x` sets an element in a vector using a Nat index which is provably valid. -By default a proof by `get_elem_tactic` is provided. +Set an element in a vector using a `Nat` index. By default, the `get_elem_tactic` is used to +synthesize a proof that the index is within bounds. -This will perform the update destructively provided that a has a reference count of 1 when called. +This will perform the update destructively provided that the vector has a reference count of 1. -/ -def setN (v : Vector α n) (i : Nat) (x : α) (h : i < n := by get_elem_tactic) : Vector α n := - v.set ⟨i, h⟩ x +@[inline] def setN (v : Vector α n) (i : Nat) (x : α) (h : i < n := by get_elem_tactic) : + Vector α n := ⟨v.toArray.setN i x (by simp_all), by simp⟩ /-- -Sets an element in a vector, or do nothing if the index is out of bounds. +Set an element in a vector using a `Nat` index. Returns the vector unchanged if the index is out of +bounds. -This will perform the update destructively provided that a has a reference count of 1 when called. +This will perform the update destructively provided that the vector has a reference count of 1. -/ -def setD (v : Vector α n) (i : Nat) (x : α) : Vector α n := - ⟨v.toArray.setD i x, by simp [v.size_eq]⟩ +@[inline] def setD (v : Vector α n) (i : Nat) (x : α) : Vector α n := + ⟨v.toArray.setD i x, by simp⟩ /-- -Sets an element in an array, or panic if the index is out of bounds. +Set an element in a vector using a `Nat` index. Panics if the index is out of bounds. -This will perform the update destructively provided that a has a reference count of 1 when called. +This will perform the update destructively provided that the vector has a reference count of 1. -/ -def set! (v : Vector α n) (i : Nat) (x : α) : Vector α n := - ⟨v.toArray.set! i x, by simp [v.size_eq]⟩ +@[inline] def set! (v : Vector α n) (i : Nat) (x : α) : Vector α n := + ⟨v.toArray.set! i x, by simp⟩ -/-- Appends a vector to another. -/ -def append : Vector α n → Vector α m → Vector α (n + m) - | ⟨a₁, _⟩, ⟨a₂, _⟩ => ⟨a₁ ++ a₂, by simp [Array.size_append, *]⟩ +/-- Append two vectors. -/ +@[inline] def append (v : Vector α n) (w : Vector α m) : Vector α (n + m) := + ⟨v.toArray ++ w.toArray, by simp⟩ instance : HAppend (Vector α n) (Vector α m) (Vector α (n + m)) where hAppend := append /-- Creates a vector from another with a provably equal length. -/ -protected def cast {n m : Nat} (h : n = m) : Vector α n → Vector α m - | ⟨x, p⟩ => ⟨x, h ▸ p⟩ +@[inline] protected def cast (h : n = m) (v : Vector α n) : Vector α m := + ⟨v.toArray, by simp [*]⟩ /-- -`extract v start halt` Returns the slice of `v` from indices `start` to `stop` (exclusive). -If `start` is greater or equal to `stop`, the result is empty. -If `stop` is greater than the size of `v`, the size is used instead. +Extracts the slice of a vector from indices `start` to `stop` (exclusive). If `start ≥ stop`, the +result is empty. If `stop` is greater than the size of the vector, the size is used instead. -/ -def extract (v : Vector α n) (start stop : Nat) : Vector α (min stop n - start) := - ⟨Array.extract v.toArray start stop, by simp [v.size_eq]⟩ +@[inline] def extract (v : Vector α n) (start stop : Nat) : Vector α (min stop n - start) := + ⟨v.toArray.extract start stop, by simp⟩ -/-- Maps a vector under a function. -/ -def map (f : α → β) (v : Vector α n) : Vector β n := - ⟨v.toArray.map f, by simp [v.size_eq]⟩ +/-- Maps elements of a vector using the function `f`. -/ +@[inline] def map (f : α → β) (v : Vector α n) : Vector β n := + ⟨v.toArray.map f, by simp⟩ -/-- Maps two vectors under a curried function of two variables. -/ -def zipWith : (a : Vector α n) → (b : Vector β n) → (f : α → β → φ) → Vector φ n - | ⟨a, h₁⟩, ⟨b, h₂⟩, f => ⟨Array.zipWith a b f, by simp [Array.size_zipWith, h₁, h₂]⟩ +/-- Maps corresponding elements of two vectors of equal size using the function `f`. -/ +@[inline] def zipWith (a : Vector α n) (b : Vector β n) (f : α → β → φ) : Vector φ n := + ⟨Array.zipWith a.toArray b.toArray f, by simp⟩ -/-- Returns a vector of length `n` from a function on `Fin n`. -/ -def ofFn (f : Fin n → α) : Vector α n := ⟨Array.ofFn f, Array.size_ofFn ..⟩ +/-- The vector of length `n` whose `i`-th element is `f i`. -/ +@[inline] def ofFn (f : Fin n → α) : Vector α n := + ⟨Array.ofFn f, by simp⟩ /-- -Swaps two entries in a Vector. +Swap two elements of a vector using `Fin` indices. -This will perform the update destructively provided that `v` has a reference count of 1 when called. +This will perform the update destructively provided that the vector has a reference count of 1. -/ -def swap (v : Vector α n) (i j : Fin n) : Vector α n := - ⟨v.toArray.swap (Fin.cast v.size_eq.symm i) (Fin.cast v.size_eq.symm j), by simp [v.size_eq]⟩ +@[inline] def swap (v : Vector α n) (i j : Fin n) : Vector α n := + ⟨v.toArray.swap (Fin.cast v.size_toArray.symm i) (Fin.cast v.size_toArray.symm j), by simp⟩ /-- -`swapN v i j hi hj` swaps two `Nat` indexed entries in a `Vector α n`. -Uses `get_elem_tactic` to supply proofs `hi` and `hj` respectively -that the indices `i` and `j` are in range. +Swap two elements of a vector using `Nat` indices. By default, the `get_elem_tactic` is used to +synthesize proofs that the indices are within bounds. -This will perform the update destructively provided that `v` has a reference count of 1 when called. +This will perform the update destructively provided that the vector has a reference count of 1. -/ -def swapN (v : Vector α n) (i j : Nat) +@[inline] def swapN (v : Vector α n) (i j : Nat) (hi : i < n := by get_elem_tactic) (hj : j < n := by get_elem_tactic) : Vector α n := - v.swap ⟨i, hi⟩ ⟨j, hj⟩ + ⟨v.toArray.swapN i j (by simp_all) (by simp_all), by simp⟩ /-- -Swaps two entries in a `Vector α n`, or panics if either index is out of bounds. +Swap two elements of a vector using `Nat` indices. Panics if either index is out of bounds. -This will perform the update destructively provided that `v` has a reference count of 1 when called. +This will perform the update destructively provided that the vector has a reference count of 1. -/ -def swap! (v : Vector α n) (i j : Nat) : Vector α n := - ⟨Array.swap! v.toArray i j, by simp [v.size_eq]⟩ +@[inline] def swap! (v : Vector α n) (i j : Nat) : Vector α n := + ⟨v.toArray.swap! i j, by simp⟩ /-- -Swaps the entry with index `i : Fin n` in the vector for a new entry. -The old entry is returned with the modified vector. +Swaps an element of a vector with a given value using a `Fin` index. The original value is returned +along with the updated vector. + +This will perform the update destructively provided that the vector has a reference count of 1. -/ -def swapAt (v : Vector α n) (i : Fin n) (x : α) : α × Vector α n := - let res := v.toArray.swapAt (Fin.cast v.size_eq.symm i) x - (res.1, ⟨res.2, by simp [Array.swapAt_def, res, v.size_eq]⟩) +@[inline] def swapAt (v : Vector α n) (i : Fin n) (x : α) : α × Vector α n := + let a := v.toArray.swapAt (Fin.cast v.size_toArray.symm i) x + ⟨a.fst, a.snd, by simp [a]⟩ /-- -Swaps the entry with index `i : Nat` in the vector for a new entry `x`. -The old entry is returned alongwith the modified vector. +Swaps an element of a vector with a given value using a `Nat` index. By default, the +`get_elem_tactic` is used to synthesise a proof that the index is within bounds. The original value +is returned along with the updated vector. -Automatically generates a proof of `i < n` with `get_elem_tactic` where feasible. +This will perform the update destructively provided that the vector has a reference count of 1. -/ -def swapAtN (v : Vector α n) (i : Nat) (x : α) (h : i < n := by get_elem_tactic) : α × Vector α n := - swapAt v ⟨i, h⟩ x +@[inline] def swapAtN (v : Vector α n) (i : Nat) (x : α) (h : i < n := by get_elem_tactic) : + α × Vector α n := + let a := v.toArray.swapAtN i x (by simp_all) + ⟨a.fst, a.snd, by simp [a]⟩ /-- -`swapAt! v i x` swaps out the entry at index `i : Nat` in the vector for `x`, if the index is valid. -Otherwise it panics The old entry is returned with the modified vector. +Swaps an element of a vector with a given value using a `Nat` index. Panics if the index is out of +bounds. The original value is returned along with the updated vector. + +This will perform the update destructively provided that the vector has a reference count of 1. -/ @[inline] def swapAt! (v : Vector α n) (i : Nat) (x : α) : α × Vector α n := - if h : i < n then - swapAt v ⟨i, h⟩ x - else - have : Inhabited α := ⟨x⟩ - panic! s!"Index {i} is out of bounds" + let a := v.toArray.swapAt! i x + ⟨a.fst, a.snd, by simp [a]⟩ -/-- `range n` returns the vector `#v[0,1,2,...,n-1]`. -/ -def range (n : Nat) : Vector Nat n := ⟨Array.range n, Array.size_range ..⟩ +/-- The vector `#v[0,1,2,...,n-1]`. -/ +@[inline] def range (n : Nat) : Vector Nat n := ⟨Array.range n, by simp⟩ /-- -`take v m` shrinks the vector to the first `m` elements if `m < n`. -Returns `v` unchanged if `m ≥ n`. +Extract the first `m` elements of a vector. If `m` is greater than or equal to the size of the +vector then the vector is returned unchanged. -/ -def take (v : Vector α n) (m : Nat) : Vector α (min n m) := - ⟨v.toArray.take m, by simp [Array.size_take, v.size_eq, Nat.min_comm]⟩ +@[inline] def take (v : Vector α n) (m : Nat) : Vector α (min m n) := + ⟨v.toArray.take m, by simp⟩ @[deprecated (since := "2024-10-22")] alias shrink := take /-- -Drops the first (up to) `i` elements from a vector of length `n`. -If `m ≥ n`, the return value is empty. +Deletes the first `m` elements of a vector. If `m` is greater than or equal to the size of the +vector then the empty vector is returned. -/ -def drop (i : Nat) (v : Vector α n) : Vector α (n - i) := - have : min n n - i = n - i := by rw [Nat.min_self] - Vector.cast this (extract v i n) +@[inline] def drop (v : Vector α n) (m : Nat) : Vector α (n - m) := + ⟨v.toArray.extract m v.size, by simp⟩ /-- -`isEqv` takes a given boolean property `p`. It returns `true` -if and only if `p a[i] b[i]` holds true for all valid indices `i`. +Compares two vectors of the same size using a given boolean relation `r`. `isEqv v w r` returns +`true` if and only if `r v[i] w[i]` is true for all indices `i`. -/ -@[inline] def isEqv (a b : Vector α n) (p : α → α → Bool) : Bool := - Array.isEqvAux a.toArray b.toArray (a.size_eq.trans b.size_eq.symm) p 0 (Nat.zero_le _) - -instance [BEq α] : BEq (Vector α n) := - ⟨fun a b => isEqv a b BEq.beq⟩ +@[inline] def isEqv (v w : Vector α n) (r : α → α → Bool) : Bool := + Array.isEqvAux v.toArray w.toArray (by simp) r 0 (by simp) -proof_wanted lawfulBEq [BEq α] [LawfulBEq α] : LawfulBEq (Vector α n) +instance [BEq α] : BEq (Vector α n) where + beq a b := isEqv a b (· == ·) -/-- `reverse v` reverses the vector `v`. -/ -def reverse (v : Vector α n) : Vector α n := - ⟨v.toArray.reverse, by simp [v.size_eq]⟩ +proof_wanted instLawfulBEq (α n) [BEq α] [LawfulBEq α] : LawfulBEq (Vector α n) -/-- `O(|v| - i)`. `feraseIdx v i` removes the element at position `i` in vector `v`. -/ -def feraseIdx (v : Vector α n) (i : Fin n) : Vector α (n-1) := - ⟨v.toArray.feraseIdx (Fin.cast v.size_eq.symm i), by simp [Array.size_feraseIdx, v.size_eq]⟩ +/-- Reverse the elements of a vector. -/ +@[inline] def reverse (v : Vector α n) : Vector α n := + ⟨v.toArray.reverse, by simp⟩ -/-- `Vector.tail` produces the tail of the vector `v`. -/ -@[inline] def tail (v : Vector α n) : Vector α (n-1) := - match n with - | 0 => v - | _ + 1 => Vector.feraseIdx v 0 +/-- Delete an element of a vector using a `Fin` index. -/ +@[inline] def feraseIdx (v : Vector α n) (i : Fin n) : Vector α (n-1) := + ⟨v.toArray.feraseIdx (Fin.cast v.size_toArray.symm i), by simp [Array.size_feraseIdx]⟩ -/-- -`O(|v| - i)`. `eraseIdx! v i` removes the element at position `i` from vector `v` of length `n`. -Panics if `i` is not less than `n`. --/ +/-- Delete an element of a vector using a `Nat` index. Panics if the index is out of bounds. -/ @[inline] def eraseIdx! (v : Vector α n) (i : Nat) : Vector α (n-1) := - if h : i < n then - feraseIdx v ⟨i,h⟩ + if _ : i < n then + ⟨v.toArray.eraseIdx i, by simp [*]⟩ else - have : Inhabited (Vector α (n-1)) := ⟨v.tail⟩ - panic! s!"Index {i} is out of bounds" + have : Inhabited (Vector α (n-1)) := ⟨v.pop⟩ + panic! "index out of bounds" -/-- -`eraseIdxN v i h` removes the element at position `i` from a vector of length `n`. -`h : i < n` has a default argument `by get_elem_tactic` which tries to supply a proof -that the index is valid. +/-- Delete the first element of a vector. Returns the empty vector if the input vector is empty. -/ +@[inline] def tail (v : Vector α n) : Vector α (n-1) := + if _ : 0 < n then + ⟨v.toArray.eraseIdx 0, by simp [*]⟩ + else + v.cast (by omega) -This function takes worst case O(n) time because it has to backshift all elements at positions -greater than i. +/-- +Delete an element of a vector using a `Nat` index. By default, the `get_elem_tactic` is used to +synthesise a proof that the index is within bounds. -/ -abbrev eraseIdxN (v : Vector α n) (i : Nat) (h : i < n := by get_elem_tactic) : Vector α (n - 1) := - v.feraseIdx ⟨i, h⟩ +@[inline] def eraseIdxN (v : Vector α n) (i : Nat) (h : i < n := by get_elem_tactic) : + Vector α (n - 1) := ⟨v.toArray.eraseIdxN i (by simp [*]), by simp⟩ /-- -If `x` is an element of vector `v` at index `j`, then `indexOf? v x` returns `some j`. -Otherwise it returns `none`. +Finds the first index of a given value in a vector using `==` for comparison. Returns `none` if the +no element of the index matches the given value. -/ -def indexOf? [BEq α] (v : Vector α n) (x : α) : Option (Fin n) := - match Array.indexOf? v.toArray x with - | some res => some (Fin.cast v.size_eq res) +@[inline] def indexOf? [BEq α] (v : Vector α n) (x : α) : Option (Fin n) := + match v.toArray.indexOf? x with + | some res => some (res.cast v.size_toArray) | none => none -/-- `isPrefixOf as bs` returns true iff vector `as` is a prefix of vector`bs` -/ -def isPrefixOf [BEq α] (as : Vector α m) (bs : Vector α n) : Bool := - Array.isPrefixOf as.toArray bs.toArray +/-- Returns `true` when `v` is a prefix of the vector `w`. -/ +@[inline] def isPrefixOf [BEq α] (v : Vector α m) (w : Vector α n) : Bool := + v.toArray.isPrefixOf w.toArray -/-- `allDiff as i` returns `true` when all elements of `v` are distinct from each other` -/ -def allDiff [BEq α] (as : Vector α n) : Bool := - Array.allDiff as.toArray +/-- +Returns `true` when all elements of the vector are pairwise distinct using `==` for comparison. +-/ +@[inline] def allDiff [BEq α] (as : Vector α n) : Bool := + as.toArray.allDiff diff --git a/Batteries/Data/Vector/Lemmas.lean b/Batteries/Data/Vector/Lemmas.lean index d40c7b79cb..6271f39f02 100644 --- a/Batteries/Data/Vector/Lemmas.lean +++ b/Batteries/Data/Vector/Lemmas.lean @@ -19,10 +19,7 @@ namespace Batteries namespace Vector -attribute [simp] Vector.size_eq - -@[simp] theorem length_toList {α n} (xs : Vector α n) : xs.toList.length = n := - xs.size_eq +theorem length_toList {α n} (xs : Vector α n) : xs.toList.length = n := by simp @[simp] theorem getElem_mk {data : Array α} {size : data.size = n} {i : Nat} (h : i < n) : (Vector.mk data size)[i] = data[i] := rfl @@ -32,9 +29,8 @@ attribute [simp] Vector.size_eq cases xs simp -@[simp] theorem getElem_toList {α n} (xs : Vector α n) (i : Nat) (h : i < xs.toList.length) : - xs.toList[i] = xs[i]'(by simpa using h) := by - simp [toList] +theorem getElem_toList {α n} (xs : Vector α n) (i : Nat) (h : i < xs.toList.length) : + xs.toList[i] = xs[i]'(by simpa using h) := by simp @[simp] theorem getElem_ofFn {α n} (f : Fin n → α) (i : Nat) (h : i < n) : (Vector.ofFn f)[i] = f ⟨i, by simpa using h⟩ := by @@ -62,9 +58,9 @@ Vectors `a` and `b` are equal to each other if their elements are equal for each protected theorem ext {a b : Vector α n} (h : (i : Nat) → (_ : i < n) → a[i] = b[i]) : a = b := by apply Vector.toArray_injective apply Array.ext - · rw [a.size_eq, b.size_eq] + · rw [a.size_toArray, b.size_toArray] · intro i hi _ - rw [a.size_eq] at hi + rw [a.size_toArray] at hi exact h i hi @[simp] theorem push_mk {data : Array α} {size : data.size = n} {x : α} : @@ -100,9 +96,9 @@ defeq issues in the implicit size argument. ext i by_cases h : i < n · simp [h] - · replace h : i = n := by omega + · replace h : i = v.size - 1 := by rw [size_toArray]; omega subst h - simp + simp [pop, back, back!, ← Array.eq_push_pop_back!_of_size_ne_zero] /-! ### Decidable quantifiers. -/ From 5a0bc0da941959c867e7fb6c815bcf034ce4201a Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Sun, 10 Nov 2024 15:45:35 -0800 Subject: [PATCH 147/177] feat: improve induction code action (#1030) --- Batteries/CodeAction/Misc.lean | 108 ++++++++++++++++++++++----------- 1 file changed, 72 insertions(+), 36 deletions(-) diff --git a/Batteries/CodeAction/Misc.lean b/Batteries/CodeAction/Misc.lean index 4ad0a8099d..1ff2505a76 100644 --- a/Batteries/CodeAction/Misc.lean +++ b/Batteries/CodeAction/Misc.lean @@ -3,11 +3,9 @@ Copyright (c) 2023 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Lean.Elab.BuiltinTerm -import Lean.Elab.BuiltinNotation +import Lean.Elab.Tactic.Induction import Batteries.Lean.Position import Batteries.CodeAction.Attr -import Lean.Meta.Tactic.TryThis import Lean.Server.CodeActions.Provider /-! @@ -236,28 +234,39 @@ def removeAfterDoneAction : TacticCodeAction := fun _ _ _ stk node => do pure #[{ eager }] /-- -Similar to `getElabInfo`, but returns the names of binders instead of just the numbers; +Similar to `getElimExprInfo`, but returns the names of binders instead of just the numbers; intended for code actions which need to name the binders. -/ -def getElimNames (inductName declName : Name) : MetaM (Array (Name × Array Name)) := do - let inductVal ← getConstInfoInduct inductName - let decl ← getConstInfo declName - forallTelescopeReducing decl.type fun xs type => do +def getElimExprNames (elimType : Expr) : MetaM (Array (Name × Array Name)) := do + -- let inductVal ← getConstInfoInduct inductName + -- let decl ← getConstInfo declName + forallTelescopeReducing elimType fun xs type => do let motive := type.getAppFn let targets := type.getAppArgs + let motiveType ← inferType motive let mut altsInfo := #[] - for i in [inductVal.numParams:xs.size] do - let x := xs[i]! + for _h : i in [:xs.size] do + let x := xs[i] if x != motive && !targets.contains x then let xDecl ← x.fvarId!.getDecl - let args ← forallTelescopeReducing xDecl.type fun args _ => do - let lctx ← getLCtx - pure <| args.filterMap fun y => - let yDecl := (lctx.find? y.fvarId!).get! - if yDecl.binderInfo.isExplicit then some yDecl.userName else none - altsInfo := altsInfo.push (xDecl.userName, args) + if xDecl.binderInfo.isExplicit then + let args ← forallTelescopeReducing xDecl.type fun args _ => do + let lctx ← getLCtx + pure <| args.filterMap fun y => + let yDecl := (lctx.find? y.fvarId!).get! + if yDecl.binderInfo.isExplicit then some yDecl.userName else none + altsInfo := altsInfo.push (xDecl.userName, args) pure altsInfo +/-- Finds the `TermInfo` for an elaborated term `stx`. -/ +def findTermInfo? (node : InfoTree) (stx : Term) : Option TermInfo := + match node.findInfo? fun + | .ofTermInfo i => i.stx.getKind == stx.raw.getKind && i.stx.getRange? == stx.raw.getRange? + | _ => false + with + | some (.ofTermInfo info) => pure info + | _ => none + /-- Invoking tactic code action "Generate an explicit pattern match for 'induction'" in the following: @@ -278,18 +287,39 @@ It also works for `cases`. @[tactic_code_action Parser.Tactic.cases Parser.Tactic.induction] def casesExpand : TacticCodeAction := fun _ snap ctx _ node => do let .node (.ofTacticInfo info) _ := node | return #[] - let (discr, induction, alts) ← match info.stx with - | `(tactic| cases $[$_ :]? $e $[$alts:inductionAlts]?) => pure (e, false, alts) - | `(tactic| induction $e $[generalizing $_*]? $[$alts:inductionAlts]?) => pure (e, true, alts) + let (targets, induction, using_, alts) ← match info.stx with + | `(tactic| cases $[$[$_ :]? $targets],* $[using $u]? $(alts)?) => + pure (targets, false, u, alts) + | `(tactic| induction $[$targets],* $[using $u]? $[generalizing $_*]? $(alts)?) => + pure (targets, true, u, alts) | _ => return #[] - if let some alts := alts then - -- this detects the incomplete syntax `cases e with` - unless alts.raw[2][0][0][0][0].isMissing do return #[] - let some (.ofTermInfo discrInfo) := node.findInfo? fun i => - i.stx.getKind == discr.raw.getKind && i.stx.getRange? == discr.raw.getRange? - | return #[] - let .const name _ := (← discrInfo.runMetaM ctx (do whnf (← inferType discrInfo.expr))).getAppFn + let some discrInfos := targets.mapM (findTermInfo? node) | return #[] + let some discr₀ := discrInfos[0]? | return #[] + let mut some ctors ← discr₀.runMetaM ctx do + let targets := discrInfos.map (·.expr) + match using_ with + | none => + if Tactic.tactic.customEliminators.get (← getOptions) then + if let some elimName ← getCustomEliminator? targets induction then + return some (← getElimExprNames (← getConstInfo elimName).type) + matchConstInduct (← whnf (← inferType discr₀.expr)).getAppFn + (fun _ => failure) fun val _ => do + let elimName := if induction then mkRecName val.name else mkCasesOnName val.name + return some (← getElimExprNames (← getConstInfo elimName).type) + | some u => + let some info := findTermInfo? node u | return none + return some (← getElimExprNames (← inferType info.expr)) | return #[] + let mut fallback := none + if let some alts := alts then + if let `(Parser.Tactic.inductionAlts| with $(_)? $alts*) := alts then + for alt in alts do + match alt with + | `(Parser.Tactic.inductionAlt| | _ $_* => $fb) => fallback := fb.raw.getRange? + | `(Parser.Tactic.inductionAlt| | $id:ident $_* => $_) => + ctors := ctors.filter (fun x => x.1 != id.getId) + | _ => pure () + if ctors.isEmpty then return #[] let tacName := info.stx.getKind.updatePrefix .anonymous let eager := { title := s!"Generate an explicit pattern match for '{tacName}'." @@ -301,15 +331,21 @@ def casesExpand : TacticCodeAction := fun _ snap ctx _ node => do lazy? := some do let tacPos := info.stx.getPos?.get! let endPos := doc.meta.text.utf8PosToLspPos info.stx.getTailPos?.get! - let startPos := if alts.isSome then - let stx' := info.stx.setArg (if induction then 4 else 3) mkNullNode - doc.meta.text.utf8PosToLspPos stx'.getTailPos?.get! - else endPos - let elimName := if induction then mkRecName name else mkCasesOnName name - let ctors ← discrInfo.runMetaM ctx (getElimNames name elimName) - let newText := if ctors.isEmpty then "" else Id.run do - let mut str := " with" - let indent := "\n".pushn ' ' (findIndentAndIsStart doc.meta.text.source tacPos).1 + let indent := "\n".pushn ' ' (findIndentAndIsStart doc.meta.text.source tacPos).1 + let (startPos, str') := if alts.isSome then + let stx' := if fallback.isSome then + info.stx.modifyArg (if induction then 4 else 3) + (·.modifyArg 0 (·.modifyArg 2 (·.modifyArgs (·.filter fun s => + !(s matches `(Parser.Tactic.inductionAlt| | _ $_* => $_)))))) + else info.stx + (doc.meta.text.utf8PosToLspPos stx'.getTailPos?.get!, "") + else (endPos, " with") + let fallback := if let some ⟨startPos, endPos⟩ := fallback then + doc.meta.text.source.extract startPos endPos + else + "sorry" + let newText := Id.run do + let mut str := str' for (name, args) in ctors do let mut ctor := toString name if let some _ := (Parser.getTokenTable snap.env).find? ctor then @@ -324,7 +360,7 @@ def casesExpand : TacticCodeAction := fun _ snap ctx _ node => do else args for arg in args do str := str ++ if arg.hasNum || arg.isInternal then " _" else s!" {arg}" - str := str ++ s!" => sorry" + str := str ++ s!" => " ++ fallback str pure { eager with edit? := some <|.ofTextEdit doc.versionedIdentifier { From f73ecd1c74373f4e9a5f90821845058bc3f82e4f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sun, 10 Nov 2024 18:46:20 -0500 Subject: [PATCH 148/177] chore: update README with new lakefile syntax (#1031) --- README.md | 17 +++++++++++------ 1 file changed, 11 insertions(+), 6 deletions(-) diff --git a/README.md b/README.md index ab8606d854..f8cc2415ac 100644 --- a/README.md +++ b/README.md @@ -5,9 +5,15 @@ The "batteries included" extended library for Lean 4. This is a collection of da # Using `batteries` To use `batteries` in your project, add the following to your `lakefile.lean`: - ```lean -require batteries from git "https://github.com/leanprover-community/batteries" @ "main" +require "leanprover-community" / "batteries" @ "main" +``` +Or add the following to your `lakefile.toml`: +```toml +[[require]] +name = "batteries" +scope = "leanprover-community" +version = "main" ``` Additionally, please make sure that you're using the version of Lean that the current version of `batteries` expects. The easiest way to do this is to copy the [`lean-toolchain`](./lean-toolchain) file from this repository to your project. Once you've added the dependency declaration, the command `lake update` checks out the current version of `batteries` and writes it the Lake manifest file. Don't run this command again unless you're prepared to potentially also update your Lean compiler version, as it will retrieve the latest version of dependencies and add them to the manifest. @@ -15,19 +21,18 @@ Additionally, please make sure that you're using the version of Lean that the cu # Build instructions * Get the newest version of `elan`. If you already have installed a version of Lean, you can run - ``` + ```sh elan self update ``` If the above command fails, or if you need to install `elan`, run - ``` + ```sh curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh ``` If this also fails, follow the instructions under `Regular install` [here](https://leanprover-community.github.io/get_started.html). * To build `batteries` run `lake build`. * To build and run all tests, run `lake test`. * To run the environment linter, run `lake lint`. -* If you added a new file, run the command `scripts/updateBatteries.sh` to update the - imports. +* If you added a new file, run the command `scripts/updateBatteries.sh` to update the imports. # Documentation From 0a62715aa3dbc19574e21c5dc7290c1b7f8c6b12 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sun, 10 Nov 2024 18:46:35 -0500 Subject: [PATCH 149/177] chore: add unlabeled PR workflow (#1033) --- .github/workflows/labels-from-status.yml | 62 ++++++++++++++++++++++++ 1 file changed, 62 insertions(+) create mode 100644 .github/workflows/labels-from-status.yml diff --git a/.github/workflows/labels-from-status.yml b/.github/workflows/labels-from-status.yml new file mode 100644 index 0000000000..8ad37be1ae --- /dev/null +++ b/.github/workflows/labels-from-status.yml @@ -0,0 +1,62 @@ +# This workflow assigns `awaiting-review` or `WIP` labels to new PRs, and it removes +# `awaiting-review`, `awaiting-author`, or `WIP` label from closed PRs. +# It does not modify labels for open PRs that already have one of the `awaiting-review`, +# `awaiting-author`, or `WIP` labels. + +name: Label PR from status change + +permissions: + contents: read + pull-requests: write + +on: + pull_request: + types: + - closed + - opened + - reopened + - converted_to_draft + - ready_for_review + branches: + - main + +jobs: + auto-label: + if: github.repository_owner == 'leanprover-community' + runs-on: ubuntu-latest + steps: + + - uses: actions/checkout@v4 + with: + fetch-depth: 0 + + - name: Unlabel closed PR + if: github.event.pull_request.state == 'closed' + uses: actions-ecosystem/action-remove-labels@v1 + with: + labels: | + WIP + awaiting-author + awaiting-review + + - name: Label unlabeled draft PR as WIP + if: | + github.event.pull_request.state == 'open' && + github.event.pull_request.draft && + ! contains(github.event.pull_request.labels.*.name, 'awaiting-author') && + ! contains(github.event.pull_request.labels.*.name, 'awaiting-review') && + ! contains(github.event.pull_request.labels.*.name, 'WIP') + uses: actions-ecosystem/action-add-labels@v1 + with: + labels: WIP + + - name: Label unlabeled other PR as awaiting-review + if: | + github.event.pull_request.state == 'open' && + ! github.event.pull_request.draft && + ! contains(github.event.pull_request.labels.*.name, 'awaiting-author') && + ! contains(github.event.pull_request.labels.*.name, 'awaiting-review') && + ! contains(github.event.pull_request.labels.*.name, 'WIP') + uses: actions-ecosystem/action-add-labels@v1 + with: + labels: awaiting-review From 325d06c3ef6bcb16c43c4ac6614976731ce6a63e Mon Sep 17 00:00:00 2001 From: Austin Letson Date: Sun, 10 Nov 2024 19:34:24 -0500 Subject: [PATCH 150/177] feat: detect default target on runLinter (#811) Co-authored-by: Mac Malone --- scripts/runLinter.lean | 94 +++++++++++++++++++++++++++++++++++------- 1 file changed, 78 insertions(+), 16 deletions(-) diff --git a/scripts/runLinter.lean b/scripts/runLinter.lean index e22dc36da5..b019a36dda 100644 --- a/scripts/runLinter.lean +++ b/scripts/runLinter.lean @@ -1,6 +1,7 @@ import Lean.Util.SearchPath import Batteries.Tactic.Lint import Batteries.Data.Array.Basic +import Lake.CLI.Main open Lean Core Elab Command Batteries.Tactic.Lint open System (FilePath) @@ -17,26 +18,65 @@ def readJsonFile (α) [FromJson α] (path : System.FilePath) : IO α := do def writeJsonFile [ToJson α] (path : System.FilePath) (a : α) : IO Unit := IO.FS.writeFile path <| toJson a |>.pretty.push '\n' +open Lake + +/-- Returns the root modules of `lean_exe` or `lean_lib` default targets in the Lake workspace. -/ +def resolveDefaultRootModules : IO (Array Name) := do + -- load the Lake workspace + let (elanInstall?, leanInstall?, lakeInstall?) ← findInstall? + let config ← MonadError.runEIO <| mkLoadConfig { elanInstall?, leanInstall?, lakeInstall? } + let some workspace ← loadWorkspace config |>.toBaseIO + | throw <| IO.userError "failed to load Lake workspace" + + -- resolve the default build specs from the Lake workspace (based on `lake build`) + let defaultBuildSpecs ← match resolveDefaultPackageTarget workspace workspace.root with + | Except.error e => IO.eprintln s!"Error getting default package target: {e}" *> IO.Process.exit 1 + | Except.ok targets => pure targets + + -- build an array of all root modules of `lean_exe` and `lean_lib` build targets + let defaultTargetModules := defaultBuildSpecs.flatMap <| + fun target => match target.info with + | BuildInfo.libraryFacet lib _ => lib.roots + | BuildInfo.leanExe exe => #[exe.config.root] + | _ => #[] + return defaultTargetModules + /-- -Usage: `runLinter [--update] [Batteries.Data.Nat.Basic]` +Parse args list for `runLinter` +and return a pair of the update and specified module arguments. -Runs the linters on all declarations in the given module (or `Batteries` by default). -If `--update` is set, the `nolints` file is updated to remove any declarations that no longer need -to be nolinted. --/ -unsafe def main (args : List String) : IO Unit := do - let (update, args) := +Throws an exception if unable to parse the arguments. +Returns `none` for the specified module if no module is specified.-/ +def parseLinterArgs (args: List String) : Except String (Bool × Option Name) := + let (update, moreArgs) := match args with | "--update" :: args => (true, args) | _ => (false, args) - let some module := - match args with - | [] => some `Batteries - | [mod] => match mod.toName with - | .anonymous => none - | name => some name - | _ => none - | IO.eprintln "Usage: runLinter [--update] [Batteries.Data.Nat.Basic]" *> IO.Process.exit 1 + match moreArgs with + | [] => Except.ok (update, none) + | [mod] => match mod.toName with + | .anonymous => Except.error "cannot convert module to Name" + | name => Except.ok (update, some name) + | _ => Except.error "cannot parse arguments" + +/-- +Return an array of the modules to lint. + +If `specifiedModule` is not `none` return an array containing only `specifiedModule`. +Otherwise, resolve the default root modules from the Lake workspace. -/ +def determineModulesToLint (specifiedModule : Option Name) : IO (Array Name) := do + match specifiedModule with + | some module => + println!"Running linter on specified module: {module}" + return #[module] + | none => + println!"Automatically detecting modules to lint" + let defaultModules ← resolveDefaultRootModules + println!"Default modules: {defaultModules}" + return defaultModules + +/-- Run the Batteries linter on a given module and update the linter if `update` is `true`. -/ +unsafe def runLinterOnModule (update : Bool) (module : Name): IO Unit := do initSearchPath (← findSysroot) let mFile ← findOLean module unless (← mFile.pathExists) do @@ -87,4 +127,26 @@ unsafe def main (args : List String) : IO Unit := do IO.print (← fmtResults.toString) IO.Process.exit 1 else - IO.println "-- Linting passed." + IO.println "-- Linting passed for {module}." + +/-- +Usage: `runLinter [--update] [Batteries.Data.Nat.Basic]` + +Runs the linters on all declarations in the given module +(or all root modules of Lake `lean_lib` and `lean_exe` default targets if no module is specified). +If `--update` is set, the `nolints` file is updated to remove any declarations that no longer need +to be nolinted. +-/ +unsafe def main (args : List String) : IO Unit := do + let linterArgs := parseLinterArgs args + let (update, specifiedModule) ← match linterArgs with + | Except.ok args => pure args + | Except.error msg => do + IO.eprintln s!"Error parsing args: {msg}" + IO.eprintln "Usage: runLinter [--update] [Batteries.Data.Nat.Basic]" + IO.Process.exit 1 + + let modulesToLint ← determineModulesToLint specifiedModule + + modulesToLint.forM <| runLinterOnModule update + From 1bfaec42a7481a81e7e12565442bd1f8dc202ae2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 11 Nov 2024 03:15:05 -0500 Subject: [PATCH 151/177] refactor: implement `BinaryHeap` using `Vector` (#850) Co-authored-by: Mario Carneiro --- Batteries/Data/BinaryHeap.lean | 174 +++++++++++++++++---------------- 1 file changed, 91 insertions(+), 83 deletions(-) diff --git a/Batteries/Data/BinaryHeap.lean b/Batteries/Data/BinaryHeap.lean index b36ed97fd3..29a273a9d0 100644 --- a/Batteries/Data/BinaryHeap.lean +++ b/Batteries/Data/BinaryHeap.lean @@ -1,78 +1,76 @@ /- Copyright (c) 2021 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro +Authors: Mario Carneiro, François G. Dorais -/ +import Batteries.Data.Vector.Basic + namespace Batteries /-- A max-heap data structure. -/ structure BinaryHeap (α) (lt : α → α → Bool) where - /-- Backing array for `BinaryHeap`. -/ + /-- `O(1)`. Get data array for a `BinaryHeap`. -/ arr : Array α namespace BinaryHeap -/-- Core operation for binary heaps, expressed directly on arrays. -Given an array which is a max-heap, push item `i` down to restore the max-heap property. -/ -def heapifyDown (lt : α → α → Bool) (a : Array α) (i : Fin a.size) : - {a' : Array α // a'.size = a.size} := +private def maxChild (lt : α → α → Bool) (a : Vector α sz) (i : Fin sz) : Option (Fin sz) := let left := 2 * i.1 + 1 let right := left + 1 - have left_le : i ≤ left := Nat.le_trans - (by rw [Nat.succ_mul, Nat.one_mul]; exact Nat.le_add_left i i) - (Nat.le_add_right ..) - have right_le : i ≤ right := Nat.le_trans left_le (Nat.le_add_right ..) - have i_le : i ≤ i := Nat.le_refl _ - have j : {j : Fin a.size // i ≤ j} := if h : left < a.size then - if lt (a.get i) (a.get ⟨left, h⟩) then ⟨⟨left, h⟩, left_le⟩ else ⟨i, i_le⟩ else ⟨i, i_le⟩ - have j := if h : right < a.size then - if lt (a.get j) (a.get ⟨right, h⟩) then ⟨⟨right, h⟩, right_le⟩ else j else j - if h : i.1 = j then ⟨a, rfl⟩ else - let a' := a.swap i j - let j' := ⟨j, by rw [a.size_swap i j]; exact j.1.2⟩ - have : a'.size - j < a.size - i := by - rw [a.size_swap i j]; exact Nat.sub_lt_sub_left i.2 <| Nat.lt_of_le_of_ne j.2 h - let ⟨a₂, h₂⟩ := heapifyDown lt a' j' - ⟨a₂, h₂.trans (a.size_swap i j)⟩ -termination_by a.size - i - -@[simp] theorem size_heapifyDown (lt : α → α → Bool) (a : Array α) (i : Fin a.size) : - (heapifyDown lt a i).1.size = a.size := (heapifyDown lt a i).2 + if hleft : left < sz then + if hright : right < sz then + if lt a[left] a[right] then + some ⟨right, hright⟩ + else + some ⟨left, hleft⟩ + else + some ⟨left, hleft⟩ + else none + +/-- Core operation for binary heaps, expressed directly on arrays. +Given an array which is a max-heap, push item `i` down to restore the max-heap property. -/ +def heapifyDown (lt : α → α → Bool) (a : Vector α sz) (i : Fin sz) : + Vector α sz := + match h : maxChild lt a i with + | none => a + | some j => + have : i < j := by + cases i; cases j + simp only [maxChild] at h + split at h + · split at h + · split at h <;> (cases h; simp_arith) + · cases h; simp_arith + · contradiction + if lt a[i] a[j] then + heapifyDown lt (a.swap i j) j + else a +termination_by sz - i /-- Core operation for binary heaps, expressed directly on arrays. Construct a heap from an unsorted array, by heapifying all the elements. -/ -def mkHeap (lt : α → α → Bool) (a : Array α) : {a' : Array α // a'.size = a.size} := - loop (a.size / 2) a (Nat.div_le_self ..) +def mkHeap (lt : α → α → Bool) (a : Vector α sz) : Vector α sz := + loop (sz / 2) a (Nat.div_le_self ..) where /-- Inner loop for `mkHeap`. -/ - loop : (i : Nat) → (a : Array α) → i ≤ a.size → {a' : Array α // a'.size = a.size} - | 0, a, _ => ⟨a, rfl⟩ + loop : (i : Nat) → (a : Vector α sz) → i ≤ sz → Vector α sz + | 0, a, _ => a | i+1, a, h => - let h := Nat.lt_of_succ_le h - let a' := heapifyDown lt a ⟨i, h⟩ - let ⟨a₂, h₂⟩ := loop i a' ((heapifyDown ..).2.symm ▸ Nat.le_of_lt h) - ⟨a₂, h₂.trans a'.2⟩ - -@[simp] theorem size_mkHeap (lt : α → α → Bool) (a : Array α) : - (mkHeap lt a).1.size = a.size := (mkHeap lt a).2 + let a' := heapifyDown lt a ⟨i, Nat.lt_of_succ_le h⟩ + loop i a' (Nat.le_trans (Nat.le_succ _) h) /-- Core operation for binary heaps, expressed directly on arrays. Given an array which is a max-heap, push item `i` up to restore the max-heap property. -/ -def heapifyUp (lt : α → α → Bool) (a : Array α) (i : Fin a.size) : - {a' : Array α // a'.size = a.size} := - if i0 : i.1 = 0 then ⟨a, rfl⟩ else - have : (i.1 - 1) / 2 < i := Nat.lt_of_le_of_lt (Nat.div_le_self ..) <| - Nat.sub_lt (Nat.pos_of_ne_zero i0) Nat.zero_lt_one - let j := ⟨(i.1 - 1) / 2, Nat.lt_trans this i.2⟩ - if lt (a.get j) (a.get i) then - let a' := a.swap i j - let ⟨a₂, h₂⟩ := heapifyUp lt a' ⟨j.1, by rw [a.size_swap i j]; exact j.2⟩ - ⟨a₂, h₂.trans (a.size_swap i j)⟩ - else ⟨a, rfl⟩ - -@[simp] theorem size_heapifyUp (lt : α → α → Bool) (a : Array α) (i : Fin a.size) : - (heapifyUp lt a i).1.size = a.size := (heapifyUp lt a i).2 +def heapifyUp (lt : α → α → Bool) (a : Vector α sz) (i : Fin sz) : + Vector α sz := + match i with + | ⟨0, _⟩ => a + | ⟨i'+1, hi⟩ => + let j := ⟨i'/2, by get_elem_tactic⟩ + if lt a[j] a[i] then + heapifyUp lt (a.swap i j) j + else a /-- `O(1)`. Build a new empty heap. -/ def empty (lt) : BinaryHeap α lt := ⟨#[]⟩ @@ -86,81 +84,91 @@ def singleton (lt) (x : α) : BinaryHeap α lt := ⟨#[x]⟩ /-- `O(1)`. Get the number of elements in a `BinaryHeap`. -/ def size (self : BinaryHeap α lt) : Nat := self.1.size +/-- `O(1)`. Get data vector of a `BinaryHeap`. -/ +def vector (self : BinaryHeap α lt) : Vector α self.size := ⟨self.1, rfl⟩ + /-- `O(1)`. Get an element in the heap by index. -/ def get (self : BinaryHeap α lt) (i : Fin self.size) : α := self.1.get i /-- `O(log n)`. Insert an element into a `BinaryHeap`, preserving the max-heap property. -/ def insert (self : BinaryHeap α lt) (x : α) : BinaryHeap α lt where - arr := let n := self.size; - heapifyUp lt (self.1.push x) ⟨n, by rw [Array.size_push]; apply Nat.lt_succ_self⟩ + arr := heapifyUp lt (self.vector.push x) ⟨_, Nat.lt_succ_self _⟩ |>.toArray @[simp] theorem size_insert (self : BinaryHeap α lt) (x : α) : (self.insert x).size = self.size + 1 := by - simp [insert, size, size_heapifyUp] + simp [size, insert] /-- `O(1)`. Get the maximum element in a `BinaryHeap`. -/ -def max (self : BinaryHeap α lt) : Option α := self.1.get? 0 - -/-- Auxiliary for `popMax`. -/ -def popMaxAux (self : BinaryHeap α lt) : {a' : BinaryHeap α lt // a'.size = self.size - 1} := - match e: self.1.size with - | 0 => ⟨self, by simp [size, e]⟩ - | n+1 => - have h0 := by rw [e]; apply Nat.succ_pos - have hn := by rw [e]; apply Nat.lt_succ_self - if hn0 : 0 < n then - let a := self.1.swap ⟨0, h0⟩ ⟨n, hn⟩ |>.pop - ⟨⟨heapifyDown lt a ⟨0, by rwa [Array.size_pop, Array.size_swap, e]⟩⟩, - by simp [size, a]⟩ - else - ⟨⟨self.1.pop⟩, by simp [size]⟩ +def max (self : BinaryHeap α lt) : Option α := self.1[0]? /-- `O(log n)`. Remove the maximum element from a `BinaryHeap`. Call `max` first to actually retrieve the maximum element. -/ -def popMax (self : BinaryHeap α lt) : BinaryHeap α lt := self.popMaxAux +def popMax (self : BinaryHeap α lt) : BinaryHeap α lt := + if h0 : self.size = 0 then self else + have hs : self.size - 1 < self.size := Nat.pred_lt h0 + have h0 : 0 < self.size := Nat.zero_lt_of_ne_zero h0 + let v := self.vector.swap ⟨_, h0⟩ ⟨_, hs⟩ |>.pop + if h : 0 < self.size - 1 then + ⟨heapifyDown lt v ⟨0, h⟩ |>.toArray⟩ + else + ⟨v.toArray⟩ @[simp] theorem size_popMax (self : BinaryHeap α lt) : - self.popMax.size = self.size - 1 := self.popMaxAux.2 + self.popMax.size = self.size - 1 := by + simp only [popMax, size] + split + · simp_arith [*] + · split <;> simp_arith [*] /-- `O(log n)`. Return and remove the maximum element from a `BinaryHeap`. -/ def extractMax (self : BinaryHeap α lt) : Option α × BinaryHeap α lt := (self.max, self.popMax) -theorem size_pos_of_max {self : BinaryHeap α lt} (e : self.max = some x) : 0 < self.size := - Decidable.of_not_not fun h : ¬ 0 < self.1.size => by simp [BinaryHeap.max, Array.get?, h] at e +theorem size_pos_of_max {self : BinaryHeap α lt} (h : self.max = some x) : 0 < self.size := by + simp only [max, getElem?_def] at h + split at h + · assumption + · contradiction /-- `O(log n)`. Equivalent to `extractMax (self.insert x)`, except that extraction cannot fail. -/ def insertExtractMax (self : BinaryHeap α lt) (x : α) : α × BinaryHeap α lt := - match e: self.max with + match e : self.max with | none => (x, self) | some m => if lt x m then - let a := self.1.set ⟨0, size_pos_of_max e⟩ x - (m, ⟨heapifyDown lt a ⟨0, by simp only [Array.size_set, a]; exact size_pos_of_max e⟩⟩) + let v := self.vector.set ⟨0, size_pos_of_max e⟩ x + (m, ⟨heapifyDown lt v ⟨0, size_pos_of_max e⟩ |>.toArray⟩) else (x, self) /-- `O(log n)`. Equivalent to `(self.max, self.popMax.insert x)`. -/ def replaceMax (self : BinaryHeap α lt) (x : α) : Option α × BinaryHeap α lt := - match e: self.max with - | none => (none, ⟨self.1.push x⟩) + match e : self.max with + | none => (none, ⟨self.vector.push x |>.toArray⟩) | some m => - let a := self.1.set ⟨0, size_pos_of_max e⟩ x - (some m, ⟨heapifyDown lt a ⟨0, by simp only [Array.size_set, a]; exact size_pos_of_max e⟩⟩) + let v := self.vector.set ⟨0, size_pos_of_max e⟩ x + (some m, ⟨heapifyDown lt v ⟨0, size_pos_of_max e⟩ |>.toArray⟩) /-- `O(log n)`. Replace the value at index `i` by `x`. Assumes that `x ≤ self.get i`. -/ def decreaseKey (self : BinaryHeap α lt) (i : Fin self.size) (x : α) : BinaryHeap α lt where - arr := heapifyDown lt (self.1.set i x) ⟨i, by rw [self.1.size_set]; exact i.2⟩ + arr := heapifyDown lt (self.vector.set i x) i |>.toArray /-- `O(log n)`. Replace the value at index `i` by `x`. Assumes that `self.get i ≤ x`. -/ def increaseKey (self : BinaryHeap α lt) (i : Fin self.size) (x : α) : BinaryHeap α lt where - arr := heapifyUp lt (self.1.set i x) ⟨i, by rw [self.1.size_set]; exact i.2⟩ + arr := heapifyUp lt (self.vector.set i x) i |>.toArray end Batteries.BinaryHeap +/-- `O(n)`. Convert an unsorted vector to a `BinaryHeap`. -/ +def Batteries.Vector.toBinaryHeap (lt : α → α → Bool) (v : Vector α n) : + Batteries.BinaryHeap α lt where + arr := BinaryHeap.mkHeap lt v |>.toArray + +open Batteries in /-- `O(n)`. Convert an unsorted array to a `BinaryHeap`. -/ def Array.toBinaryHeap (lt : α → α → Bool) (a : Array α) : Batteries.BinaryHeap α lt where - arr := Batteries.BinaryHeap.mkHeap lt a + arr := BinaryHeap.mkHeap lt ⟨a, rfl⟩ |>.toArray +open Batteries in /-- `O(n log n)`. Sort an array using a `BinaryHeap`. -/ @[specialize] def Array.heapSort (a : Array α) (lt : α → α → Bool) : Array α := loop (a.toBinaryHeap (flip lt)) #[] From 12381da57b6d3e42fa6f454d865ffaf8ecc687a8 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 12 Nov 2024 03:00:02 +0000 Subject: [PATCH 152/177] perf: inline the `Decidable (Coprime _ _)` instance (#1036) --- Batteries/Data/Nat/Gcd.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Batteries/Data/Nat/Gcd.lean b/Batteries/Data/Nat/Gcd.lean index 708585f5c3..6bc9f43129 100644 --- a/Batteries/Data/Nat/Gcd.lean +++ b/Batteries/Data/Nat/Gcd.lean @@ -19,7 +19,8 @@ See also `nat.coprime_of_dvd` and `nat.coprime_of_dvd'` to prove `nat.Coprime m /-- `m` and `n` are coprime, or relatively prime, if their `gcd` is 1. -/ @[reducible] def Coprime (m n : Nat) : Prop := gcd m n = 1 -instance (m n : Nat) : Decidable (Coprime m n) := inferInstanceAs (Decidable (_ = 1)) +-- if we don't inline this, then the compiler computes the GCD even if it already has it +@[inline] instance (m n : Nat) : Decidable (Coprime m n) := inferInstanceAs (Decidable (_ = 1)) theorem coprime_iff_gcd_eq_one : Coprime m n ↔ gcd m n = 1 := .rfl From 44f2360f50d0d3aea24653103a3f11ef5de9eb90 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Tue, 12 Nov 2024 04:19:01 +0100 Subject: [PATCH 153/177] chore(Data/Rat): move Float functions to separate file (#1035) Co-authored-by: Kim Morrison --- Batteries/Data/Rat.lean | 1 + Batteries/Data/Rat/Basic.lean | 17 +---------------- Batteries/Data/Rat/Float.lean | 29 +++++++++++++++++++++++++++++ 3 files changed, 31 insertions(+), 16 deletions(-) create mode 100644 Batteries/Data/Rat/Float.lean diff --git a/Batteries/Data/Rat.lean b/Batteries/Data/Rat.lean index 686b74eed3..8045da7952 100644 --- a/Batteries/Data/Rat.lean +++ b/Batteries/Data/Rat.lean @@ -1,2 +1,3 @@ import Batteries.Data.Rat.Basic +import Batteries.Data.Rat.Float import Batteries.Data.Rat.Lemmas diff --git a/Batteries/Data/Rat/Basic.lean b/Batteries/Data/Rat/Basic.lean index 01cbd07859..37951fbec6 100644 --- a/Batteries/Data/Rat/Basic.lean +++ b/Batteries/Data/Rat/Basic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Batteries.Data.Nat.Gcd -import Batteries.Lean.Float /-! # Basics for the Rational Numbers -/ @@ -278,18 +277,4 @@ protected def ceil (a : Rat) : Int := else a.num / a.den + 1 -/-- Convert this rational number to a `Float` value. -/ -protected def toFloat (a : Rat) : Float := a.num.divFloat a.den - -/-- Convert this floating point number to a rational value. -/ -protected def _root_.Float.toRat? (a : Float) : Option Rat := - a.toRatParts.map fun (v, exp) => - mkRat (v.sign * v.natAbs <<< exp.toNat) (1 <<< (-exp).toNat) - -/-- -Convert this floating point number to a rational value, -mapping non-finite values (`inf`, `-inf`, `nan`) to 0. --/ -protected def _root_.Float.toRat0 (a : Float) : Rat := a.toRat?.getD 0 - -instance : Coe Rat Float := ⟨Rat.toFloat⟩ +end Rat diff --git a/Batteries/Data/Rat/Float.lean b/Batteries/Data/Rat/Float.lean new file mode 100644 index 0000000000..271a6b40fb --- /dev/null +++ b/Batteries/Data/Rat/Float.lean @@ -0,0 +1,29 @@ +/- +Copyright (c) 2022 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ +import Batteries.Data.Rat.Basic +import Batteries.Lean.Float + +/-! # Rational Numbers and Float -/ + +namespace Rat + +/-- Convert this rational number to a `Float` value. -/ +protected def toFloat (a : Rat) : Float := a.num.divFloat a.den + +/-- Convert this floating point number to a rational value. -/ +protected def _root_.Float.toRat? (a : Float) : Option Rat := + a.toRatParts.map fun (v, exp) => + mkRat (v.sign * v.natAbs <<< exp.toNat) (1 <<< (-exp).toNat) + +/-- +Convert this floating point number to a rational value, +mapping non-finite values (`inf`, `-inf`, `nan`) to 0. +-/ +protected def _root_.Float.toRat0 (a : Float) : Rat := a.toRat?.getD 0 + +instance : Coe Rat Float := ⟨Rat.toFloat⟩ + +end Rat From fd970b247cdb7dca34f4949a6bd50e63ac9b10ae Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 12 Nov 2024 21:23:35 +1100 Subject: [PATCH 154/177] chore: robustify some String proofs (#1038) --- Batteries/Data/String/Lemmas.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Batteries/Data/String/Lemmas.lean b/Batteries/Data/String/Lemmas.lean index 5854890f25..431591ebc1 100644 --- a/Batteries/Data/String/Lemmas.lean +++ b/Batteries/Data/String/Lemmas.lean @@ -64,7 +64,7 @@ theorem utf8Len_reverseAux (cs₁ cs₂) : @[simp] theorem utf8Len_reverse (cs) : utf8Len cs.reverse = utf8Len cs := utf8Len_reverseAux .. @[simp] theorem utf8Len_eq_zero : utf8Len l = 0 ↔ l = [] := by - cases l <;> simp [Nat.ne_of_gt add_utf8Size_pos] + cases l <;> simp [Nat.ne_zero_iff_zero_lt.mpr (Char.utf8Size_pos _)] section open List @@ -256,7 +256,7 @@ theorem back_eq (s : String) : back s = s.1.getLastD default := by theorem atEnd_of_valid (cs : List Char) (cs' : List Char) : atEnd ⟨cs ++ cs'⟩ ⟨utf8Len cs⟩ ↔ cs' = [] := by rw [atEnd_iff] - cases cs' <;> simp [Nat.lt_add_of_pos_right add_utf8Size_pos] + cases cs' <;> simp [add_utf8Size_pos] unseal posOfAux findAux in theorem posOfAux_eq (s c) : posOfAux s c = findAux s (· == c) := rfl From e6047ba5227277c06879db3dd6533d8fc722fcb3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Tue, 12 Nov 2024 18:13:55 -0500 Subject: [PATCH 155/177] feat: generate docs from subdirectory (#1028) Co-authored-by: Kim Morrison --- .github/workflows/docs-deploy.yml | 38 ++++++++++++++++++ .github/workflows/docs-release.yml | 46 ++++++++++++++++++++++ README.md | 30 ++++----------- docs/README.md | 1 + docs/lake-manifest.json | 62 ++++++++++++++++++++++++++++++ docs/lakefile.toml | 13 +++++++ 6 files changed, 168 insertions(+), 22 deletions(-) create mode 100644 .github/workflows/docs-deploy.yml create mode 100644 .github/workflows/docs-release.yml create mode 120000 docs/README.md create mode 100644 docs/lake-manifest.json create mode 100644 docs/lakefile.toml diff --git a/.github/workflows/docs-deploy.yml b/.github/workflows/docs-deploy.yml new file mode 100644 index 0000000000..39746f19d3 --- /dev/null +++ b/.github/workflows/docs-deploy.yml @@ -0,0 +1,38 @@ +name: Deploy Docs + +on: + workflow_dispatch: + schedule: + - cron: '0 10 * * *' # daily (UTC 10:00) + +permissions: + contents: write + +jobs: + deploy-docs: + runs-on: ubuntu-latest + if: github.repository_owner == 'leanprover-community' + steps: + + - name: Checkout + uses: actions/checkout@v4 + + - name: Install Lean + uses: leanprover/lean-action@v1 + with: + test: false + lint: false + use-github-cache: true + + - name: Build Docs + working-directory: docs + run: lake build -q Batteries:docs + + - name: Deploy Docs + run: | + git config user.name "leanprover-community-batteries-bot" + git config user.email "leanprover-community-batteries-bot@users.noreply.github.com" + git checkout -b docs + git add docs/doc docs/doc-data + git commit -m "chore: generate docs" + git push origin docs --force diff --git a/.github/workflows/docs-release.yml b/.github/workflows/docs-release.yml new file mode 100644 index 0000000000..2926a31340 --- /dev/null +++ b/.github/workflows/docs-release.yml @@ -0,0 +1,46 @@ +name: Release Docs + +on: + push: + tags: + - "v[0-9]+.[0-9]+.[0-9]+" + - "v[0-9]+.[0-9]+.[0-9]+-rc[0-9]+" + +permissions: + contents: write + +jobs: + build-docs: + runs-on: ubuntu-latest + if: github.repository_owner == 'leanprover-community' + steps: + + - name: Checkout + uses: actions/checkout@v4 + + - name: Install Lean + uses: leanprover/lean-action@v1 + with: + test: false + lint: false + use-github-cache: true + + - name: Build Docs + working-directory: docs + run: lake build -q Batteries:docs + + - name: Compress Docs + working-directory: docs + env: + TAG_NAME: ${{ github.ref_name }} + run: | + tar -czf docs-${TAG_NAME}.tar.gz doc doc-data + zip -rq docs-${TAG_NAME}.zip doc doc-data + + - name: Release Docs + uses: softprops/action-gh-release@v2 + with: + files: | + docs/docs-${{ github.ref_name }}.tar.gz + docs/docs-${{ github.ref_name }}.zip + fail_on_unmatched_files: true diff --git a/README.md b/README.md index f8cc2415ac..2873f65536 100644 --- a/README.md +++ b/README.md @@ -36,31 +36,17 @@ Additionally, please make sure that you're using the version of Lean that the cu # Documentation -You can generate `batteries`' documentation with - -```text -# if you're generating documentation for the first time -> lake -R -Kdoc=on update -... -# actually generate the documentation -> lake -R -Kdoc=on build Batteries:docs -... -> ls build/doc/index.html -build/doc/index.html -``` - -After generating the docs, run `lake build -R` to reset the configuration. +You can generate `batteries` documentation with -The top-level HTML file will be located at `build/doc/Batteries.html`, though to actually expose the -documentation as a server you need to - -```text -> cd build/doc -> python3 -m http.server -Serving HTTP on :: port 8000 (http://[::]:8000/) ... +```sh +cd docs +lake build Batteries:docs ``` -Note that documentation for the latest nightly of `batteries` is available as part of [the Mathlib 4 +The top-level HTML file will be located at `docs/doc/index.html`, though to actually expose the +documentation you need to run a HTTP server (e.g. `python3 -m http.server`) in the `docs/doc` directory. + +Note that documentation for the latest nightly of `batteries` is also available as part of [the Mathlib 4 documentation][mathlib4 docs]. [mathlib4 docs]: https://leanprover-community.github.io/mathlib4_docs/Batteries.html diff --git a/docs/README.md b/docs/README.md new file mode 120000 index 0000000000..32d46ee883 --- /dev/null +++ b/docs/README.md @@ -0,0 +1 @@ +../README.md \ No newline at end of file diff --git a/docs/lake-manifest.json b/docs/lake-manifest.json new file mode 100644 index 0000000000..9c082aa9e8 --- /dev/null +++ b/docs/lake-manifest.json @@ -0,0 +1,62 @@ +{"version": "1.1.0", + "packagesDir": "../.lake/packages", + "packages": + [{"url": "https://github.com/acmepjz/md4lean", + "type": "git", + "subDir": null, + "scope": "", + "rev": "5e95f4776be5e048364f325c7e9d619bb56fb005", + "name": "MD4Lean", + "manifestFile": "lake-manifest.json", + "inputRev": "main", + "inherited": true, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/fgdorais/lean4-unicode-basic", + "type": "git", + "subDir": null, + "scope": "", + "rev": "b41bc9cec7f433d6e1d74ff3b59edaaf58ad2915", + "name": "UnicodeBasic", + "manifestFile": "lake-manifest.json", + "inputRev": "main", + "inherited": true, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/dupuisf/BibtexQuery", + "type": "git", + "subDir": null, + "scope": "", + "rev": "bdc2fc30b1e834b294759a5d391d83020a90058e", + "name": "BibtexQuery", + "manifestFile": "lake-manifest.json", + "inputRev": "master", + "inherited": true, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/mhuisi/lean4-cli", + "type": "git", + "subDir": null, + "scope": "", + "rev": "726b3c9ad13acca724d4651f14afc4804a7b0e4d", + "name": "Cli", + "manifestFile": "lake-manifest.json", + "inputRev": "main", + "inherited": true, + "configFile": "lakefile.toml"}, + {"url": "https://github.com/leanprover/doc-gen4", + "type": "git", + "subDir": null, + "scope": "", + "rev": "8add673e2ea4da0929103ad19dc824e1c0b7437d", + "name": "«doc-gen4»", + "manifestFile": "lake-manifest.json", + "inputRev": "main", + "inherited": false, + "configFile": "lakefile.lean"}, + {"type": "path", + "scope": "", + "name": "batteries", + "manifestFile": "lake-manifest.json", + "inherited": false, + "dir": "./..", + "configFile": "lakefile.toml"}], + "name": "docs", + "lakeDir": ".lake"} diff --git a/docs/lakefile.toml b/docs/lakefile.toml new file mode 100644 index 0000000000..0726079709 --- /dev/null +++ b/docs/lakefile.toml @@ -0,0 +1,13 @@ +name = "docs" +reservoir = false +packagesDir = "../.lake/packages" +buildDir = "." + +[[require]] +scope = "leanprover" +name = "doc-gen4" +rev = "main" + +[[require]] +name = "batteries" +path = ".." From 66a3bd2d823920fcc33f32dbf317f9a8339f49e0 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 13 Nov 2024 10:14:26 +1100 Subject: [PATCH 156/177] feat: List.SatisfiesM_foldlM (#1034) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: François G. Dorais Co-authored-by: Eric Wieser --- Batteries/Data/List.lean | 1 + Batteries/Data/List/Monadic.lean | 41 ++++++++++++++++++++++++++++++++ 2 files changed, 42 insertions(+) create mode 100644 Batteries/Data/List/Monadic.lean diff --git a/Batteries/Data/List.lean b/Batteries/Data/List.lean index f93f90a6c2..3429039dc9 100644 --- a/Batteries/Data/List.lean +++ b/Batteries/Data/List.lean @@ -3,6 +3,7 @@ import Batteries.Data.List.Count import Batteries.Data.List.EraseIdx import Batteries.Data.List.Init.Lemmas import Batteries.Data.List.Lemmas +import Batteries.Data.List.Monadic import Batteries.Data.List.OfFn import Batteries.Data.List.Pairwise import Batteries.Data.List.Perm diff --git a/Batteries/Data/List/Monadic.lean b/Batteries/Data/List/Monadic.lean new file mode 100644 index 0000000000..00eb213755 --- /dev/null +++ b/Batteries/Data/List/Monadic.lean @@ -0,0 +1,41 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kim Morrison +-/ +import Batteries.Classes.SatisfiesM + +/-! +# Results about monadic operations on `List`, in terms of `SatisfiesM`. + +-/ + +namespace List + +theorem satisfiesM_foldlM [Monad m] [LawfulMonad m] {f : β → α → m β} (h₀ : motive b) + (h₁ : ∀ (b) (_ : motive b) (a : α) (_ : a ∈ l), SatisfiesM motive (f b a)) : + SatisfiesM motive (List.foldlM f b l) := by + have g b hb a am := Classical.indefiniteDescription _ (h₁ b hb a am) + clear h₁ + induction l generalizing b with + | nil => exact SatisfiesM.pure h₀ + | cons hd tl ih => + simp only [foldlM_cons] + apply SatisfiesM.bind_pre + let ⟨q, qh⟩ := g b h₀ hd (mem_cons_self hd tl) + exact ⟨(fun ⟨b, bh⟩ => ⟨b, ih bh (fun b bh a am => g b bh a (mem_cons_of_mem hd am))⟩) <$> q, + by simpa using qh⟩ + +theorem satisfiesM_foldrM [Monad m] [LawfulMonad m] {f : α → β → m β} (h₀ : motive b) + (h₁ : ∀ (a : α) (_ : a ∈ l) (b) (_ : motive b), SatisfiesM motive (f a b)) : + SatisfiesM motive (List.foldrM f b l) := by + induction l with + | nil => exact SatisfiesM.pure h₀ + | cons hd tl ih => + simp only [foldrM_cons] + apply SatisfiesM.bind_pre + let ⟨q, qh⟩ := ih (fun a am b hb => h₁ a (mem_cons_of_mem hd am) b hb) + exact ⟨(fun ⟨b, bh⟩ => ⟨b, h₁ hd (mem_cons_self hd tl) b bh⟩) <$> q, + by simpa using qh⟩ + +end List From e0d84490d3dc76e242e49bf303d22411ff449e16 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 13 Nov 2024 11:03:05 +1100 Subject: [PATCH 157/177] chore: cleanup proof of satisfiesM_foldlM (#1039) --- Batteries/Data/List/Monadic.lean | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/Batteries/Data/List/Monadic.lean b/Batteries/Data/List/Monadic.lean index 00eb213755..f0fa51e370 100644 --- a/Batteries/Data/List/Monadic.lean +++ b/Batteries/Data/List/Monadic.lean @@ -15,15 +15,13 @@ namespace List theorem satisfiesM_foldlM [Monad m] [LawfulMonad m] {f : β → α → m β} (h₀ : motive b) (h₁ : ∀ (b) (_ : motive b) (a : α) (_ : a ∈ l), SatisfiesM motive (f b a)) : SatisfiesM motive (List.foldlM f b l) := by - have g b hb a am := Classical.indefiniteDescription _ (h₁ b hb a am) - clear h₁ induction l generalizing b with | nil => exact SatisfiesM.pure h₀ | cons hd tl ih => simp only [foldlM_cons] apply SatisfiesM.bind_pre - let ⟨q, qh⟩ := g b h₀ hd (mem_cons_self hd tl) - exact ⟨(fun ⟨b, bh⟩ => ⟨b, ih bh (fun b bh a am => g b bh a (mem_cons_of_mem hd am))⟩) <$> q, + let ⟨q, qh⟩ := h₁ b h₀ hd (mem_cons_self hd tl) + exact ⟨(fun ⟨b, bh⟩ => ⟨b, ih bh (fun b bh a am => h₁ b bh a (mem_cons_of_mem hd am))⟩) <$> q, by simpa using qh⟩ theorem satisfiesM_foldrM [Monad m] [LawfulMonad m] {f : α → β → m β} (h₀ : motive b) From 11da075dd931360a652129305eb045e5b1126ffd Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 13 Nov 2024 11:23:56 +1100 Subject: [PATCH 158/177] feat: sugar for SatisfiesM (#1029) Co-authored-by: Mario Carneiro --- Batteries.lean | 3 + Batteries/Classes/SatisfiesM.lean | 130 ++++++++++++++++++++++++++++-- Batteries/Data/HashMap/Basic.lean | 2 +- Batteries/Lean/EStateM.lean | 40 +++++++++ Batteries/Lean/Except.lean | 51 +++++++++++- Batteries/Lean/LawfulMonad.lean | 30 +++++++ Batteries/Lean/SatisfiesM.lean | 35 ++++++++ BatteriesTest/satisfying.lean | 8 ++ 8 files changed, 291 insertions(+), 8 deletions(-) create mode 100644 Batteries/Lean/EStateM.lean create mode 100644 Batteries/Lean/LawfulMonad.lean create mode 100644 Batteries/Lean/SatisfiesM.lean create mode 100644 BatteriesTest/satisfying.lean diff --git a/Batteries.lean b/Batteries.lean index 7e5c73f5fd..e6464503cb 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -38,6 +38,7 @@ import Batteries.Data.UnionFind import Batteries.Data.Vector import Batteries.Lean.AttributeExtra import Batteries.Lean.Delaborator +import Batteries.Lean.EStateM import Batteries.Lean.Except import Batteries.Lean.Expr import Batteries.Lean.Float @@ -45,6 +46,7 @@ import Batteries.Lean.HashMap import Batteries.Lean.HashSet import Batteries.Lean.IO.Process import Batteries.Lean.Json +import Batteries.Lean.LawfulMonad import Batteries.Lean.Meta.Basic import Batteries.Lean.Meta.DiscrTree import Batteries.Lean.Meta.Expr @@ -59,6 +61,7 @@ import Batteries.Lean.NameMapAttribute import Batteries.Lean.PersistentHashMap import Batteries.Lean.PersistentHashSet import Batteries.Lean.Position +import Batteries.Lean.SatisfiesM import Batteries.Lean.Syntax import Batteries.Lean.System.IO import Batteries.Lean.TagAttribute diff --git a/Batteries/Classes/SatisfiesM.lean b/Batteries/Classes/SatisfiesM.lean index 89f264963a..834cbbea68 100644 --- a/Batteries/Classes/SatisfiesM.lean +++ b/Batteries/Classes/SatisfiesM.lean @@ -1,8 +1,11 @@ /- Copyright (c) 2022 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro +Authors: Mario Carneiro, Kim Morrison -/ +import Batteries.Lean.EStateM +import Batteries.Lean.Except +import Batteries.Tactic.Lint /-! ## SatisfiesM @@ -12,6 +15,13 @@ and enables Hoare-like reasoning over monadic expressions. For example, given a function `f : α → m β`, to say that the return value of `f` satisfies `Q` whenever the input satisfies `P`, we write `∀ a, P a → SatisfiesM Q (f a)`. +For any monad equipped with `MonadSatisfying m` +one can lift `SatisfiesM` to a monadic value in `Subtype`, +using `satisfying x h : m {a // p a}`, where `x : m α` and `h : SatisfiesM p x`. +This includes `Option`, `ReaderT`, `StateT`, and `ExceptT`, and the Lean monad stack. +(Although it is not entirely clear one should treat the Lean monad stack as lawful, +even though Lean accepts this.) + ## Notes `SatisfiesM` is not yet a satisfactory solution for verifying the behaviour of large scale monadic @@ -23,7 +33,7 @@ presumably requiring more syntactic support (and smarter `do` blocks) from Lean. Or it may be that such a solution will look different! This is an open research program, and for now one should not be overly ambitious using `SatisfiesM`. -In particular lemmas about pure operations on data structures in `batteries` except for `HashMap` +In particular lemmas about pure operations on data structures in `Batteries` except for `HashMap` should avoid `SatisfiesM` for now, so that it is easy to migrate to other approaches in future. -/ @@ -158,15 +168,33 @@ end SatisfiesM ⟨by revert x; intro | .ok _, ⟨.ok ⟨_, h⟩, rfl⟩, _, rfl => exact h, fun h => match x with | .ok a => ⟨.ok ⟨a, h _ rfl⟩, rfl⟩ | .error e => ⟨.error e, rfl⟩⟩ +theorem SatisfiesM_EStateM_eq : + SatisfiesM (m := EStateM ε σ) p x ↔ ∀ s a s', x.run s = .ok a s' → p a := by + constructor + · rintro ⟨x, rfl⟩ s a s' h + match w : x.run s with + | .ok a s' => simp at h; exact h.1 + | .error e s' => simp [w] at h + · intro w + refine ⟨?_, ?_⟩ + · intro s + match q : x.run s with + | .ok a s' => exact .ok ⟨a, w s a s' q⟩ s' + | .error e s' => exact .error e s' + · ext s + rw [EStateM.run_map, EStateM.run] + split <;> simp_all + @[simp] theorem SatisfiesM_ReaderT_eq [Monad m] : - SatisfiesM (m := ReaderT ρ m) p x ↔ ∀ s, SatisfiesM p (x s) := + SatisfiesM (m := ReaderT ρ m) p x ↔ ∀ s, SatisfiesM p (x.run s) := (exists_congr fun a => by exact ⟨fun eq _ => eq ▸ rfl, funext⟩).trans Classical.skolem.symm theorem SatisfiesM_StateRefT_eq [Monad m] : - SatisfiesM (m := StateRefT' ω σ m) p x ↔ ∀ s, SatisfiesM p (x s) := by simp + SatisfiesM (m := StateRefT' ω σ m) p x ↔ ∀ s, SatisfiesM p (x s) := by simp [ReaderT.run] @[simp] theorem SatisfiesM_StateT_eq [Monad m] [LawfulMonad m] : - SatisfiesM (m := StateT ρ m) (α := α) p x ↔ ∀ s, SatisfiesM (m := m) (p ·.1) (x s) := by + SatisfiesM (m := StateT ρ m) (α := α) p x ↔ ∀ s, SatisfiesM (m := m) (p ·.1) (x.run s) := by + change SatisfiesM (m := StateT ρ m) (α := α) p x ↔ ∀ s, SatisfiesM (m := m) (p ·.1) (x s) refine .trans ⟨fun ⟨f, eq⟩ => eq ▸ ?_, fun ⟨f, h⟩ => ?_⟩ Classical.skolem.symm · refine ⟨fun s => (fun ⟨⟨a, h⟩, s'⟩ => ⟨⟨a, s'⟩, h⟩) <$> f s, fun s => ?_⟩ rw [← comp_map, map_eq_pure_bind]; rfl @@ -174,9 +202,99 @@ theorem SatisfiesM_StateRefT_eq [Monad m] : show _ >>= _ = _; simp [← h] @[simp] theorem SatisfiesM_ExceptT_eq [Monad m] [LawfulMonad m] : - SatisfiesM (m := ExceptT ρ m) (α := α) p x ↔ SatisfiesM (m := m) (∀ a, · = .ok a → p a) x := by + SatisfiesM (m := ExceptT ρ m) (α := α) p x ↔ + SatisfiesM (m := m) (∀ a, · = .ok a → p a) x.run := by + change _ ↔ SatisfiesM (m := m) (∀ a, · = .ok a → p a) x refine ⟨fun ⟨f, eq⟩ => eq ▸ ?_, fun ⟨f, eq⟩ => eq ▸ ?_⟩ · exists (fun | .ok ⟨a, h⟩ => ⟨.ok a, fun | _, rfl => h⟩ | .error e => ⟨.error e, nofun⟩) <$> f show _ = _ >>= _; rw [← comp_map, map_eq_pure_bind]; congr; funext a; cases a <;> rfl · exists ((fun | ⟨.ok a, h⟩ => .ok ⟨a, h _ rfl⟩ | ⟨.error e, _⟩ => .error e) <$> f : m _) show _ >>= _ = _; simp [← comp_map, ← bind_pure_comp]; congr; funext ⟨a, h⟩; cases a <;> rfl + +/-- +If a monad has `MonadSatisfying m`, then we can lift a `h : SatisfiesM (m := m) p x` predicate +to monadic value `satisfying x p : m { x // p x }`. + +Reader, state, and exception monads have `MonadSatisfying` instances if the base monad does. +-/ +class MonadSatisfying (m : Type u → Type v) [Functor m] [LawfulFunctor m] where + /-- Lift a `SatisfiesM` predicate to a monadic value. -/ + satisfying {p : α → Prop} {x : m α} (h : SatisfiesM (m := m) p x) : m {a // p a} + /-- The value of the lifted monadic value is equal to the original monadic value. -/ + val_eq {p : α → Prop} {x : m α} (h : SatisfiesM (m := m) p x) : Subtype.val <$> satisfying h = x + +export MonadSatisfying (satisfying) + +namespace MonadSatisfying + +instance : MonadSatisfying Id where + satisfying {α p x} h := ⟨x, by obtain ⟨⟨_, h⟩, rfl⟩ := h; exact h⟩ + val_eq {α p x} h := rfl + +instance : MonadSatisfying Option where + satisfying {α p x?} h := + have h' := SatisfiesM_Option_eq.mp h + match x? with + | none => none + | some x => some ⟨x, h' x rfl⟩ + val_eq {α p x?} h := by cases x? <;> simp + +instance : MonadSatisfying (Except ε) where + satisfying {α p x?} h := + have h' := SatisfiesM_Except_eq.mp h + match x? with + | .ok x => .ok ⟨x, h' x rfl⟩ + | .error e => .error e + val_eq {α p x?} h := by cases x? <;> simp + +-- This will be redundant after nightly-2024-11-08. +attribute [ext] ReaderT.ext + +instance [Monad m] [LawfulMonad m][MonadSatisfying m] : MonadSatisfying (ReaderT ρ m) where + satisfying {α p x} h := + have h' := SatisfiesM_ReaderT_eq.mp h + fun r => satisfying (h' r) + val_eq {α p x} h := by + have h' := SatisfiesM_ReaderT_eq.mp h + ext r + rw [ReaderT.run_map, ← MonadSatisfying.val_eq (h' r)] + rfl + +instance [Monad m] [LawfulMonad m] [MonadSatisfying m] : MonadSatisfying (StateRefT' ω σ m) := + inferInstanceAs <| MonadSatisfying (ReaderT _ _) + +-- This will be redundant after nightly-2024-11-08. +attribute [ext] StateT.ext + +instance [Monad m] [LawfulMonad m] [MonadSatisfying m] : MonadSatisfying (StateT ρ m) where + satisfying {α p x} h := + have h' := SatisfiesM_StateT_eq.mp h + fun r => (fun ⟨⟨a, r'⟩, h⟩ => ⟨⟨a, h⟩, r'⟩) <$> satisfying (h' r) + val_eq {α p x} h := by + have h' := SatisfiesM_StateT_eq.mp h + ext r + rw [← MonadSatisfying.val_eq (h' r), StateT.run_map] + simp [StateT.run] + +instance [Monad m] [LawfulMonad m] [MonadSatisfying m] : MonadSatisfying (ExceptT ε m) where + satisfying {α p x} h := + let x' := satisfying (SatisfiesM_ExceptT_eq.mp h) + ExceptT.mk ((fun ⟨y, w⟩ => y.pmap fun a h => ⟨a, w _ h⟩) <$> x') + val_eq {α p x} h:= by + ext + rw [← MonadSatisfying.val_eq (SatisfiesM_ExceptT_eq.mp h)] + simp + +instance : MonadSatisfying (EStateM ε σ) where + satisfying {α p x} h := + have h' := SatisfiesM_EStateM_eq.mp h + fun s => match w : x.run s with + | .ok a s' => .ok ⟨a, h' s a s' w⟩ s' + | .error e s' => .error e s' + val_eq {α p x} h := by + ext s + rw [EStateM.run_map, EStateM.run] + simp only + split <;> simp_all + +end MonadSatisfying diff --git a/Batteries/Data/HashMap/Basic.lean b/Batteries/Data/HashMap/Basic.lean index 3fd1d3350d..de029c4276 100644 --- a/Batteries/Data/HashMap/Basic.lean +++ b/Batteries/Data/HashMap/Basic.lean @@ -202,7 +202,7 @@ Applies `f` to each key-value pair `a, b` in the map. If it returns `some c` the have : m'.1.size > 0 := by have := Array.size_mapM (m := StateT (ULift Nat) Id) (go .nil) m.buckets.1 simp [SatisfiesM_StateT_eq, SatisfiesM_Id_eq] at this - simp [this, Id.run, StateT.run, m.2.2, m'] + simp [this, Id.run, m.2.2, m'] ⟨m'.2.1, m'.1, this⟩ where /-- Inner loop of `filterMap`. Note that this reverses the bucket lists, diff --git a/Batteries/Lean/EStateM.lean b/Batteries/Lean/EStateM.lean new file mode 100644 index 0000000000..84049c3560 --- /dev/null +++ b/Batteries/Lean/EStateM.lean @@ -0,0 +1,40 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kim Morrison +-/ + +namespace EStateM + +namespace Result + +/-- Map a function over an `EStateM.Result`, preserving states and errors. -/ +def map {ε σ α β : Type u} (f : α → β) (x : Result ε σ α) : Result ε σ β := + match x with + | .ok a s' => .ok (f a) s' + | .error e s' => .error e s' + +@[simp] theorem map_ok {ε σ α β : Type u} (f : α → β) (a : α) (s : σ) : + (Result.ok a s : Result ε σ α).map f = .ok (f a) s := rfl + +@[simp] theorem map_error {ε σ α β : Type u} (f : α → β) (e : ε) (s : σ) : + (Result.error e s : Result ε σ α).map f = .error e s := rfl + +@[simp] theorem map_eq_ok {ε σ α β : Type u} (f : α → β) (x : Result ε σ α) (b : β) (s : σ) : + x.map f = .ok b s ↔ ∃ a, x = .ok a s ∧ b = f a := by + cases x <;> simp [and_assoc, and_comm, eq_comm] + +@[simp] theorem map_eq_error {ε σ α β : Type u} (f : α → β) (x : Result ε σ α) (e : ε) (s : σ) : + x.map f = .error e s ↔ x = .error e s := by + cases x <;> simp [eq_comm] + +end Result + +@[simp] theorem run_map (f : α → β) (x : EStateM ε σ α) : + (f <$> x).run s = (x.run s).map f := rfl + +@[ext] theorem ext {ε σ α : Type u} (x y : EStateM ε σ α) (h : ∀ s, x.run s = y.run s) : x = y := by + funext s + exact h s + +end EStateM diff --git a/Batteries/Lean/Except.lean b/Batteries/Lean/Except.lean index 802d6b5161..e0eb4c79f5 100644 --- a/Batteries/Lean/Except.lean +++ b/Batteries/Lean/Except.lean @@ -7,7 +7,56 @@ import Lean.Util.Trace open Lean +namespace Except + /-- Visualize an `Except` using a checkmark or a cross. -/ -def Except.emoji : Except ε α → String +def emoji : Except ε α → String | .error _ => crossEmoji | .ok _ => checkEmoji + +@[simp] theorem map_error {ε : Type u} (f : α → β) (e : ε) : + f <$> (.error e : Except ε α) = .error e := rfl + +@[simp] theorem map_ok {ε : Type u} (f : α → β) (x : α) : + f <$> (.ok x : Except ε α) = .ok (f x) := rfl + +/-- Map a function over an `Except` value, using a proof that the value is `.ok`. -/ +def pmap {ε : Type u} {α β : Type v} (x : Except ε α) (f : (a : α) → x = .ok a → β) : Except ε β := + match x with + | .error e => .error e + | .ok a => .ok (f a rfl) + +@[simp] theorem pmap_error {ε : Type u} {α β : Type v} (e : ε) + (f : (a : α) → Except.error e = Except.ok a → β) : + Except.pmap (.error e) f = .error e := rfl + +@[simp] theorem pmap_ok {ε : Type u} {α β : Type v} (a : α) + (f : (a' : α) → (.ok a : Except ε α) = .ok a' → β) : + Except.pmap (.ok a) f = .ok (f a rfl) := rfl + +@[simp] theorem pmap_id {ε : Type u} {α : Type v} (x : Except ε α) : + x.pmap (fun a _ => a) = x := by cases x <;> simp + +@[simp] theorem map_pmap (g : β → γ) (x : Except ε α) (f : (a : α) → x = .ok a → β) : + g <$> x.pmap f = x.pmap fun a h => g (f a h) := by + cases x <;> simp + +end Except + +namespace ExceptT + +-- This will be redundant after nightly-2024-11-08. +attribute [ext] ExceptT.ext + +@[simp] theorem run_mk {m : Type u → Type v} (x : m (Except ε α)) : (ExceptT.mk x).run = x := rfl +@[simp] theorem mk_run (x : ExceptT ε m α) : ExceptT.mk (ExceptT.run x) = x := rfl + +@[simp] +theorem map_mk [Monad m] [LawfulMonad m] (f : α → β) (x : m (Except ε α)) : + f <$> ExceptT.mk x = ExceptT.mk ((f <$> · ) <$> x) := by + simp only [Functor.map, Except.map, ExceptT.map, map_eq_pure_bind] + congr + funext a + split <;> simp + +end ExceptT diff --git a/Batteries/Lean/LawfulMonad.lean b/Batteries/Lean/LawfulMonad.lean new file mode 100644 index 0000000000..a69d5ad18f --- /dev/null +++ b/Batteries/Lean/LawfulMonad.lean @@ -0,0 +1,30 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kim Morrison +-/ +import Batteries.Classes.SatisfiesM +import Lean.Elab.Command + +/-! +# Construct `LawfulMonad` instances for the Lean monad stack. +-/ + +open Lean Elab Term Tactic Command + +instance : LawfulMonad (EIO ε) := inferInstanceAs <| LawfulMonad (EStateM _ _) +instance : LawfulMonad BaseIO := inferInstanceAs <| LawfulMonad (EIO _) +instance : LawfulMonad IO := inferInstanceAs <| LawfulMonad (EIO _) + +instance : LawfulMonad (EST ε σ) := inferInstanceAs <| LawfulMonad (EStateM _ _) + +instance : LawfulMonad CoreM := + inferInstanceAs <| LawfulMonad (ReaderT _ <| StateRefT' _ _ (EIO Exception)) +instance : LawfulMonad MetaM := + inferInstanceAs <| LawfulMonad (ReaderT _ <| StateRefT' _ _ CoreM) +instance : LawfulMonad TermElabM := + inferInstanceAs <| LawfulMonad (ReaderT _ <| StateRefT' _ _ MetaM) +instance : LawfulMonad TacticM := + inferInstanceAs <| LawfulMonad (ReaderT _ $ StateRefT' _ _ $ TermElabM) +instance : LawfulMonad CommandElabM := + inferInstanceAs <| LawfulMonad (ReaderT _ $ StateRefT' _ _ $ EIO _) diff --git a/Batteries/Lean/SatisfiesM.lean b/Batteries/Lean/SatisfiesM.lean new file mode 100644 index 0000000000..189e138091 --- /dev/null +++ b/Batteries/Lean/SatisfiesM.lean @@ -0,0 +1,35 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kim Morrison +-/ +import Batteries.Classes.SatisfiesM +import Batteries.Lean.LawfulMonad +import Lean.Elab.Command + +/-! +# Construct `MonadSatisfying` instances for the Lean monad stack. +-/ + +open Lean Elab Term Tactic Command + +instance : MonadSatisfying (EIO ε) := inferInstanceAs <| MonadSatisfying (EStateM _ _) +instance : MonadSatisfying BaseIO := inferInstanceAs <| MonadSatisfying (EIO _) +instance : MonadSatisfying IO := inferInstanceAs <| MonadSatisfying (EIO _) + +instance : MonadSatisfying (EST ε σ) := inferInstanceAs <| MonadSatisfying (EStateM _ _) + +instance : MonadSatisfying CoreM := + inferInstanceAs <| MonadSatisfying (ReaderT _ <| StateRefT' _ _ (EIO _)) + +instance : MonadSatisfying MetaM := + inferInstanceAs <| MonadSatisfying (ReaderT _ <| StateRefT' _ _ CoreM) + +instance : MonadSatisfying TermElabM := + inferInstanceAs <| MonadSatisfying (ReaderT _ <| StateRefT' _ _ MetaM) + +instance : MonadSatisfying TacticM := + inferInstanceAs <| MonadSatisfying (ReaderT _ $ StateRefT' _ _ TermElabM) + +instance : MonadSatisfying CommandElabM := + inferInstanceAs <| MonadSatisfying (ReaderT _ $ StateRefT' _ _ (EIO _)) diff --git a/BatteriesTest/satisfying.lean b/BatteriesTest/satisfying.lean new file mode 100644 index 0000000000..eadfbc6a66 --- /dev/null +++ b/BatteriesTest/satisfying.lean @@ -0,0 +1,8 @@ +import Batteries.Lean.SatisfiesM +import Batteries.Data.Array.Monadic + +open Lean Meta Array Elab Term Tactic Command + +example (xs : Array Expr) : MetaM { ts : Array Expr // ts.size = xs.size } := do + let r ← satisfying (xs.size_mapM inferType) + return r From eeb25d3911f9835cea005eec28f9dac0933dd5b6 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 13 Nov 2024 16:32:00 +1100 Subject: [PATCH 159/177] chore: remove unnecessary import in Batteries.Lean.LawfulMonad (#1042) --- Batteries/Lean/LawfulMonad.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Batteries/Lean/LawfulMonad.lean b/Batteries/Lean/LawfulMonad.lean index a69d5ad18f..03221f5c86 100644 --- a/Batteries/Lean/LawfulMonad.lean +++ b/Batteries/Lean/LawfulMonad.lean @@ -3,7 +3,6 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Batteries.Classes.SatisfiesM import Lean.Elab.Command /-! From 44484f9bbb217b8adbbe58879cd87a09e511e0df Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 13 Nov 2024 16:46:49 +1100 Subject: [PATCH 160/177] chore: remove duplicate NameMap forIn instance (#1043) --- Batteries/Lean/NameMap.lean | 3 --- 1 file changed, 3 deletions(-) diff --git a/Batteries/Lean/NameMap.lean b/Batteries/Lean/NameMap.lean index 7ca7e7acea..0c6d341925 100644 --- a/Batteries/Lean/NameMap.lean +++ b/Batteries/Lean/NameMap.lean @@ -13,9 +13,6 @@ We provide `NameMap.filter` and `NameMap.filterMap`. namespace Lean.NameMap -instance : ForIn m (NameMap β) (Name × β) := - inferInstanceAs <| ForIn m (RBMap Name β _) _ - /-- `filter f m` returns the `NameMap` consisting of all "`key`/`val`"-pairs in `m` where `f key val` returns `true`. From 6d0acd1b53b4b87e41bec1c0229924b9ab1ef84f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 13 Nov 2024 17:08:49 +1100 Subject: [PATCH 161/177] chore: update authors lines (#1045) --- Batteries/Control/Nondet/Basic.lean | 4 ++-- Batteries/Data/HashMap/Lemmas.lean | 2 +- Batteries/Data/MLList/Basic.lean | 4 ++-- Batteries/Data/MLList/Heartbeats.lean | 4 ++-- Batteries/Data/MLList/IO.lean | 4 ++-- Batteries/Lean/Except.lean | 2 +- Batteries/Lean/IO/Process.lean | 4 ++-- Batteries/Lean/Meta/DiscrTree.lean | 2 +- Batteries/Lean/Meta/Simp.lean | 4 ++-- Batteries/Lean/System/IO.lean | 4 ++-- BatteriesTest/Internal/DummyLabelAttr.lean | 2 +- BatteriesTest/omega/benchmark.lean | 2 +- BatteriesTest/show_term.lean | 4 ++-- BatteriesTest/solve_by_elim.lean | 4 ++-- 14 files changed, 23 insertions(+), 23 deletions(-) diff --git a/Batteries/Control/Nondet/Basic.lean b/Batteries/Control/Nondet/Basic.lean index c6d384b9dd..aa616a2857 100644 --- a/Batteries/Control/Nondet/Basic.lean +++ b/Batteries/Control/Nondet/Basic.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2023 Scott Morrison. All rights reserved. +Copyright (c) 2023 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison +Authors: Kim Morrison -/ import Batteries.Tactic.Lint.Misc import Batteries.Data.MLList.Basic diff --git a/Batteries/Data/HashMap/Lemmas.lean b/Batteries/Data/HashMap/Lemmas.lean index 79107fc0bf..2e9df0770d 100644 --- a/Batteries/Data/HashMap/Lemmas.lean +++ b/Batteries/Data/HashMap/Lemmas.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Lean FRO, LLC. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison +Authors: Kim Morrison -/ import Batteries.Data.HashMap.Basic diff --git a/Batteries/Data/MLList/Basic.lean b/Batteries/Data/MLList/Basic.lean index 8f58c11fb8..1ed6d7e256 100644 --- a/Batteries/Data/MLList/Basic.lean +++ b/Batteries/Data/MLList/Basic.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2018 Scott Morrison. All rights reserved. +Copyright (c) 2018 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro, Keeley Hoek, Simon Hudon, Scott Morrison +Authors: Mario Carneiro, Keeley Hoek, Simon Hudon, Kim Morrison -/ /-! # Monadic lazy lists. diff --git a/Batteries/Data/MLList/Heartbeats.lean b/Batteries/Data/MLList/Heartbeats.lean index d898bbdb30..21cb18dfa0 100644 --- a/Batteries/Data/MLList/Heartbeats.lean +++ b/Batteries/Data/MLList/Heartbeats.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2023 Scott Morrison. All rights reserved. +Copyright (c) 2023 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison +Authors: Kim Morrison -/ import Batteries.Data.MLList.Basic import Lean.Util.Heartbeats diff --git a/Batteries/Data/MLList/IO.lean b/Batteries/Data/MLList/IO.lean index f230fa2184..1153e1fbc6 100644 --- a/Batteries/Data/MLList/IO.lean +++ b/Batteries/Data/MLList/IO.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2023 Scott Morrison. All rights reserved. +Copyright (c) 2023 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison +Authors: Kim Morrison -/ import Batteries.Lean.System.IO import Batteries.Data.MLList.Basic diff --git a/Batteries/Lean/Except.lean b/Batteries/Lean/Except.lean index e0eb4c79f5..214ae74364 100644 --- a/Batteries/Lean/Except.lean +++ b/Batteries/Lean/Except.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Lean FRO, LLC. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison +Authors: Kim Morrison -/ import Lean.Util.Trace diff --git a/Batteries/Lean/IO/Process.lean b/Batteries/Lean/IO/Process.lean index f5eb335b7f..4c21582912 100644 --- a/Batteries/Lean/IO/Process.lean +++ b/Batteries/Lean/IO/Process.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2023 Scott Morrison. All rights reserved. +Copyright (c) 2023 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison +Authors: Kim Morrison -/ /-! diff --git a/Batteries/Lean/Meta/DiscrTree.lean b/Batteries/Lean/Meta/DiscrTree.lean index a7a49bd54e..ed86bce7ce 100644 --- a/Batteries/Lean/Meta/DiscrTree.lean +++ b/Batteries/Lean/Meta/DiscrTree.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2022 Jannis Limperg. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jannis Limperg, Scott Morrison +Authors: Jannis Limperg, Kim Morrison -/ import Lean.Meta.DiscrTree diff --git a/Batteries/Lean/Meta/Simp.lean b/Batteries/Lean/Meta/Simp.lean index d968ae9328..7ea81e87fa 100644 --- a/Batteries/Lean/Meta/Simp.lean +++ b/Batteries/Lean/Meta/Simp.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Scott Morrison. All rights reserved. +Copyright (c) 2022 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison, Gabriel Ebner, Floris van Doorn +Authors: Kim Morrison, Gabriel Ebner, Floris van Doorn -/ import Lean.Elab.Tactic.Simp import Batteries.Tactic.OpenPrivate diff --git a/Batteries/Lean/System/IO.lean b/Batteries/Lean/System/IO.lean index 40dcda5365..d0848bf9cd 100644 --- a/Batteries/Lean/System/IO.lean +++ b/Batteries/Lean/System/IO.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2023 Scott Morrison. All rights reserved. +Copyright (c) 2023 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison +Authors: Kim Morrison -/ /-! diff --git a/BatteriesTest/Internal/DummyLabelAttr.lean b/BatteriesTest/Internal/DummyLabelAttr.lean index 9cd76815cc..d996128a73 100644 --- a/BatteriesTest/Internal/DummyLabelAttr.lean +++ b/BatteriesTest/Internal/DummyLabelAttr.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Lean FRO, LLC. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison +Authors: Kim Morrison -/ import Lean.LabelAttribute diff --git a/BatteriesTest/omega/benchmark.lean b/BatteriesTest/omega/benchmark.lean index f3a5113dd1..c1e50b1a2a 100644 --- a/BatteriesTest/omega/benchmark.lean +++ b/BatteriesTest/omega/benchmark.lean @@ -4,7 +4,7 @@ As it's important that `omega` is fast, particularly when it has nothing to do, this file maintains a benchmark suite for `omega`. It is particularly low-tech, -and currently only reproducible on Scott Morrison's FRO machine; +and currently only reproducible on Kim Morrison's FRO machine; nevertheless it seems useful to keep the benchmark history in the repository. The benchmark file consists of the test suite from `omega`'s initial release, diff --git a/BatteriesTest/show_term.lean b/BatteriesTest/show_term.lean index 8557e8229b..7288c369e9 100644 --- a/BatteriesTest/show_term.lean +++ b/BatteriesTest/show_term.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Scott Morrison. All rights reserved. +Copyright (c) 2021 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison +Authors: Kim Morrison -/ /-- info: Try this: exact (n, 37) -/ diff --git a/BatteriesTest/solve_by_elim.lean b/BatteriesTest/solve_by_elim.lean index 6bfe72bf02..cdad389de9 100644 --- a/BatteriesTest/solve_by_elim.lean +++ b/BatteriesTest/solve_by_elim.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Scott Morrison. All rights reserved. +Copyright (c) 2021 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison +Authors: Kim Morrison -/ import Batteries.Tactic.PermuteGoals import BatteriesTest.Internal.DummyLabelAttr From e2b30dce0e92d4fd12de8dec3e8194c36b4ed36b Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 14 Nov 2024 10:21:05 +1100 Subject: [PATCH 162/177] chore: remove >6 month old deprecations (#1044) --- Batteries.lean | 2 -- Batteries/Data/Array/Merge.lean | 17 --------- Batteries/Data/List/Basic.lean | 4 --- Batteries/Data/List/Lemmas.lean | 3 -- Batteries/Data/RBMap/Basic.lean | 9 ----- Batteries/Lean/Delaborator.lean | 13 ------- Batteries/Lean/HashSet.lean | 3 -- Batteries/Logic.lean | 2 -- Batteries/StdDeprecations.lean | 63 --------------------------------- 9 files changed, 116 deletions(-) delete mode 100644 Batteries/Lean/Delaborator.lean delete mode 100644 Batteries/StdDeprecations.lean diff --git a/Batteries.lean b/Batteries.lean index e6464503cb..b645334bf4 100644 --- a/Batteries.lean +++ b/Batteries.lean @@ -37,7 +37,6 @@ import Batteries.Data.UInt import Batteries.Data.UnionFind import Batteries.Data.Vector import Batteries.Lean.AttributeExtra -import Batteries.Lean.Delaborator import Batteries.Lean.EStateM import Batteries.Lean.Except import Batteries.Lean.Expr @@ -70,7 +69,6 @@ import Batteries.Linter import Batteries.Linter.UnnecessarySeqFocus import Batteries.Linter.UnreachableTactic import Batteries.Logic -import Batteries.StdDeprecations import Batteries.Tactic.Alias import Batteries.Tactic.Basic import Batteries.Tactic.Case diff --git a/Batteries/Data/Array/Merge.lean b/Batteries/Data/Array/Merge.lean index 4b806a35d6..70034111ad 100644 --- a/Batteries/Data/Array/Merge.lean +++ b/Batteries/Data/Array/Merge.lean @@ -26,11 +26,6 @@ where if lt x y then go (acc.push x) (i + 1) j else go (acc.push y) i (j + 1) termination_by xs.size + ys.size - (i + j) -set_option linter.unusedVariables false in -@[deprecated merge (since := "2024-04-24"), inherit_doc merge] -def mergeSortedPreservingDuplicates [ord : Ord α] (xs ys : Array α) : Array α := - merge (compare · · |>.isLT) xs ys - -- We name `ord` so it can be provided as a named argument. set_option linter.unusedVariables.funArgs false in /-- @@ -57,8 +52,6 @@ where | .eq => go (acc.push (merge x y)) (i + 1) (j + 1) termination_by xs.size + ys.size - (i + j) -@[deprecated (since := "2024-04-24")] alias mergeSortedMergingDuplicates := mergeDedupWith - /-- `O(|xs| + |ys|)`. Merge arrays `xs` and `ys`, which must be sorted according to `compare` and must not contain duplicates. If an element appears in both `xs` and `ys`, only one copy is kept. @@ -66,8 +59,6 @@ not contain duplicates. If an element appears in both `xs` and `ys`, only one co @[inline] def mergeDedup [ord : Ord α] (xs ys : Array α) : Array α := mergeDedupWith (ord := ord) xs ys fun x _ => x -@[deprecated (since := "2024-04-24")] alias mergeSortedDeduplicating := mergeDedup - set_option linter.unusedVariables false in /-- `O(|xs| * |ys|)`. Merge `xs` and `ys`, which do not need to be sorted. Elements which occur in @@ -85,8 +76,6 @@ where ys.foldl (init := xs) fun xs y => if xs.any (· == y) (stop := xsSize) then xs else xs.push y -@[deprecated (since := "2024-04-24")] alias mergeUnsortedDeduplicating := mergeUnsortedDedup - -- We name `eq` so it can be provided as a named argument. set_option linter.unusedVariables.funArgs false in /-- @@ -105,8 +94,6 @@ where acc.push hd termination_by xs.size - i -@[deprecated (since := "2024-04-24")] alias mergeAdjacentDuplicates := mergeAdjacentDups - /-- `O(|xs|)`. Deduplicate a sorted array. The array must be sorted with to an order which agrees with `==`, i.e. whenever `x == y` then `compare x y == .eq`. @@ -114,13 +101,9 @@ where def dedupSorted [eq : BEq α] (xs : Array α) : Array α := xs.mergeAdjacentDups (eq := eq) fun x _ => x -@[deprecated (since := "2024-04-24")] alias deduplicateSorted := dedupSorted - /-- `O(|xs| log |xs|)`. Sort and deduplicate an array. -/ def sortDedup [ord : Ord α] (xs : Array α) : Array α := have := ord.toBEq dedupSorted <| xs.qsort (compare · · |>.isLT) -@[deprecated (since := "2024-04-24")] alias sortAndDeduplicate := sortDedup - end Array diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index 3e2460f18c..9c0ab14d46 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -53,10 +53,6 @@ drop_while (· != 1) [0, 1, 2, 3] = [1, 2, 3] | [] => [] | x :: xs => bif p x then xs else after p xs -@[deprecated (since := "2024-05-06")] alias removeNth := eraseIdx -@[deprecated (since := "2024-05-06")] alias removeNthTR := eraseIdxTR -@[deprecated (since := "2024-05-06")] alias removeNth_eq_removeNthTR := eraseIdx_eq_eraseIdxTR - /-- Replaces the first element of the list for which `f` returns `some` with the returned value. -/ @[simp] def replaceF (f : α → Option α) : List α → List α | [] => [] diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index 89f1de368d..c8fba52299 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -524,7 +524,6 @@ theorem get?_inj @[deprecated (since := "2024-10-21")] alias modifyNth_succ_cons := modify_succ_cons @[deprecated (since := "2024-10-21")] alias modifyNthTail_id := modifyTailIdx_id @[deprecated (since := "2024-10-21")] alias eraseIdx_eq_modifyNthTail := eraseIdx_eq_modifyTailIdx -@[deprecated (since := "2024-05-06")] alias removeNth_eq_nth_tail := eraseIdx_eq_modifyTailIdx @[deprecated (since := "2024-10-21")] alias getElem?_modifyNth := getElem?_modify @[deprecated getElem?_modify (since := "2024-06-12")] theorem get?_modifyNth (f : α → α) (n) (l : List α) (m) : @@ -574,5 +573,3 @@ theorem get?_set_ne (a : α) {m n} (l : List α) (h : m ≠ n) : (set l m a).get theorem get?_set (a : α) {m n} (l : List α) : (set l m a).get? n = if m = n then (fun _ => a) <$> l.get? n else l.get? n := by simp [getElem?_set']; rfl -@[deprecated (since := "2024-05-06")] alias length_removeNth := length_eraseIdx -@[deprecated (since := "2024-04-22")] alias sublist.erase := Sublist.erase diff --git a/Batteries/Data/RBMap/Basic.lean b/Batteries/Data/RBMap/Basic.lean index 8eb202591c..e2c9a32c5f 100644 --- a/Batteries/Data/RBMap/Basic.lean +++ b/Batteries/Data/RBMap/Basic.lean @@ -67,9 +67,6 @@ protected def max? : RBNode α → Option α | node _ _ v nil => some v | node _ _ _ r => r.max? -@[deprecated (since := "2024-04-17")] protected alias min := RBNode.min? -@[deprecated (since := "2024-04-17")] protected alias max := RBNode.max? - /-- Fold a function in tree order along the nodes. `v₀` is used at `nil` nodes and `f` is used to combine results at branching nodes. @@ -670,9 +667,6 @@ instance : ToStream (RBSet α cmp) (RBNode.Stream α) := ⟨fun x => x.1.toStrea /-- `O(log n)`. Returns the entry `a` such that `a ≥ k` for all keys in the RBSet. -/ @[inline] protected def max? (t : RBSet α cmp) : Option α := t.1.max? -@[deprecated (since := "2024-04-17")] protected alias min := RBSet.min? -@[deprecated (since := "2024-04-17")] protected alias max := RBSet.max? - instance [Repr α] : Repr (RBSet α cmp) where reprPrec m prec := Repr.addAppParen ("RBSet.ofList " ++ repr m.toList) prec @@ -1055,9 +1049,6 @@ instance : Stream (Values.Stream α β) β := ⟨Values.Stream.next?⟩ /-- `O(log n)`. Returns the key-value pair `(a, b)` such that `a ≥ k` for all keys in the RBMap. -/ @[inline] protected def max? : RBMap α β cmp → Option (α × β) := RBSet.max? -@[deprecated (since := "2024-04-17")] protected alias min := RBMap.min? -@[deprecated (since := "2024-04-17")] protected alias max := RBMap.max? - instance [Repr α] [Repr β] : Repr (RBMap α β cmp) where reprPrec m prec := Repr.addAppParen ("RBMap.ofList " ++ repr m.toList) prec diff --git a/Batteries/Lean/Delaborator.lean b/Batteries/Lean/Delaborator.lean deleted file mode 100644 index e09f2198de..0000000000 --- a/Batteries/Lean/Delaborator.lean +++ /dev/null @@ -1,13 +0,0 @@ -/- -Copyright (c) 2022 Mario Carneiro. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro, Kyle Miller --/ -import Lean.PrettyPrinter - -open Lean PrettyPrinter Delaborator SubExpr - -/-- Abbreviation for `Lean.MessageData.ofConst`. -/ -@[deprecated Lean.MessageData.ofConst (since := "2024-05-18")] -def Lean.ppConst (e : Expr) : MessageData := - Lean.MessageData.ofConst e diff --git a/Batteries/Lean/HashSet.lean b/Batteries/Lean/HashSet.lean index 2a6ee47909..d0afac657b 100644 --- a/Batteries/Lean/HashSet.lean +++ b/Batteries/Lean/HashSet.lean @@ -10,9 +10,6 @@ namespace Std.HashSet variable [BEq α] [Hashable α] -instance : Singleton α (HashSet α) := ⟨fun x => HashSet.empty.insert x⟩ -instance : Insert α (HashSet α) := ⟨fun a s => s.insert a⟩ - /-- `O(n)`. Returns `true` if `f` returns `true` for any element of the set. -/ diff --git a/Batteries/Logic.lean b/Batteries/Logic.lean index 4bc5f9f102..0be01f3b28 100644 --- a/Batteries/Logic.lean +++ b/Batteries/Logic.lean @@ -8,8 +8,6 @@ import Batteries.Tactic.Alias instance {f : α → β} [DecidablePred p] : DecidablePred (p ∘ f) := inferInstanceAs <| DecidablePred fun x => p (f x) -@[deprecated (since := "2024-03-15")] alias proofIrrel := proof_irrel - /-! ## id -/ theorem Function.id_def : @id α = fun x => x := rfl diff --git a/Batteries/StdDeprecations.lean b/Batteries/StdDeprecations.lean deleted file mode 100644 index 74a00f4f5a..0000000000 --- a/Batteries/StdDeprecations.lean +++ /dev/null @@ -1,63 +0,0 @@ -/- -Copyright (c) 2024 Kim Morrison. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kim Morrison --/ -import Batteries.Tactic.Alias -import Batteries.Data.DList.Basic -import Batteries.Data.PairingHeap -import Batteries.Data.BinomialHeap.Basic -import Batteries.Data.HashMap.Basic -import Batteries.Data.RBMap.Basic -import Batteries.Data.UnionFind.Basic - -/-! -# We set up deprecations for identifiers formerly in the `Std` namespace. - -Note that we have not moved anything in the `Std.CodeAction` or `Std.Linter` namespace. - -These deprecations do not cover `Std.Tactic`, the contents of which has been moved, -but it would be much harder to generate the deprecations. -Let's hope that people using the tactic implementations can work this out themselves. --/ - -@[deprecated (since := "2024-05-07")] alias Std.AssocList := Batteries.AssocList -@[deprecated (since := "2024-05-07")] alias Std.mkHashMap := Batteries.mkHashMap -@[deprecated (since := "2024-05-07")] alias Std.DList := Batteries.DList -@[deprecated (since := "2024-05-07")] alias Std.PairingHeapImp.Heap := Batteries.PairingHeapImp.Heap -@[deprecated (since := "2024-05-07")] alias Std.TotalBLE := Batteries.TotalBLE -@[deprecated (since := "2024-05-07")] alias Std.OrientedCmp := Batteries.OrientedCmp -@[deprecated (since := "2024-05-07")] alias Std.TransCmp := Batteries.TransCmp -@[deprecated (since := "2024-05-07")] alias Std.BEqCmp := Batteries.BEqCmp -@[deprecated (since := "2024-05-07")] alias Std.LTCmp := Batteries.LTCmp -@[deprecated (since := "2024-05-07")] alias Std.LECmp := Batteries.LECmp -@[deprecated (since := "2024-05-07")] alias Std.LawfulCmp := Batteries.LawfulCmp -@[deprecated (since := "2024-05-07")] alias Std.OrientedOrd := Batteries.OrientedOrd -@[deprecated (since := "2024-05-07")] alias Std.TransOrd := Batteries.TransOrd -@[deprecated (since := "2024-05-07")] alias Std.BEqOrd := Batteries.BEqOrd -@[deprecated (since := "2024-05-07")] alias Std.LTOrd := Batteries.LTOrd -@[deprecated (since := "2024-05-07")] alias Std.LEOrd := Batteries.LEOrd -@[deprecated (since := "2024-05-07")] alias Std.LawfulOrd := Batteries.LawfulOrd -@[deprecated (since := "2024-05-07")] -alias Std.compareOfLessAndEq_eq_lt := Batteries.compareOfLessAndEq_eq_lt -@[deprecated (since := "2024-05-07")] alias Std.RBColor := Batteries.RBColor -@[deprecated (since := "2024-05-07")] alias Std.RBNode := Batteries.RBNode -@[deprecated (since := "2024-05-07")] alias Std.RBSet := Batteries.RBSet -@[deprecated (since := "2024-05-07")] alias Std.mkRBSet := Batteries.mkRBSet -@[deprecated (since := "2024-05-07")] alias Std.RBMap := Batteries.RBMap -@[deprecated (since := "2024-05-07")] alias Std.mkRBMap := Batteries.mkRBMap -@[deprecated (since := "2024-05-07")] alias Std.BinomialHeap := Batteries.BinomialHeap -@[deprecated (since := "2024-05-07")] alias Std.mkBinomialHeap := Batteries.mkBinomialHeap -@[deprecated (since := "2024-05-07")] alias Std.UFNode := Batteries.UFNode -@[deprecated (since := "2024-05-07")] alias Std.UnionFind := Batteries.UnionFind - --- Check that these generate usable deprecated hints --- when referring to names inside these namespaces. -set_option warningAsError true in -/-- -error: `Std.UnionFind` has been deprecated, use `Batteries.UnionFind` instead ---- -error: unknown constant 'Std.UnionFind.find' --/ -#guard_msgs in -#eval Std.UnionFind.find From b100ff2565805e9f30a482788b3fc66937a7f38a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 14 Nov 2024 13:28:59 +1100 Subject: [PATCH 163/177] chore: deprecate Expr.forallArity, per todo (#1046) --- Batteries/Lean/Expr.lean | 13 ++----------- 1 file changed, 2 insertions(+), 11 deletions(-) diff --git a/Batteries/Lean/Expr.lean b/Batteries/Lean/Expr.lean index 9ac9a09238..98781ab0aa 100644 --- a/Batteries/Lean/Expr.lean +++ b/Batteries/Lean/Expr.lean @@ -27,17 +27,8 @@ def toSyntax (e : Expr) : TermElabM Syntax.Term := withFreshMacroScope do @[deprecated (since := "2024-10-16"), inherit_doc getNumHeadLambdas] abbrev lambdaArity := @getNumHeadLambdas -/-- -Returns the number of leading `∀` binders of an expression. Ignores metadata. --/ -def forallArity : Expr → Nat - | mdata _ b => forallArity b - | forallE _ _ body _ => 1 + forallArity body - | _ => 0 - --- TODO: replace `forallArity` after https://github.com/leanprover/lean4/pull/5729 --- @[deprecated (since := "2024-10-16"), inherit_doc getNumHeadForalls] --- abbrev forallArity := @getNumHeadForalls +@[deprecated (since := "2024-11-13"), inherit_doc getNumHeadForalls] +abbrev forallArity := @getNumHeadForalls /-- Like `withApp` but ignores metadata. -/ @[inline] From 9dcf29471c06329e21b9e4ceba391601bec4c6e3 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 15 Nov 2024 03:48:59 +0000 Subject: [PATCH 164/177] fix: run `Lean.enableInitializersExecution` (#1047) --- scripts/runLinter.lean | 50 +++++++++++++++++++++--------------------- 1 file changed, 25 insertions(+), 25 deletions(-) diff --git a/scripts/runLinter.lean b/scripts/runLinter.lean index b019a36dda..2a88e18eb5 100644 --- a/scripts/runLinter.lean +++ b/scripts/runLinter.lean @@ -104,30 +104,31 @@ unsafe def runLinterOnModule (update : Bool) (module : Name): IO Unit := do readJsonFile NoLints nolintsFile else pure #[] - withImportModules #[module, lintModule] {} (trustLevel := 1024) fun env => - let ctx := { fileName := "", fileMap := default } - let state := { env } - Prod.fst <$> (CoreM.toIO · ctx state) do - let decls ← getDeclsInPackage module.getRoot - let linters ← getChecks (slow := true) (runAlways := none) (runOnly := none) - let results ← lintCore decls linters - if update then - writeJsonFile (α := NoLints) nolintsFile <| - .qsort (lt := fun (a, b) (c, d) => a.lt c || (a == c && b.lt d)) <| - .flatten <| results.map fun (linter, decls) => - decls.fold (fun res decl _ => res.push (linter.name, decl)) #[] - let results := results.map fun (linter, decls) => - .mk linter <| nolints.foldl (init := decls) fun decls (linter', decl') => - if linter.name == linter' then decls.erase decl' else decls - let failed := results.any (!·.2.isEmpty) - if failed then - let fmtResults ← - formatLinterResults results decls (groupByFilename := true) (useErrorFormat := true) - s!"in {module}" (runSlowLinters := true) .medium linters.size - IO.print (← fmtResults.toString) - IO.Process.exit 1 - else - IO.println "-- Linting passed for {module}." + unsafe Lean.enableInitializersExecution + let env ← importModules #[module, lintModule] {} (trustLevel := 1024) + let ctx := { fileName := "", fileMap := default } + let state := { env } + Prod.fst <$> (CoreM.toIO · ctx state) do + let decls ← getDeclsInPackage module.getRoot + let linters ← getChecks (slow := true) (runAlways := none) (runOnly := none) + let results ← lintCore decls linters + if update then + writeJsonFile (α := NoLints) nolintsFile <| + .qsort (lt := fun (a, b) (c, d) => a.lt c || (a == c && b.lt d)) <| + .flatten <| results.map fun (linter, decls) => + decls.fold (fun res decl _ => res.push (linter.name, decl)) #[] + let results := results.map fun (linter, decls) => + .mk linter <| nolints.foldl (init := decls) fun decls (linter', decl') => + if linter.name == linter' then decls.erase decl' else decls + let failed := results.any (!·.2.isEmpty) + if failed then + let fmtResults ← + formatLinterResults results decls (groupByFilename := true) (useErrorFormat := true) + s!"in {module}" (runSlowLinters := true) .medium linters.size + IO.print (← fmtResults.toString) + IO.Process.exit 1 + else + IO.println s!"-- Linting passed for {module}." /-- Usage: `runLinter [--update] [Batteries.Data.Nat.Basic]` @@ -149,4 +150,3 @@ unsafe def main (args : List String) : IO Unit := do let modulesToLint ← determineModulesToLint specifiedModule modulesToLint.forM <| runLinterOnModule update - From 01f4969b6e861db6a99261ea5eadd5a9bb63011b Mon Sep 17 00:00:00 2001 From: David Thrane Christiansen Date: Fri, 15 Nov 2024 08:55:10 +0100 Subject: [PATCH 165/177] fix: use two tokens for open private/export private (#1050) --- Batteries/Tactic/OpenPrivate.lean | 4 ++-- BatteriesTest/OpenPrivateDefs.lean | 4 ++++ BatteriesTest/openPrivate.lean | 37 ++++++++++++++++++++++++++++++ 3 files changed, 43 insertions(+), 2 deletions(-) create mode 100644 BatteriesTest/OpenPrivateDefs.lean create mode 100644 BatteriesTest/openPrivate.lean diff --git a/Batteries/Tactic/OpenPrivate.lean b/Batteries/Tactic/OpenPrivate.lean index 7ade0b94d9..a7e126133f 100644 --- a/Batteries/Tactic/OpenPrivate.lean +++ b/Batteries/Tactic/OpenPrivate.lean @@ -99,7 +99,7 @@ name component. It is also possible to specify the module instead with `open private a b c from Other.Module`. -/ -syntax (name := openPrivate) "open private" (ppSpace ident)* +syntax (name := openPrivate) "open" ppSpace "private" (ppSpace ident)* (" in" (ppSpace ident)*)? (" from" (ppSpace ident)*)? : command /-- Elaborator for `open private`. -/ @@ -119,7 +119,7 @@ It will also open the newly created alias definition under the provided short na It is also possible to specify the module instead with `export private a b c from Other.Module`. -/ -syntax (name := exportPrivate) "export private" (ppSpace ident)* +syntax (name := exportPrivate) "export" ppSpace "private" (ppSpace ident)* (" in" (ppSpace ident)*)? (" from" (ppSpace ident)*)? : command /-- Elaborator for `export private`. -/ diff --git a/BatteriesTest/OpenPrivateDefs.lean b/BatteriesTest/OpenPrivateDefs.lean new file mode 100644 index 0000000000..f5d1a483e4 --- /dev/null +++ b/BatteriesTest/OpenPrivateDefs.lean @@ -0,0 +1,4 @@ +/-! +This file contains a private declaration. It's tested in `openPrivate.lean`. +-/ +private def secretNumber : Nat := 2 diff --git a/BatteriesTest/openPrivate.lean b/BatteriesTest/openPrivate.lean new file mode 100644 index 0000000000..36f042c634 --- /dev/null +++ b/BatteriesTest/openPrivate.lean @@ -0,0 +1,37 @@ + +import Batteries.Tactic.OpenPrivate + +import BatteriesTest.OpenPrivateDefs + + + +/-- error: unknown identifier 'secretNumber' -/ +#guard_msgs in +#eval secretNumber + + +-- It works with one space between the tokens +/-- info: 2 -/ +#guard_msgs in +open private secretNumber from BatteriesTest.OpenPrivateDefs in +#eval secretNumber + + +-- It also works with other kinds of whitespace between the tokens + +/-- info: 2 -/ +#guard_msgs in +open private secretNumber from BatteriesTest.OpenPrivateDefs in +#eval secretNumber + + +/-- info: 2 -/ +#guard_msgs in +open + private secretNumber from BatteriesTest.OpenPrivateDefs in +#eval secretNumber + +/-- info: 2 -/ +#guard_msgs in +open /- Being sneaky! -/ private secretNumber from BatteriesTest.OpenPrivateDefs in +#eval secretNumber From 485efbc439ee0ebdeae8afb0acd24a5e82e2f771 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 19 Nov 2024 20:53:19 +1100 Subject: [PATCH 166/177] chore: remove @[simp] attributes from monad-specific SatisfiesM lemmas (#1054) --- Batteries/Classes/SatisfiesM.lean | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/Batteries/Classes/SatisfiesM.lean b/Batteries/Classes/SatisfiesM.lean index 834cbbea68..9b34cf475b 100644 --- a/Batteries/Classes/SatisfiesM.lean +++ b/Batteries/Classes/SatisfiesM.lean @@ -185,14 +185,15 @@ theorem SatisfiesM_EStateM_eq : rw [EStateM.run_map, EStateM.run] split <;> simp_all -@[simp] theorem SatisfiesM_ReaderT_eq [Monad m] : +theorem SatisfiesM_ReaderT_eq [Monad m] : SatisfiesM (m := ReaderT ρ m) p x ↔ ∀ s, SatisfiesM p (x.run s) := (exists_congr fun a => by exact ⟨fun eq _ => eq ▸ rfl, funext⟩).trans Classical.skolem.symm theorem SatisfiesM_StateRefT_eq [Monad m] : - SatisfiesM (m := StateRefT' ω σ m) p x ↔ ∀ s, SatisfiesM p (x s) := by simp [ReaderT.run] + SatisfiesM (m := StateRefT' ω σ m) p x ↔ ∀ s, SatisfiesM p (x s) := by + simp [SatisfiesM_ReaderT_eq, ReaderT.run] -@[simp] theorem SatisfiesM_StateT_eq [Monad m] [LawfulMonad m] : +theorem SatisfiesM_StateT_eq [Monad m] [LawfulMonad m] : SatisfiesM (m := StateT ρ m) (α := α) p x ↔ ∀ s, SatisfiesM (m := m) (p ·.1) (x.run s) := by change SatisfiesM (m := StateT ρ m) (α := α) p x ↔ ∀ s, SatisfiesM (m := m) (p ·.1) (x s) refine .trans ⟨fun ⟨f, eq⟩ => eq ▸ ?_, fun ⟨f, h⟩ => ?_⟩ Classical.skolem.symm @@ -201,7 +202,7 @@ theorem SatisfiesM_StateRefT_eq [Monad m] : · refine ⟨fun s => (fun ⟨⟨a, s'⟩, h⟩ => ⟨⟨a, h⟩, s'⟩) <$> f s, funext fun s => ?_⟩ show _ >>= _ = _; simp [← h] -@[simp] theorem SatisfiesM_ExceptT_eq [Monad m] [LawfulMonad m] : +theorem SatisfiesM_ExceptT_eq [Monad m] [LawfulMonad m] : SatisfiesM (m := ExceptT ρ m) (α := α) p x ↔ SatisfiesM (m := m) (∀ a, · = .ok a → p a) x.run := by change _ ↔ SatisfiesM (m := m) (∀ a, · = .ok a → p a) x From a822446d61ad7e7f5e843365c7041c326553050a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 21 Nov 2024 15:53:39 +1100 Subject: [PATCH 167/177] chore: cleanup breaking non-terminal simps (#1057) --- Batteries/Data/Range/Lemmas.lean | 18 +++++++++++++----- 1 file changed, 13 insertions(+), 5 deletions(-) diff --git a/Batteries/Data/Range/Lemmas.lean b/Batteries/Data/Range/Lemmas.lean index a075e0b6b4..3fd893daa9 100644 --- a/Batteries/Data/Range/Lemmas.lean +++ b/Batteries/Data/Range/Lemmas.lean @@ -57,12 +57,15 @@ theorem forIn'_eq_forIn_range' [Monad m] (r : Std.Range) suffices ∀ H, forIn' [start:stop:step] init f = forIn (L.pmap Subtype.mk H) init f' from this _ intro H; dsimp only [forIn', Range.forIn'] if h : start < stop then - simp [numElems, Nat.not_le.2 h, L]; split + simp only [numElems, Nat.not_le.2 h, ↓reduceIte, L]; split · subst step suffices ∀ n H init, forIn'.loop start stop 0 f n start (Nat.le_refl _) init = forIn ((List.range' start n 0).pmap Subtype.mk H) init f' from this _ .. - intro n; induction n with (intro H init; unfold forIn'.loop; simp [*]) + intro n; induction n with + (intro H init + unfold forIn'.loop + simp only [↓reduceDIte, Nat.add_zero, List.pmap, List.forIn_cons, List.forIn_nil, h]) | succ n ih => simp [ih (List.forall_mem_cons.1 H).2]; rfl · next step0 => have hstep := Nat.pos_of_ne_zero step0 @@ -84,10 +87,15 @@ theorem forIn'_eq_forIn_range' [Monad m] (r : Std.Range) have ih := ih _ _ (Nat.le_trans hle (Nat.le_add_right ..)) (List.forall_mem_cons.1 H).2 (Nat.le_of_succ_le_succ h1) fun i => by rw [Nat.add_right_comm, Nat.add_assoc, ← Nat.mul_succ, h2, Nat.succ_le_succ_iff] - have := h2 0; simp at this - rw [forIn'.loop]; simp [this, ih]; rfl + have := h2 0 + simp only [Nat.mul_zero, Nat.add_zero, Nat.le_zero_eq, Nat.add_one_ne_zero, iff_false, + Nat.not_le] at this + rw [forIn'.loop] + simp only [this, ↓reduceDIte, ih, List.pmap, List.forIn_cons] + rfl else - simp [List.range', h, numElems_stop_le_start ⟨start, stop, step⟩ (Nat.not_lt.1 h), L] + simp only [numElems_stop_le_start ⟨start, stop, step⟩ (Nat.not_lt.1 h), List.range'_zero, + List.pmap, List.forIn_nil, L] cases stop <;> unfold forIn'.loop <;> simp [List.forIn', h] theorem forIn_eq_forIn_range' [Monad m] (r : Std.Range) From 33d7f346440869364a2cd077bde8cebf73243aaa Mon Sep 17 00:00:00 2001 From: Jared Corduan Date: Thu, 21 Nov 2024 07:31:31 -0500 Subject: [PATCH 168/177] fix: Readme.md needs "git" token in lakefile.lean (#1056) --- README.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/README.md b/README.md index 2873f65536..e20b87c09f 100644 --- a/README.md +++ b/README.md @@ -6,7 +6,7 @@ The "batteries included" extended library for Lean 4. This is a collection of da To use `batteries` in your project, add the following to your `lakefile.lean`: ```lean -require "leanprover-community" / "batteries" @ "main" +require "leanprover-community" / "batteries" @ git "main" ``` Or add the following to your `lakefile.toml`: ```toml From 8b52587ff32e2e443cce6109b5305341289339e7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Fri, 22 Nov 2024 19:38:59 -0500 Subject: [PATCH 169/177] chore: fix workflow token (#1059) --- .github/workflows/labels-from-status.yml | 3 +++ 1 file changed, 3 insertions(+) diff --git a/.github/workflows/labels-from-status.yml b/.github/workflows/labels-from-status.yml index 8ad37be1ae..147af67f34 100644 --- a/.github/workflows/labels-from-status.yml +++ b/.github/workflows/labels-from-status.yml @@ -34,6 +34,7 @@ jobs: if: github.event.pull_request.state == 'closed' uses: actions-ecosystem/action-remove-labels@v1 with: + github_token: ${{ secrets.GITHUB_TOKEN }} labels: | WIP awaiting-author @@ -48,6 +49,7 @@ jobs: ! contains(github.event.pull_request.labels.*.name, 'WIP') uses: actions-ecosystem/action-add-labels@v1 with: + github_token: ${{ secrets.GITHUB_TOKEN }} labels: WIP - name: Label unlabeled other PR as awaiting-review @@ -59,4 +61,5 @@ jobs: ! contains(github.event.pull_request.labels.*.name, 'WIP') uses: actions-ecosystem/action-add-labels@v1 with: + github_token: ${{ secrets.GITHUB_TOKEN }} labels: awaiting-review From 0dc51ac7947ff6aa2c16bcffb64c46c7149d1276 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 23 Nov 2024 20:03:43 +1100 Subject: [PATCH 170/177] chore: fix two deprecations (#1061) --- Batteries/Lean/Expr.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Batteries/Lean/Expr.lean b/Batteries/Lean/Expr.lean index 98781ab0aa..17761cfaa0 100644 --- a/Batteries/Lean/Expr.lean +++ b/Batteries/Lean/Expr.lean @@ -24,10 +24,10 @@ def toSyntax (e : Expr) : TermElabM Syntax.Term := withFreshMacroScope do mvar.mvarId!.assign e pure stx -@[deprecated (since := "2024-10-16"), inherit_doc getNumHeadLambdas] +@[deprecated getNumHeadLambdas (since := "2024-10-16"), inherit_doc getNumHeadLambdas] abbrev lambdaArity := @getNumHeadLambdas -@[deprecated (since := "2024-11-13"), inherit_doc getNumHeadForalls] +@[deprecated getNumHeadForalls(since := "2024-11-13"), inherit_doc getNumHeadForalls] abbrev forallArity := @getNumHeadForalls /-- Like `withApp` but ignores metadata. -/ From 44e2d2e643fd2618b01f9a0592d7dcbd3ffa22de Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Sun, 24 Nov 2024 18:42:43 -0500 Subject: [PATCH 171/177] feat: more vector lemmas (#1062) --- Batteries/Data/Vector/Basic.lean | 7 +- Batteries/Data/Vector/Lemmas.lean | 261 ++++++++++++++++++++++++------ 2 files changed, 215 insertions(+), 53 deletions(-) diff --git a/Batteries/Data/Vector/Basic.lean b/Batteries/Data/Vector/Basic.lean index 0ae5871356..0273c93c18 100644 --- a/Batteries/Data/Vector/Basic.lean +++ b/Batteries/Data/Vector/Basic.lean @@ -9,6 +9,7 @@ import Batteries.Data.List.Basic import Batteries.Data.List.Lemmas import Batteries.Tactic.Alias import Batteries.Tactic.Lint.Misc +import Batteries.Tactic.PrintPrefix /-! # Vectors @@ -252,7 +253,7 @@ Compares two vectors of the same size using a given boolean relation `r`. `isEqv `true` if and only if `r v[i] w[i]` is true for all indices `i`. -/ @[inline] def isEqv (v w : Vector α n) (r : α → α → Bool) : Bool := - Array.isEqvAux v.toArray w.toArray (by simp) r 0 (by simp) + Array.isEqvAux v.toArray w.toArray (by simp) r n (by simp) instance [BEq α] : BEq (Vector α n) where beq a b := isEqv a b (· == ·) @@ -294,9 +295,7 @@ Finds the first index of a given value in a vector using `==` for comparison. Re no element of the index matches the given value. -/ @[inline] def indexOf? [BEq α] (v : Vector α n) (x : α) : Option (Fin n) := - match v.toArray.indexOf? x with - | some res => some (res.cast v.size_toArray) - | none => none + (v.toArray.indexOf? x).map (Fin.cast v.size_toArray) /-- Returns `true` when `v` is a prefix of the vector `w`. -/ @[inline] def isPrefixOf [BEq α] (v : Vector α m) (w : Vector α n) : Bool := diff --git a/Batteries/Data/Vector/Lemmas.lean b/Batteries/Data/Vector/Lemmas.lean index 6271f39f02..9a14c4a548 100644 --- a/Batteries/Data/Vector/Lemmas.lean +++ b/Batteries/Data/Vector/Lemmas.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2024 Shreyas Srinivas. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Shreyas Srinivas, Francois Dorais +Authors: Shreyas Srinivas, François G. Dorais -/ import Batteries.Data.Vector.Basic @@ -10,50 +10,13 @@ import Batteries.Data.List.Lemmas import Batteries.Data.Array.Lemmas import Batteries.Tactic.Lint.Simp -/-! -## Vectors -Lemmas about `Vector α n` --/ - namespace Batteries namespace Vector -theorem length_toList {α n} (xs : Vector α n) : xs.toList.length = n := by simp - -@[simp] theorem getElem_mk {data : Array α} {size : data.size = n} {i : Nat} (h : i < n) : - (Vector.mk data size)[i] = data[i] := rfl - -@[simp] theorem getElem_toArray {α n} (xs : Vector α n) (i : Nat) (h : i < xs.toArray.size) : - xs.toArray[i] = xs[i]'(by simpa using h) := by - cases xs - simp - -theorem getElem_toList {α n} (xs : Vector α n) (i : Nat) (h : i < xs.toList.length) : - xs.toList[i] = xs[i]'(by simpa using h) := by simp - -@[simp] theorem getElem_ofFn {α n} (f : Fin n → α) (i : Nat) (h : i < n) : - (Vector.ofFn f)[i] = f ⟨i, by simpa using h⟩ := by - simp [ofFn] - -/-- An `empty` vector maps to a `empty` vector. -/ -@[simp] -theorem map_empty (f : α → β) : map f empty = empty := by - rw [map, empty, mk.injEq] - exact Array.map_empty f - theorem toArray_injective : ∀ {v w : Vector α n}, v.toArray = w.toArray → v = w | {..}, {..}, rfl => rfl -/-- A vector of length `0` is an `empty` vector. -/ -protected theorem eq_empty (v : Vector α 0) : v = empty := by - apply Vector.toArray_injective - apply Array.eq_empty_of_size_eq_zero v.2 - -/-- -`Vector.ext` is an extensionality theorem. -Vectors `a` and `b` are equal to each other if their elements are equal for each valid index. --/ @[ext] protected theorem ext {a b : Vector α n} (h : (i : Nat) → (_ : i < n) → a[i] = b[i]) : a = b := by apply Vector.toArray_injective @@ -63,26 +26,217 @@ protected theorem ext {a b : Vector α n} (h : (i : Nat) → (_ : i < n) → a[i rw [a.size_toArray] at hi exact h i hi -@[simp] theorem push_mk {data : Array α} {size : data.size = n} {x : α} : - (Vector.mk data size).push x = - Vector.mk (data.push x) (by simp [size, Nat.succ_eq_add_one]) := rfl +/-! ### mk lemmas -/ + +theorem toArray_mk (a : Array α) (h : a.size = n) : (Vector.mk a h).toArray = a := rfl + +@[simp] theorem allDiff_mk [BEq α] (a : Array α) (h : a.size = n) : + (Vector.mk a h).allDiff = a.allDiff := rfl + +@[simp] theorem mk_append_mk (a b : Array α) (ha : a.size = n) (hb : b.size = m) : + Vector.mk a ha ++ Vector.mk b hb = Vector.mk (a ++ b) (by simp [ha, hb]) := rfl + +@[simp] theorem back!_mk [Inhabited α] (a : Array α) (h : a.size = n) : + (Vector.mk a h).back! = a.back! := rfl + +@[simp] theorem back?_mk (a : Array α) (h : a.size = n) : + (Vector.mk a h).back? = a.back? := rfl + +@[simp] theorem drop_mk (a : Array α) (h : a.size = n) (m) : + (Vector.mk a h).drop m = Vector.mk (a.extract m a.size) (by simp [h]) := rfl + +@[simp] theorem eraseIdx!_mk (a : Array α) (h : a.size = n) (i) (hi : i < n) : + (Vector.mk a h).eraseIdx! i = Vector.mk (a.eraseIdx! i) (by simp [h, hi]) := by + simp [Vector.eraseIdx!, hi] + +@[simp] theorem feraseIdx_mk (a : Array α) (h : a.size = n) (i) : + (Vector.mk a h).feraseIdx i = Vector.mk (a.feraseIdx (i.cast h.symm)) (by simp [h]) := rfl + +@[simp] theorem extract_mk (a : Array α) (h : a.size = n) (start stop) : + (Vector.mk a h).extract start stop = Vector.mk (a.extract start stop) (by simp [h]) := rfl + +@[simp] theorem getElem_mk (a : Array α) (h : a.size = n) (i) (hi : i < n) : + (Vector.mk a h)[i] = a[i] := rfl + +@[simp] theorem get_mk (a : Array α) (h : a.size = n) (i) : + (Vector.mk a h).get i = a.get (i.cast h.symm) := rfl + +@[simp] theorem getD_mk (a : Array α) (h : a.size = n) (i x) : + (Vector.mk a h).getD i x = a.getD i x := rfl + +@[simp] theorem uget_mk (a : Array α) (h : a.size = n) (i) (hi : i.toNat < n) : + (Vector.mk a h).uget i hi = a.uget i (by simp [h, hi]) := rfl + +@[simp] theorem indexOf?_mk [BEq α] (a : Array α) (h : a.size = n) (x : α) : + (Vector.mk a h).indexOf? x = (a.indexOf? x).map (Fin.cast h) := rfl + +@[simp] theorem mk_isEqv_mk (r : α → α → Bool) (a b : Array α) (ha : a.size = n) (hb : b.size = n) : + Vector.isEqv (Vector.mk a ha) (Vector.mk b hb) r = Array.isEqv a b r := by + simp [Vector.isEqv, Array.isEqv, ha, hb] + +@[simp] theorem mk_isPrefixOf_mk [BEq α] (a b : Array α) (ha : a.size = n) (hb : b.size = m) : + (Vector.mk a ha).isPrefixOf (Vector.mk b hb) = a.isPrefixOf b := rfl + +@[simp] theorem map_mk (a : Array α) (h : a.size = n) (f : α → β) : + (Vector.mk a h).map f = Vector.mk (a.map f) (by simp [h]) := rfl + +@[simp] theorem pop_mk (a : Array α) (h : a.size = n) : + (Vector.mk a h).pop = Vector.mk a.pop (by simp [h]) := rfl + +@[simp] theorem push_mk (a : Array α) (h : a.size = n) (x) : + (Vector.mk a h).push x = Vector.mk (a.push x) (by simp [h]) := rfl + +@[simp] theorem reverse_mk (a : Array α) (h : a.size = n) : + (Vector.mk a h).reverse = Vector.mk a.reverse (by simp [h]) := rfl + +@[simp] theorem set_mk (a : Array α) (h : a.size = n) (i x) : + (Vector.mk a h).set i x = Vector.mk (a.set (i.cast h.symm) x) (by simp [h]) := rfl + +@[simp] theorem set!_mk (a : Array α) (h : a.size = n) (i x) : + (Vector.mk a h).set! i x = Vector.mk (a.set! i x) (by simp [h]) := rfl + +@[simp] theorem setD_mk (a : Array α) (h : a.size = n) (i x) : + (Vector.mk a h).setD i x = Vector.mk (a.setD i x) (by simp [h]) := rfl + +@[simp] theorem setN_mk (a : Array α) (h : a.size = n) (i x) (hi : i < n) : + (Vector.mk a h).setN i x = Vector.mk (a.setN i x) (by simp [h]) := rfl + +@[simp] theorem swap_mk (a : Array α) (h : a.size = n) (i j) : + (Vector.mk a h).swap i j = Vector.mk (a.swap (i.cast h.symm) (j.cast h.symm)) (by simp [h]) := + rfl + +@[simp] theorem swap!_mk (a : Array α) (h : a.size = n) (i j) : + (Vector.mk a h).swap! i j = Vector.mk (a.swap! i j) (by simp [h]) := rfl + +@[simp] theorem swapN_mk (a : Array α) (h : a.size = n) (i j) (hi : i < n) (hj : j < n) : + (Vector.mk a h).swapN i j = Vector.mk (a.swapN i j) (by simp [h]) := rfl + +@[simp] theorem swapAt_mk (a : Array α) (h : a.size = n) (i x) : (Vector.mk a h).swapAt i x = + ((a.swapAt (i.cast h.symm) x).fst, Vector.mk (a.swapAt (i.cast h.symm) x).snd (by simp [h])) := + rfl -@[simp] theorem pop_mk {data : Array α} {size : data.size = n} : - (Vector.mk data size).pop = Vector.mk data.pop (by simp [size]) := rfl +@[simp] theorem swapAt!_mk (a : Array α) (h : a.size = n) (i x) : (Vector.mk a h).swapAt! i x = + ((a.swapAt! i x).fst, Vector.mk (a.swapAt! i x).snd (by simp [h])) := rfl + +@[simp] theorem swapAtN_mk (a : Array α) (h : a.size = n) (i x) (hi : i < n) : + (Vector.mk a h).swapAtN i x = + ((a.swapAtN i x).fst, Vector.mk (a.swapAtN i x).snd (by simp [h])) := rfl + +@[simp] theorem take_mk (a : Array α) (h : a.size = n) (m) : + (Vector.mk a h).take m = Vector.mk (a.take m) (by simp [h]) := rfl + +@[simp] theorem mk_zipWith_mk (f : α → β → γ) (a : Array α) (b : Array β) + (ha : a.size = n) (hb : b.size = n) : zipWith (Vector.mk a ha) (Vector.mk b hb) f = + Vector.mk (Array.zipWith a b f) (by simp [ha, hb]) := rfl + +@[simp] theorem getElem_toArray {α n} (xs : Vector α n) (i : Nat) (h : i < xs.toArray.size) : + xs.toArray[i] = xs[i]'(by simpa using h) := by + cases xs; simp + +/-! ### toArray lemmas -/ + +@[simp] theorem toArray_append (a : Vector α m) (b : Vector α n) : + (a ++ b).toArray = a.toArray ++ b.toArray := rfl + +@[simp] theorem toArray_drop (a : Vector α n) (m) : + (a.drop m).toArray = a.toArray.extract m a.size := rfl + +@[simp] theorem toArray_empty : (Vector.empty (α := α)).toArray = #[] := rfl + +@[simp] theorem toArray_mkEmpty (cap) : + (Vector.mkEmpty (α := α) cap).toArray = Array.mkEmpty cap := rfl + +@[simp] theorem toArray_eraseIdx! (a : Vector α n) (i) (hi : i < n) : + (a.eraseIdx! i).toArray = a.toArray.eraseIdx! i := by + cases a; simp [hi] + +@[simp] theorem toArray_eraseIdxN (a : Vector α n) (i) (hi : i < n) : + (a.eraseIdxN i).toArray = a.toArray.eraseIdxN i (by simp [hi]) := rfl + +@[simp] theorem toArray_feraseIdx (a : Vector α n) (i) : + (a.feraseIdx i).toArray = a.toArray.feraseIdx (i.cast a.size_toArray.symm) := rfl + +@[simp] theorem toArray_extract (a : Vector α n) (start stop) : + (a.extract start stop).toArray = a.toArray.extract start stop := rfl + +@[simp] theorem toArray_map (f : α → β) (a : Vector α n) : + (a.map f).toArray = a.toArray.map f := rfl + +@[simp] theorem toArray_ofFn (f : Fin n → α) : (Vector.ofFn f).toArray = Array.ofFn f := rfl + +@[simp] theorem toArray_pop (a : Vector α n) : a.pop.toArray = a.toArray.pop := rfl + +@[simp] theorem toArray_push (a : Vector α n) (x) : (a.push x).toArray = a.toArray.push x := rfl + +@[simp] theorem toArray_range : (Vector.range n).toArray = Array.range n := rfl + +@[simp] theorem toArray_reverse (a : Vector α n) : a.reverse.toArray = a.toArray.reverse := rfl + +@[simp] theorem toArray_set (a : Vector α n) (i x) : + (a.set i x).toArray = a.toArray.set (i.cast a.size_toArray.symm) x := rfl + +@[simp] theorem toArray_set! (a : Vector α n) (i x) : + (a.set! i x).toArray = a.toArray.set! i x := rfl + +@[simp] theorem toArray_setD (a : Vector α n) (i x) : + (a.setD i x).toArray = a.toArray.setD i x := rfl + +@[simp] theorem toArray_setN (a : Vector α n) (i x) (hi : i < n) : + (a.setN i x).toArray = a.toArray.setN i x (by simp [hi]) := rfl + +@[simp] theorem toArray_singleton (x : α) : (Vector.singleton x).toArray = #[x] := rfl + +@[simp] theorem toArray_swap (a : Vector α n) (i j) : (a.swap i j).toArray = + a.toArray.swap (i.cast a.size_toArray.symm) (j.cast a.size_toArray.symm) := rfl + +@[simp] theorem toArray_swap! (a : Vector α n) (i j) : + (a.swap! i j).toArray = a.toArray.swap! i j := rfl + +@[simp] theorem toArray_swapN (a : Vector α n) (i j) (hi : i < n) (hj : j < n) : + (a.swapN i j).toArray = a.toArray.swapN i j (by simp [hi]) (by simp [hj]) := rfl + +@[simp] theorem toArray_swapAt (a : Vector α n) (i x) : + ((a.swapAt i x).fst, (a.swapAt i x).snd.toArray) = + ((a.toArray.swapAt (i.cast a.size_toArray.symm) x).fst, + (a.toArray.swapAt (i.cast a.size_toArray.symm) x).snd) := rfl + +@[simp] theorem toArray_swapAt! (a : Vector α n) (i x) : + ((a.swapAt! i x).fst, (a.swapAt! i x).snd.toArray) = + ((a.toArray.swapAt! i x).fst, (a.toArray.swapAt! i x).snd) := rfl + +@[simp] theorem toArray_swapAtN (a : Vector α n) (i x) (hi : i < n) : + ((a.swapAtN i x).fst, (a.swapAtN i x).snd.toArray) = + ((a.toArray.swapAtN i x (by simp [hi])).fst, + (a.toArray.swapAtN i x (by simp [hi])).snd) := rfl + +@[simp] theorem toArray_take (a : Vector α n) (m) : (a.take m).toArray = a.toArray.take m := rfl + +@[simp] theorem toArray_zipWith (f : α → β → γ) (a : Vector α n) (b : Vector β n) : + (Vector.zipWith a b f).toArray = Array.zipWith a.toArray b.toArray f := rfl + +/-! ### toList lemmas -/ + +theorem length_toList {α n} (xs : Vector α n) : xs.toList.length = n := by simp + +theorem getElem_toList {α n} (xs : Vector α n) (i : Nat) (h : i < xs.toList.length) : + xs.toList[i] = xs[i]'(by simpa using h) := by simp + +/-! ### getElem lemmas -/ + +@[simp] theorem getElem_ofFn {α n} (f : Fin n → α) (i : Nat) (h : i < n) : + (Vector.ofFn f)[i] = f ⟨i, by simpa using h⟩ := by + simp [ofFn] @[simp] theorem getElem_push_last {v : Vector α n} {x : α} : (v.push x)[n] = x := by - rcases v with ⟨data, rfl⟩ - simp + rcases v with ⟨_, rfl⟩; simp -- The `simpNF` linter incorrectly claims that this lemma can not be applied by `simp`. @[simp, nolint simpNF] theorem getElem_push_lt {v : Vector α n} {x : α} {i : Nat} (h : i < n) : (v.push x)[i] = v[i] := by - rcases v with ⟨data, rfl⟩ - simp [Array.getElem_push_lt, h] + rcases v with ⟨_, rfl⟩; simp [Array.getElem_push_lt, h] @[simp] theorem getElem_pop {v : Vector α n} {i : Nat} (h : i < n - 1) : (v.pop)[i] = v[i] := by - rcases v with ⟨data, rfl⟩ - simp + rcases v with ⟨_, rfl⟩; simp /-- Variant of `getElem_pop` that will sometimes fire when `getElem_pop` gets stuck because of @@ -100,6 +254,15 @@ defeq issues in the implicit size argument. subst h simp [pop, back, back!, ← Array.eq_push_pop_back!_of_size_ne_zero] +/-! ### empty lemmas -/ + +@[simp] theorem map_empty (f : α → β) : map f empty = empty := by + apply toArray_injective; simp + +protected theorem eq_empty (v : Vector α 0) : v = empty := by + apply toArray_injective + apply Array.eq_empty_of_size_eq_zero v.2 + /-! ### Decidable quantifiers. -/ theorem forall_zero_iff {P : Vector α 0 → Prop} : From 7488499a8aad6ffada87ab6db73673d88dc04c97 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Mon, 25 Nov 2024 20:45:41 -0500 Subject: [PATCH 172/177] feat: add KMP frontend for lists (#1065) --- Batteries/Data/List.lean | 1 + Batteries/Data/List/Matcher.lean | 85 ++++++++++++++++++++++++++++++++ BatteriesTest/kmp_matcher.lean | 17 ++++++- 3 files changed, 102 insertions(+), 1 deletion(-) create mode 100644 Batteries/Data/List/Matcher.lean diff --git a/Batteries/Data/List.lean b/Batteries/Data/List.lean index 3429039dc9..248240370a 100644 --- a/Batteries/Data/List.lean +++ b/Batteries/Data/List.lean @@ -3,6 +3,7 @@ import Batteries.Data.List.Count import Batteries.Data.List.EraseIdx import Batteries.Data.List.Init.Lemmas import Batteries.Data.List.Lemmas +import Batteries.Data.List.Matcher import Batteries.Data.List.Monadic import Batteries.Data.List.OfFn import Batteries.Data.List.Pairwise diff --git a/Batteries/Data/List/Matcher.lean b/Batteries/Data/List/Matcher.lean new file mode 100644 index 0000000000..cc3b6db42f --- /dev/null +++ b/Batteries/Data/List/Matcher.lean @@ -0,0 +1,85 @@ +/- +Copyright (c) 2024 François G. Dorais. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: François G. Dorais +-/ +import Batteries.Data.Array.Match +import Batteries.Data.String.Basic + +namespace List + +/-- Knuth-Morris-Pratt matcher type + +This type is used to keep data for running the Knuth-Morris-Pratt (KMP) list matching algorithm. +KMP is a linear time algorithm to locate all contiguous sublists of a list that match a given +pattern. Generating the algorithm data is also linear in the length of the pattern but the data can +be re-used to match the same pattern over multiple lists. + +The KMP data for a pattern can be generated using `Matcher.ofList`. Then `Matcher.find?` and +`Matcher.findAll` can be used to run the algorithm on an input list. +``` +def m := Matcher.ofList [0,1,1,0] + +#eval Option.isSome <| m.find? [2,1,1,0,1,1,2] -- false +#eval Option.isSome <| m.find? [0,0,1,1,0,0] -- true + +#eval Array.size <| m.findAll [0,1,1,0,1,1,0] -- 2 +#eval Array.size <| m.findAll [0,1,1,0,1,1,0,1,1,0] -- 3 +``` +-/ +structure Matcher (α : Type _) extends Array.Matcher α where + /-- The pattern for the matcher -/ + pattern : List α + +/-- Make KMP matcher from list pattern. -/ +@[inline] def Matcher.ofList [BEq α] (pattern : List α) : Matcher α where + toMatcher := Array.Matcher.ofStream pattern + pattern := pattern + +/-- List stream that keeps count of items read. -/ +local instance (α) : Stream (List α × Nat) α where + next? + | ([], _) => none + | (x::xs, n) => (x, xs, n+1) + +/-- +Find all start and end positions of all infix sublists of `l` matching `m.pattern`. +The sublists may be overlapping. +-/ +partial def Matcher.findAll [BEq α] (m : Matcher α) (l : List α) : Array (Nat × Nat) := + loop (l, 0) m.toMatcher #[] +where + /-- Accumulator loop for `List.Matcher.findAll` -/ + loop (l : List α × Nat) (am : Array.Matcher α) (occs : Array (Nat × Nat)) : Array (Nat × Nat) := + match am.next? l with + | none => occs + | some (l, am) => loop l am (occs.push (l.snd - m.table.size, l.snd)) + +/-- +Find the start and end positions of the first infix sublist of `l` matching `m.pattern`, +or `none` if there is no such sublist. +-/ +def Matcher.find? [BEq α] (m : Matcher α) (l : List α) : Option (Nat × Nat) := + match m.next? (l, 0) with + | none => none + | some (l, _) => some (l.snd - m.table.size, l.snd) + +/-- +Returns all the start and end positions of all infix sublists of of `l` that match `pattern`. +The sublists may be overlapping. +-/ +@[inline] def findAllInfix [BEq α] (l pattern : List α) : Array (Nat × Nat) := + (Matcher.ofList pattern).findAll l + +/-- +Returns the start and end positions of the first infix sublist of `l` that matches `pattern`, +or `none` if there is no such sublist. +-/ +@[inline] def findInfix? [BEq α] (l pattern : List α) : Option (Nat × Nat) := + (Matcher.ofList pattern).find? l + +/-- +Returns true iff `pattern` occurs as an infix sublist of `l`. +-/ +@[inline] def containsInfix [BEq α] (l pattern : List α) : Bool := + findInfix? l pattern |>.isSome diff --git a/BatteriesTest/kmp_matcher.lean b/BatteriesTest/kmp_matcher.lean index f61e42343f..c15a47befd 100644 --- a/BatteriesTest/kmp_matcher.lean +++ b/BatteriesTest/kmp_matcher.lean @@ -1,6 +1,9 @@ import Batteries.Data.String.Matcher +import Batteries.Data.List.Matcher -/-! Tests for Knuth-Morris-Pratt matching algorithm -/ +/-! # Tests for the Knuth-Morris-Pratt (KMP) matching algorithm -/ + +/-! ### String API -/ /-- Matcher for pattern "abba" -/ def m := String.Matcher.ofString "abba" @@ -24,3 +27,15 @@ def m := String.Matcher.ofString "abba" #guard Array.size ("xyxyyxxyx".findAllSubstr "xyx") = 2 #guard Array.size ("xyxyxyyxxyxyx".findAllSubstr "xyx") = 4 + +/-! ### List API -/ + +def lm := List.Matcher.ofList [0,1,1,0] + +#guard lm.find? [2,1,1,0,1,1,2] == none + +#guard lm.find? [0,0,1,1,0,0] == some (1, 5) + +#guard (lm.findAll [0,1,1,0,1,1,0]).size == 2 + +#guard (lm.findAll [0,1,1,0,1,1,0,1,1,0]).size == 3 From f6d16c2a073e0385a362ad3e5d9e7e8260e298d6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Tue, 26 Nov 2024 01:07:46 -0500 Subject: [PATCH 173/177] refactor: add List.finRange and Array.finRange (#1055) --- Batteries/Data/Array/Basic.lean | 3 ++ Batteries/Data/Array/Lemmas.lean | 18 +++++++++ Batteries/Data/Fin.lean | 1 + Batteries/Data/Fin/Basic.lean | 11 ++++-- Batteries/Data/Fin/Fold.lean | 47 ++++++++++++++++++++++ Batteries/Data/Fin/Lemmas.lean | 65 ------------------------------- Batteries/Data/List.lean | 1 + Batteries/Data/List/Basic.lean | 3 ++ Batteries/Data/List/FinRange.lean | 60 ++++++++++++++++++++++++++++ 9 files changed, 140 insertions(+), 69 deletions(-) create mode 100644 Batteries/Data/Fin/Fold.lean create mode 100644 Batteries/Data/List/FinRange.lean diff --git a/Batteries/Data/Array/Basic.lean b/Batteries/Data/Array/Basic.lean index 403c42d9e2..3f05693926 100644 --- a/Batteries/Data/Array/Basic.lean +++ b/Batteries/Data/Array/Basic.lean @@ -175,6 +175,9 @@ def eraseIdx! (a : Array α) (i : Nat) : Array α := have : Inhabited (Array α) := ⟨a⟩ panic! s!"index {i} out of bounds" +/-- `finRange n` is the array of all elements of `Fin n` in order. -/ +protected def finRange (n : Nat) : Array (Fin n) := ofFn fun i => i + end Array diff --git a/Batteries/Data/Array/Lemmas.lean b/Batteries/Data/Array/Lemmas.lean index 8a5544089d..7022ef1cc2 100644 --- a/Batteries/Data/Array/Lemmas.lean +++ b/Batteries/Data/Array/Lemmas.lean @@ -5,6 +5,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Gabriel Ebner -/ import Batteries.Data.List.Lemmas +import Batteries.Data.List.FinRange import Batteries.Data.Array.Basic import Batteries.Tactic.SeqFocus import Batteries.Util.ProofWanted @@ -240,3 +241,20 @@ theorem getElem_insertAt_eq (as : Array α) (i : Fin (as.size+1)) (v : α) rw [get_insertAt_loop_eq, Fin.getElem_fin, getElem_push_eq] exact heq exact Nat.le_of_lt_succ i.is_lt + +/-! ### ofFn -/ + +@[simp] theorem toList_ofFn (f : Fin n → α) : (ofFn f).toList = List.ofFn f := by + apply List.ext_getElem <;> simp + +/-! ### finRange -/ + +@[simp] theorem size_finRange (n) : (Array.finRange n).size = n := by + simp [Array.finRange] + +@[simp] theorem getElem_finRange (n i) (h : i < (Array.finRange n).size) : + (Array.finRange n)[i] = ⟨i, size_finRange n ▸ h⟩ := by + simp [Array.finRange] + +@[simp] theorem toList_finRange (n) : (Array.finRange n).toList = List.finRange n := by + simp [Array.finRange, List.finRange] diff --git a/Batteries/Data/Fin.lean b/Batteries/Data/Fin.lean index 5fe5cc41ca..7a5b9c16e8 100644 --- a/Batteries/Data/Fin.lean +++ b/Batteries/Data/Fin.lean @@ -1,2 +1,3 @@ import Batteries.Data.Fin.Basic +import Batteries.Data.Fin.Fold import Batteries.Data.Fin.Lemmas diff --git a/Batteries/Data/Fin/Basic.lean b/Batteries/Data/Fin/Basic.lean index b61481e33a..f1357f50b0 100644 --- a/Batteries/Data/Fin/Basic.lean +++ b/Batteries/Data/Fin/Basic.lean @@ -3,17 +3,20 @@ Copyright (c) 2017 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Robert Y. Lewis, Keeley Hoek, Mario Carneiro -/ +import Batteries.Tactic.Alias +import Batteries.Data.Array.Basic +import Batteries.Data.List.Basic namespace Fin /-- `min n m` as an element of `Fin (m + 1)` -/ def clamp (n m : Nat) : Fin (m + 1) := ⟨min n m, Nat.lt_succ_of_le (Nat.min_le_right ..)⟩ -/-- `enum n` is the array of all elements of `Fin n` in order -/ -def enum (n) : Array (Fin n) := Array.ofFn id +@[deprecated (since := "2024-11-15")] +alias enum := Array.finRange -/-- `list n` is the list of all elements of `Fin n` in order -/ -def list (n) : List (Fin n) := (enum n).toList +@[deprecated (since := "2024-11-15")] +alias list := List.finRange /-- Folds a monadic function over `Fin n` from left to right: diff --git a/Batteries/Data/Fin/Fold.lean b/Batteries/Data/Fin/Fold.lean new file mode 100644 index 0000000000..ed37c6b948 --- /dev/null +++ b/Batteries/Data/Fin/Fold.lean @@ -0,0 +1,47 @@ +/- +Copyright (c) 2024 François G. Dorais. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: François G. Dorais +-/ +import Batteries.Data.List.FinRange + +namespace Fin + +theorem foldlM_eq_foldlM_finRange [Monad m] (f : α → Fin n → m α) (x) : + foldlM n f x = (List.finRange n).foldlM f x := by + induction n generalizing x with + | zero => simp + | succ n ih => + rw [foldlM_succ, List.finRange_succ, List.foldlM_cons] + congr; funext + rw [List.foldlM_map, ih] + +@[deprecated (since := "2024-11-19")] +alias foldlM_eq_foldlM_list := foldlM_eq_foldlM_finRange + +theorem foldrM_eq_foldrM_finRange [Monad m] [LawfulMonad m] (f : Fin n → α → m α) (x) : + foldrM n f x = (List.finRange n).foldrM f x := by + induction n with + | zero => simp + | succ n ih => rw [foldrM_succ, ih, List.finRange_succ, List.foldrM_cons, List.foldrM_map] + +@[deprecated (since := "2024-11-19")] +alias foldrM_eq_foldrM_list := foldrM_eq_foldrM_finRange + +theorem foldl_eq_foldl_finRange (f : α → Fin n → α) (x) : + foldl n f x = (List.finRange n).foldl f x := by + induction n generalizing x with + | zero => rw [foldl_zero, List.finRange_zero, List.foldl_nil] + | succ n ih => rw [foldl_succ, ih, List.finRange_succ, List.foldl_cons, List.foldl_map] + +@[deprecated (since := "2024-11-19")] +alias foldl_eq_foldl_list := foldl_eq_foldl_finRange + +theorem foldr_eq_foldr_finRange (f : Fin n → α → α) (x) : + foldr n f x = (List.finRange n).foldr f x := by + induction n with + | zero => rw [foldr_zero, List.finRange_zero, List.foldr_nil] + | succ n ih => rw [foldr_succ, ih, List.finRange_succ, List.foldr_cons, List.foldr_map] + +@[deprecated (since := "2024-11-19")] +alias foldr_eq_foldr_list := foldr_eq_foldr_finRange diff --git a/Batteries/Data/Fin/Lemmas.lean b/Batteries/Data/Fin/Lemmas.lean index 5010e1310f..90a574cf2c 100644 --- a/Batteries/Data/Fin/Lemmas.lean +++ b/Batteries/Data/Fin/Lemmas.lean @@ -14,46 +14,6 @@ attribute [norm_cast] val_last @[simp] theorem coe_clamp (n m : Nat) : (clamp n m : Nat) = min n m := rfl -/-! ### enum/list -/ - -@[simp] theorem size_enum (n) : (enum n).size = n := Array.size_ofFn .. - -@[simp] theorem enum_zero : (enum 0) = #[] := by simp [enum, Array.ofFn, Array.ofFn.go] - -@[simp] theorem getElem_enum (i) (h : i < (enum n).size) : (enum n)[i] = ⟨i, size_enum n ▸ h⟩ := - Array.getElem_ofFn .. - -@[simp] theorem length_list (n) : (list n).length = n := by simp [list] - -@[simp] theorem getElem_list (i : Nat) (h : i < (list n).length) : - (list n)[i] = cast (length_list n) ⟨i, h⟩ := by - simp only [list]; rw [← Array.getElem_eq_getElem_toList, getElem_enum, cast_mk] - -@[deprecated getElem_list (since := "2024-06-12")] -theorem get_list (i : Fin (list n).length) : (list n).get i = i.cast (length_list n) := by - simp [getElem_list] - -@[simp] theorem list_zero : list 0 = [] := by simp [list] - -theorem list_succ (n) : list (n+1) = 0 :: (list n).map Fin.succ := by - apply List.ext_get; simp; intro i; cases i <;> simp - -theorem list_succ_last (n) : list (n+1) = (list n).map castSucc ++ [last n] := by - rw [list_succ] - induction n with - | zero => simp - | succ n ih => - rw [list_succ, List.map_cons castSucc, ih] - simp [Function.comp_def, succ_castSucc] - -theorem list_reverse (n) : (list n).reverse = (list n).map rev := by - induction n with - | zero => simp - | succ n ih => - conv => lhs; rw [list_succ_last] - conv => rhs; rw [list_succ] - simp [← List.map_reverse, ih, Function.comp_def, rev_succ] - /-! ### foldlM -/ theorem foldlM_loop_lt [Monad m] (f : α → Fin n → m α) (x) (h : i < n) : @@ -82,15 +42,6 @@ termination_by n - i theorem foldlM_succ [Monad m] (f : α → Fin (n+1) → m α) (x) : foldlM (n+1) f x = f x 0 >>= foldlM n (fun x j => f x j.succ) := foldlM_loop .. -theorem foldlM_eq_foldlM_list [Monad m] (f : α → Fin n → m α) (x) : - foldlM n f x = (list n).foldlM f x := by - induction n generalizing x with - | zero => simp - | succ n ih => - rw [foldlM_succ, list_succ, List.foldlM_cons] - congr; funext - rw [List.foldlM_map, ih] - /-! ### foldrM -/ theorem foldrM_loop_zero [Monad m] (f : Fin n → α → m α) (x) : @@ -119,12 +70,6 @@ theorem foldrM_loop [Monad m] [LawfulMonad m] (f : Fin (n+1) → α → m α) (x theorem foldrM_succ [Monad m] [LawfulMonad m] (f : Fin (n+1) → α → m α) (x) : foldrM (n+1) f x = foldrM n (fun i => f i.succ) x >>= f 0 := foldrM_loop .. -theorem foldrM_eq_foldrM_list [Monad m] [LawfulMonad m] (f : Fin n → α → m α) (x) : - foldrM n f x = (list n).foldrM f x := by - induction n with - | zero => simp - | succ n ih => rw [foldrM_succ, ih, list_succ, List.foldrM_cons, List.foldrM_map] - /-! ### foldl -/ theorem foldl_loop_lt (f : α → Fin n → α) (x) (h : i < n) : @@ -162,11 +107,6 @@ theorem foldl_eq_foldlM (f : α → Fin n → α) (x) : foldl n f x = foldlM (m:=Id) n f x := by induction n generalizing x <;> simp [foldl_succ, foldlM_succ, *] -theorem foldl_eq_foldl_list (f : α → Fin n → α) (x) : foldl n f x = (list n).foldl f x := by - induction n generalizing x with - | zero => rw [foldl_zero, list_zero, List.foldl_nil] - | succ n ih => rw [foldl_succ, ih, list_succ, List.foldl_cons, List.foldl_map] - /-! ### foldr -/ theorem foldr_loop_zero (f : Fin n → α → α) (x) : @@ -198,11 +138,6 @@ theorem foldr_eq_foldrM (f : Fin n → α → α) (x) : foldr n f x = foldrM (m:=Id) n f x := by induction n <;> simp [foldr_succ, foldrM_succ, *] -theorem foldr_eq_foldr_list (f : Fin n → α → α) (x) : foldr n f x = (list n).foldr f x := by - induction n with - | zero => rw [foldr_zero, list_zero, List.foldr_nil] - | succ n ih => rw [foldr_succ, ih, list_succ, List.foldr_cons, List.foldr_map] - theorem foldl_rev (f : Fin n → α → α) (x) : foldl n (fun x i => f i.rev x) x = foldr n f x := by induction n generalizing x with diff --git a/Batteries/Data/List.lean b/Batteries/Data/List.lean index 248240370a..b4081d125e 100644 --- a/Batteries/Data/List.lean +++ b/Batteries/Data/List.lean @@ -1,6 +1,7 @@ import Batteries.Data.List.Basic import Batteries.Data.List.Count import Batteries.Data.List.EraseIdx +import Batteries.Data.List.FinRange import Batteries.Data.List.Init.Lemmas import Batteries.Data.List.Lemmas import Batteries.Data.List.Matcher diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index 9c0ab14d46..7f3e66e1a6 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -1064,3 +1064,6 @@ where loop : List α → List α → List α | [], r => reverseAux (a :: r) [] -- Note: `reverseAux` is tail recursive. | b :: l, r => bif p b then reverseAux (a :: r) (b :: l) else loop l (b :: r) + +/-- `finRange n` lists all elements of `Fin n` in order -/ +def finRange (n : Nat) : List (Fin n) := ofFn fun i => i diff --git a/Batteries/Data/List/FinRange.lean b/Batteries/Data/List/FinRange.lean new file mode 100644 index 0000000000..432e3133af --- /dev/null +++ b/Batteries/Data/List/FinRange.lean @@ -0,0 +1,60 @@ +/- +Copyright (c) 2024 François G. Dorais. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: François G. Dorais +-/ +import Batteries.Data.List.OfFn + +namespace List + +@[simp] theorem length_finRange (n) : (List.finRange n).length = n := by + simp [List.finRange] + +@[deprecated (since := "2024-11-19")] +alias length_list := length_finRange + +@[simp] theorem getElem_finRange (i : Nat) (h : i < (List.finRange n).length) : + (finRange n)[i] = Fin.cast (length_finRange n) ⟨i, h⟩ := by + simp [List.finRange] + +@[deprecated (since := "2024-11-19")] +alias getElem_list := getElem_finRange + +@[simp] theorem finRange_zero : finRange 0 = [] := by simp [finRange, ofFn] + +@[deprecated (since := "2024-11-19")] +alias list_zero := finRange_zero + +theorem finRange_succ (n) : finRange (n+1) = 0 :: (finRange n).map Fin.succ := by + apply List.ext_getElem; simp; intro i; cases i <;> simp + +@[deprecated (since := "2024-11-19")] +alias list_succ := finRange_succ + +theorem finRange_succ_last (n) : + finRange (n+1) = (finRange n).map Fin.castSucc ++ [Fin.last n] := by + apply List.ext_getElem + · simp + · intros + simp only [List.finRange, List.getElem_ofFn, getElem_append, length_map, length_ofFn, + getElem_map, Fin.castSucc_mk, getElem_singleton] + split + · rfl + · next h => exact Fin.eq_last_of_not_lt h + +@[deprecated (since := "2024-11-19")] +alias list_succ_last := finRange_succ_last + +theorem finRange_reverse (n) : (finRange n).reverse = (finRange n).map Fin.rev := by + induction n with + | zero => simp + | succ n ih => + conv => lhs; rw [finRange_succ_last] + conv => rhs; rw [finRange_succ] + rw [reverse_append, reverse_cons, reverse_nil, nil_append, singleton_append, ← map_reverse, + map_cons, ih, map_map, map_map] + congr; funext + simp [Fin.rev_succ] + +@[deprecated (since := "2024-11-19")] +alias list_reverse := finRange_reverse From c933dd9b00271d869e22b802a015092d1e8e454a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 27 Nov 2024 22:33:25 +1100 Subject: [PATCH 174/177] feat: List.dropPrefix? / dropSuffix? / dropInfix? and specification lemmas (#1066) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: François G. Dorais --- Batteries/Data/List/Basic.lean | 32 +++++++ Batteries/Data/List/Lemmas.lean | 145 ++++++++++++++++++++++++++++++++ 2 files changed, 177 insertions(+) diff --git a/Batteries/Data/List/Basic.lean b/Batteries/Data/List/Basic.lean index 7f3e66e1a6..8dbcabb4a8 100644 --- a/Batteries/Data/List/Basic.lean +++ b/Batteries/Data/List/Basic.lean @@ -1065,5 +1065,37 @@ where | [], r => reverseAux (a :: r) [] -- Note: `reverseAux` is tail recursive. | b :: l, r => bif p b then reverseAux (a :: r) (b :: l) else loop l (b :: r) +/-- `dropPrefix? l p` returns +`some r` if `l = p' ++ r` for some `p'` which is paiwise `==` to `p`, +and `none` otherwise. -/ +def dropPrefix? [BEq α] : List α → List α → Option (List α) + | list, [] => some list + | [], _ :: _ => none + | a :: as, b :: bs => if a == b then dropPrefix? as bs else none + +/-- `dropSuffix? l s` returns +`some r` if `l = r ++ s'` for some `s'` which is paiwise `==` to `s`, +and `none` otherwise. -/ +def dropSuffix? [BEq α] (l s : List α) : Option (List α) := + let (r, s') := l.splitAt (l.length - s.length) + if s' == s then some r else none + +/-- `dropInfix? l i` returns +`some (p, s)` if `l = p ++ i' ++ s` for some `i'` which is paiwise `==` to `i`, +and `none` otherwise. + +Note that this is an inefficient implementation, and if computation time is a concern you should be +using the Knuth-Morris-Pratt algorithm as implemented in `Batteries.Data.List.Matcher`. +-/ +def dropInfix? [BEq α] (l i : List α) : Option (List α × List α) := + go l [] +where + /-- Inner loop for `dropInfix?`. -/ + go : List α → List α → Option (List α × List α) + | [], acc => if i.isEmpty then some (acc.reverse, []) else none + | a :: as, acc => match (a :: as).dropPrefix? i with + | none => go as (a :: acc) + | some s => (acc.reverse, s) + /-- `finRange n` lists all elements of `Fin n` in order -/ def finRange (n : Nat) : List (Fin n) := ofFn fun i => i diff --git a/Batteries/Data/List/Lemmas.lean b/Batteries/Data/List/Lemmas.lean index c8fba52299..840940dbaf 100644 --- a/Batteries/Data/List/Lemmas.lean +++ b/Batteries/Data/List/Lemmas.lean @@ -15,6 +15,26 @@ namespace List @[simp] theorem getElem_mk {xs : List α} {i : Nat} (h : i < xs.length) : (Array.mk xs)[i] = xs[i] := rfl +/-! ### == -/ + +@[simp] theorem beq_nil_iff [BEq α] {l : List α} : (l == []) = l.isEmpty := by + cases l <;> rfl + +@[simp] theorem nil_beq_iff [BEq α] {l : List α} : ([] == l) = l.isEmpty := by + cases l <;> rfl + +@[simp] theorem cons_beq_cons [BEq α] {a b : α} {l₁ l₂ : List α} : + (a :: l₁ == b :: l₂) = (a == b && l₁ == l₂) := rfl + +theorem length_eq_of_beq [BEq α] {l₁ l₂ : List α} (h : l₁ == l₂) : l₁.length = l₂.length := + match l₁, l₂ with + | [], [] => rfl + | [], _ :: _ => by simp [beq_nil_iff] at h + | _ :: _, [] => by simp [nil_beq_iff] at h + | a :: l₁, b :: l₂ => by + simp at h + simpa using length_eq_of_beq h.2 + /-! ### next? -/ @[simp] theorem next?_nil : @next? α [] = none := rfl @@ -508,6 +528,131 @@ theorem insertP_loop (a : α) (l r : List α) : induction l with simp [insertP, insertP.loop, cond] | cons _ _ ih => split <;> simp [insertP_loop, ih] +/-! ### dropPrefix?, dropSuffix?, dropInfix?-/ + +open Option + +@[simp] theorem dropPrefix?_nil [BEq α] {p : List α} : dropPrefix? p [] = some p := by + simp [dropPrefix?] + +theorem dropPrefix?_eq_some_iff [BEq α] {l p s : List α} : + dropPrefix? l p = some s ↔ ∃ p', l = p' ++ s ∧ p' == p := by + unfold dropPrefix? + split + · simp + · simp + · rename_i a as b bs + simp only [ite_none_right_eq_some] + constructor + · rw [dropPrefix?_eq_some_iff] + rintro ⟨w, p', rfl, h⟩ + refine ⟨a :: p', by simp_all⟩ + · rw [dropPrefix?_eq_some_iff] + rintro ⟨p, h, w⟩ + rw [cons_eq_append_iff] at h + obtain (⟨rfl, rfl⟩ | ⟨a', rfl, rfl⟩) := h + · simp at w + · simp only [cons_beq_cons, Bool.and_eq_true] at w + refine ⟨w.1, a', rfl, w.2⟩ + +theorem dropPrefix?_append_of_beq [BEq α] {l₁ l₂ : List α} (p : List α) (h : l₁ == l₂) : + dropPrefix? (l₁ ++ p) l₂ = some p := by + simp [dropPrefix?_eq_some_iff, h] + +theorem dropSuffix?_eq_some_iff [BEq α] {l p s : List α} : + dropSuffix? l s = some p ↔ ∃ s', l = p ++ s' ∧ s' == s := by + unfold dropSuffix? + rw [splitAt_eq] + simp only [ite_none_right_eq_some, some.injEq] + constructor + · rintro ⟨w, rfl⟩ + refine ⟨_, by simp, w⟩ + · rintro ⟨s', rfl, w⟩ + simp [length_eq_of_beq w, w] + +@[simp] theorem dropSuffix?_nil [BEq α] {s : List α} : dropSuffix? s [] = some s := by + simp [dropSuffix?_eq_some_iff] + +theorem dropInfix?_go_eq_some_iff [BEq α] {i l acc p s : List α} : + dropInfix?.go i l acc = some (p, s) ↔ ∃ p', + p = acc.reverse ++ p' ∧ + -- `i` is an infix up to `==` + (∃ i', l = p' ++ i' ++ s ∧ i' == i) ∧ + -- and there is no shorter prefix for which that is the case + (∀ p'' i'' s'', l = p'' ++ i'' ++ s'' → i'' == i → p''.length ≥ p'.length) := by + unfold dropInfix?.go + split + · simp only [isEmpty_eq_true, ite_none_right_eq_some, some.injEq, Prod.mk.injEq, nil_eq, + append_assoc, append_eq_nil, ge_iff_le, and_imp] + constructor + · rintro ⟨rfl, rfl, rfl⟩ + simp + · rintro ⟨p', rfl, ⟨_, ⟨rfl, rfl, rfl⟩, h⟩, w⟩ + simp_all + · rename_i a t + split <;> rename_i h + · rw [dropInfix?_go_eq_some_iff] + constructor + · rintro ⟨p', rfl, ⟨i', rfl, h₂⟩, w⟩ + refine ⟨a :: p', ?_⟩ + simp [h₂] + intro p'' i'' s'' h₁ h₂ + rw [cons_eq_append_iff] at h₁ + obtain (⟨rfl, h₁⟩ | ⟨p'', rfl, h₁⟩) := h₁ + · rw [append_assoc, ← h₁] at h + have := dropPrefix?_append_of_beq s'' h₂ + simp_all + · simpa using w p'' i'' s'' (by simpa using h₁) h₂ + · rintro ⟨p', rfl, ⟨i', h₁, h₂⟩, w⟩ + rw [cons_eq_append_iff] at h₁ + simp at h₁ + obtain (⟨⟨rfl, rfl⟩, rfl⟩ | ⟨a', h₁, rfl⟩) := h₁ + · simp only [nil_beq_iff, isEmpty_eq_true] at h₂ + simp only [h₂] at h + simp at h + · rw [append_eq_cons_iff] at h₁ + obtain (⟨rfl, rfl⟩ | ⟨p', rfl, rfl⟩) := h₁ + · rw [← cons_append] at h + have := dropPrefix?_append_of_beq s h₂ + simp_all + · refine ⟨p', ?_⟩ + simp only [reverse_cons, append_assoc, singleton_append, append_cancel_left_eq, + append_cancel_right_eq, exists_eq_left', ge_iff_le, true_and] + refine ⟨h₂, ?_⟩ + intro p'' i'' s'' h₃ h₄ + rw [← append_assoc] at h₃ + rw [h₃] at w + simpa using w (a :: p'') i'' s'' (by simp) h₄ + · rename_i s' + simp only [some.injEq, Prod.mk.injEq, append_assoc, ge_iff_le] + rw [dropPrefix?_eq_some_iff] at h + obtain ⟨p', h, w⟩ := h + constructor + · rintro ⟨rfl, rfl⟩ + simpa using ⟨p', by simp_all⟩ + · rintro ⟨p'', rfl, ⟨i', h₁, h₂⟩, w'⟩ + specialize w' [] p' s' (by simpa using h) w + simp at w' + simp [w'] at h₁ ⊢ + rw [h] at h₁ + apply append_inj_right h₁ + replace w := length_eq_of_beq w + replace h₂ := length_eq_of_beq h₂ + simp_all + +theorem dropInfix?_eq_some_iff [BEq α] {l i p s : List α} : + dropInfix? l i = some (p, s) ↔ + -- `i` is an infix up to `==` + (∃ i', l = p ++ i' ++ s ∧ i' == i) ∧ + -- and there is no shorter prefix for which that is the case + (∀ p' i' s', l = p' ++ i' ++ s' → i' == i → p'.length ≥ p.length) := by + unfold dropInfix? + rw [dropInfix?_go_eq_some_iff] + simp + +@[simp] theorem dropInfix?_nil [BEq α] {s : List α} : dropInfix? s [] = some ([], s) := by + simp [dropInfix?_eq_some_iff] + /-! ### deprecations -/ @[deprecated (since := "2024-08-15")] alias isEmpty_iff_eq_nil := isEmpty_iff From dfbe9387054e2972449de7bf346059d3609529fa Mon Sep 17 00:00:00 2001 From: Julian Berman Date: Fri, 29 Nov 2024 02:12:39 -0500 Subject: [PATCH 175/177] fix: dependency syntax for a reservoir+git lakefile.toml (#1067) --- README.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/README.md b/README.md index e20b87c09f..36d46eee4d 100644 --- a/README.md +++ b/README.md @@ -13,7 +13,7 @@ Or add the following to your `lakefile.toml`: [[require]] name = "batteries" scope = "leanprover-community" -version = "main" +rev = "main" ``` Additionally, please make sure that you're using the version of Lean that the current version of `batteries` expects. The easiest way to do this is to copy the [`lean-toolchain`](./lean-toolchain) file from this repository to your project. Once you've added the dependency declaration, the command `lake update` checks out the current version of `batteries` and writes it the Lake manifest file. Don't run this command again unless you're prepared to potentially also update your Lean compiler version, as it will retrieve the latest version of dependencies and add them to the manifest. From 5de63200cb5bb3460e09225874a6e2fe6ca9da5d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Fri, 29 Nov 2024 08:31:03 -0500 Subject: [PATCH 176/177] chore: bump toolchain to v4.14.0-rc3 (#1068) --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 0bef727630..6d9e70f733 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.14.0-rc1 +leanprover/lean4:v4.14.0-rc3 From f46c0445cc86ad2bc973edf8c5b2bd88bd91913b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Fri, 29 Nov 2024 09:41:07 -0500 Subject: [PATCH 177/177] chore: update docs-release action (#1069) --- .github/workflows/docs-release.yml | 2 ++ 1 file changed, 2 insertions(+) diff --git a/.github/workflows/docs-release.yml b/.github/workflows/docs-release.yml index 2926a31340..2c8a5c57f6 100644 --- a/.github/workflows/docs-release.yml +++ b/.github/workflows/docs-release.yml @@ -40,6 +40,8 @@ jobs: - name: Release Docs uses: softprops/action-gh-release@v2 with: + prerelease: ${{ contains(github.ref, 'rc') }} + make_latest: ${{ !contains(github.ref, 'rc') }} files: | docs/docs-${{ github.ref_name }}.tar.gz docs/docs-${{ github.ref_name }}.zip