diff --git a/.gitignore b/.gitignore index 9cf8557aec..84244fad11 100644 --- a/.gitignore +++ b/.gitignore @@ -5,8 +5,3 @@ /lakefile.olean # After v4.3.0-rc2 lake stores its files here: /.lake/ -# Note: because std has no dependencies, the lake-manifest contains no information -# that is not regenerated automatically by lake, so this avoids conflicts when -# lake decides to upgrade its manifest version. -# We may revisit this if lake decides to store more useful information in the manifest. -lake-manifest.json diff --git a/GNUmakefile b/GNUmakefile index a8ea2d849e..d0fba7074e 100644 --- a/GNUmakefile +++ b/GNUmakefile @@ -10,10 +10,13 @@ all: build test build: lake build -test: $(addsuffix .run, $(TESTS)) +# Find all .lean files in test and subdirectories, replace .lean with .run +TESTS := $(shell find test -type f -name '*.lean' | sed 's/\.lean$$/.run/') -test/%.run: build - lake env lean test/$* +test: $(TESTS) + +%.run: build + lake env lean $(@:.run=.lean) lint: build ./.lake/build/bin/runLinter diff --git a/README.md b/README.md index 8b5fa74ffe..b460508a8f 100644 --- a/README.md +++ b/README.md @@ -68,3 +68,7 @@ One of the easiest ways to contribute is to find a missing proof and complete it [`proof_wanted`](https://github.com/search?q=repo%3Aleanprover%2Fstd4+proof_wanted+language%3ALean&type=code&l=Lean) declaration documents statements that have been identified as being useful, but that have not yet been proven. + +In contrast to mathlib, `std` uses pull requests from forks of this repository. Hence, no special permissions on this repository are required for new contributors. + +You can change the labels on PRs by commenting one of `awaiting-review`, `awaiting-author`, or `WIP`. This is helpful for triage. diff --git a/Std.lean b/Std.lean index 2b6980dcf3..69530ea732 100644 --- a/Std.lean +++ b/Std.lean @@ -75,6 +75,7 @@ import Std.Lean.PersistentHashMap import Std.Lean.PersistentHashSet import Std.Lean.Position import Std.Lean.SMap +import Std.Lean.Syntax import Std.Lean.System.IO import Std.Lean.Tactic import Std.Lean.TagAttribute @@ -89,6 +90,7 @@ import Std.Tactic.Basic import Std.Tactic.ByCases import Std.Tactic.Case import Std.Tactic.Change +import Std.Tactic.Classical import Std.Tactic.CoeExt import Std.Tactic.Congr import Std.Tactic.Exact @@ -98,6 +100,7 @@ import Std.Tactic.FalseOrByContra import Std.Tactic.GuardExpr import Std.Tactic.GuardMsgs import Std.Tactic.HaveI +import Std.Tactic.Init import Std.Tactic.Instances import Std.Tactic.LabelAttr import Std.Tactic.LeftRight @@ -145,6 +148,7 @@ import Std.Tactic.Unreachable import Std.Tactic.Where import Std.Test.Internal.DummyLabelAttr import Std.Util.Cache +import Std.Util.CheckTactic import Std.Util.ExtendedBinder import Std.Util.LibraryNote import Std.Util.Pickle diff --git a/Std/CodeAction/Basic.lean b/Std/CodeAction/Basic.lean index 17f8b2cac8..7302aaa9c5 100644 --- a/Std/CodeAction/Basic.lean +++ b/Std/CodeAction/Basic.lean @@ -147,7 +147,7 @@ partial def findInfoTree? (kind : SyntaxNodeKind) (tgtRange : String.Range) (f : ContextInfo → Info → Bool) (canonicalOnly := false) : Option (ContextInfo × InfoTree) := match t with - | .context ctx t => findInfoTree? kind tgtRange ctx t f canonicalOnly + | .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 diff --git a/Std/CodeAction/Misc.lean b/Std/CodeAction/Misc.lean index 56b0914b3e..7c063fee5b 100644 --- a/Std/CodeAction/Misc.lean +++ b/Std/CodeAction/Misc.lean @@ -38,9 +38,17 @@ In the following: instance : Monad Id := _ ``` -invoking the hole code action "Generate a skeleton for the structure under construction." produces: +invoking the hole code action "Generate a (minimal) skeleton for the structure under construction." +produces: ```lean -instance : Monad Id := { +instance : Monad Id where + pure := _ + bind := _ +``` + +and invoking "Generate a (maximal) skeleton for the structure under construction." produces: +```lean +instance : Monad Id where map := _ mapConst := _ pure := _ @@ -48,22 +56,24 @@ instance : Monad Id := { seqLeft := _ seqRight := _ bind := _ -} ``` -/ -@[hole_code_action] partial def instanceStub : HoleCodeAction := fun params snap ctx info => do +@[hole_code_action] partial def instanceStub : HoleCodeAction := fun _ snap ctx info => do let some ty := info.expectedType? | return #[] let .const name _ := (← info.runMetaM ctx (whnf ty)).getAppFn | return #[] unless isStructure snap.env name do return #[] - let eager := { - title := "Generate a skeleton for the structure under construction." - kind? := "quickfix" - isPreferred? := true - } let doc ← readDoc - pure #[{ - eager - lazy? := some do + let fields := collectFields snap.env name #[] [] + let only := !fields.any fun (_, auto) => auto + let mkAutofix minimal := + let eager := { + title := s!"\ + Generate a {if only then "" else if minimal then "(minimal) " else "(maximal) "}\ + skeleton for the structure under construction." + kind? := "quickfix" + isPreferred? := minimal + } + let lazy? := some do let useWhere := do let _ :: (stx, _) :: _ ← findStack? snap.stx info.stx | none guard (stx.getKind == ``Parser.Command.declValSimple) @@ -73,7 +83,8 @@ instance : Monad Id := { let indent := "\n".pushn ' ' indent let mut str := if useWhere.isSome then "where" else "{" let mut first := useWhere.isNone && isStart - for field in collectFields snap.env name #[] do + for (field, auto) in fields do + if minimal && auto then continue if first then str := str ++ " " first := false @@ -92,14 +103,25 @@ instance : Monad Id := { newText := str } } - }] + { eager, lazy? } + pure <| if only then #[mkAutofix true] else #[mkAutofix true, mkAutofix false] where + /-- Returns true if this field is an autoParam or optParam, or if it is given an optional value + in a child struct. -/ + isAutofillable (env : Environment) (fieldInfo : StructureFieldInfo) (stack : List Name) : Bool := + fieldInfo.autoParam?.isSome || env.contains (mkDefaultFnOfProjFn fieldInfo.projFn) + || stack.any fun struct => env.contains (mkDefaultFnOfProjFn (struct ++ fieldInfo.fieldName)) + /-- Returns the fields of a structure, unfolding parent structures. -/ - collectFields (env : Environment) (structName : Name) (fields : Array Name) : Array Name := + collectFields (env : Environment) (structName : Name) + (fields : Array (Name × Bool)) (stack : List Name) : Array (Name × Bool) := (getStructureFields env structName).foldl (init := fields) fun fields field => - match isSubobjectField? env structName field with - | some substructName => collectFields env substructName fields - | none => fields.push field + if let some fieldInfo := getFieldInfo? env structName field then + if let some substructName := fieldInfo.subobject? then + collectFields env substructName fields (structName :: stack) + else + fields.push (field, isAutofillable env fieldInfo stack) + else fields /-- Returns the explicit arguments given a type. -/ def getExplicitArgs : Expr → Array Name → Array Name diff --git a/Std/Data/Array/Basic.lean b/Std/Data/Array/Basic.lean index 6b82ab02a3..04975b66a2 100644 --- a/Std/Data/Array/Basic.lean +++ b/Std/Data/Array/Basic.lean @@ -36,6 +36,18 @@ 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. +If `start` and `stop` are given, only the subarray `xs[start:stop]` is +considered (in addition to `d`). +-/ +@[inline] +protected def minWith [ord : Ord α] + (xs : Array α) (d : α) (start := 0) (stop := xs.size) : α := + xs.foldl (init := d) (start := start) (stop := stop) fun min x => + if compare x min |>.isLT then x else min + set_option linter.unusedVariables.funArgs false in /-- Find the first minimal element of an array. If the array is empty, `d` is @@ -45,8 +57,10 @@ considered. @[inline] protected def minD [ord : Ord α] (xs : Array α) (d : α) (start := 0) (stop := xs.size) : α := - xs.foldl (init := d) (start := start) (stop := stop) fun min x => - if compare x min |>.isLT then x else min + if h: start < xs.size ∧ start < stop then + xs.minWith (xs.get ⟨start, h.left⟩) (start + 1) stop + else + d set_option linter.unusedVariables.funArgs false in /-- @@ -57,8 +71,8 @@ considered. @[inline] protected def min? [ord : Ord α] (xs : Array α) (start := 0) (stop := xs.size) : Option α := - if h : start < xs.size then - some $ xs.minD (xs.get ⟨start, h⟩) start stop + if h : start < xs.size ∧ start < stop then + some $ xs.minD (xs.get ⟨start, h.left⟩) start stop else none @@ -73,6 +87,17 @@ protected def minI [ord : Ord α] [Inhabited α] (xs : Array α) (start := 0) (stop := xs.size) : α := xs.minD default start stop +set_option linter.unusedVariables.funArgs false in +/-- +Returns the first maximal element among `d` and elements of the array. +If `start` and `stop` are given, only the subarray `xs[start:stop]` is +considered (in addition to `d`). +-/ +@[inline] +protected def maxWith [ord : Ord α] + (xs : Array α) (d : α) (start := 0) (stop := xs.size) : α := + xs.minWith (ord := ord.opposite) d start stop + set_option linter.unusedVariables.funArgs false in /-- Find the first maximal element of an array. If the array is empty, `d` is diff --git a/Std/Data/Array/Init/Basic.lean b/Std/Data/Array/Init/Basic.lean index 444b5aec4e..5b44d9227d 100644 --- a/Std/Data/Array/Init/Basic.lean +++ b/Std/Data/Array/Init/Basic.lean @@ -29,7 +29,7 @@ def ofFn {n} (f : Fin n → α) : Array α := go 0 (mkEmpty n) where /-- Auxiliary for `ofFn`. `ofFn.go f i acc = acc ++ #[f i, ..., f(n - 1)]` -/ go (i : Nat) (acc : Array α) : Array α := if h : i < n then go (i+1) (acc.push (f ⟨i, h⟩)) else acc -termination_by _ => n - i +termination_by n - i /-- The array `#[0, 1, ..., n - 1]`. -/ def range (n : Nat) : Array Nat := diff --git a/Std/Data/Array/Init/Lemmas.lean b/Std/Data/Array/Init/Lemmas.lean index ba9ae7ec4e..01d6358ccc 100644 --- a/Std/Data/Array/Init/Lemmas.lean +++ b/Std/Data/Array/Init/Lemmas.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro -/ import Std.Tactic.NoMatch import Std.Tactic.HaveI +import Std.Tactic.ByCases import Std.Classes.LawfulMonad import Std.Data.Fin.Init.Lemmas import Std.Data.Nat.Init.Lemmas @@ -152,7 +153,7 @@ where · rw [← List.get_drop_eq_drop _ i ‹_›] simp [aux (i+1), map_eq_pure_bind]; rfl · rw [List.drop_length_le (Nat.ge_of_not_lt ‹_›)]; rfl -termination_by aux => arr.size - i + termination_by arr.size - i theorem SatisfiesM_mapM [Monad m] [LawfulMonad m] (as : Array α) (f : α → m β) (motive : Nat → Prop) (h0 : motive 0) @@ -262,9 +263,9 @@ theorem SatisfiesM_anyM [Monad m] [LawfulMonad m] (p : α → m Bool) (as : Arra | true, h => .pure h | false, h => go hj hstop h hp · next hj => exact .pure <| Nat.le_antisymm hj' (Nat.ge_of_not_lt hj) ▸ h0 + termination_by stop - j simp only [Array.anyM_eq_anyM_loop] exact go hstart _ h0 fun i hi => hp i <| Nat.lt_of_lt_of_le hi <| Nat.min_le_left .. -termination_by go => stop - j theorem SatisfiesM_anyM_iff_exists [Monad m] [LawfulMonad m] (p : α → m Bool) (as : Array α) (start stop) (q : Fin as.size → Prop) diff --git a/Std/Data/Array/Lemmas.lean b/Std/Data/Array/Lemmas.lean index 4a85a12919..101a85889d 100644 --- a/Std/Data/Array/Lemmas.lean +++ b/Std/Data/Array/Lemmas.lean @@ -7,6 +7,7 @@ Authors: Mario Carneiro, Gabriel Ebner import Std.Data.Nat.Lemmas import Std.Data.List.Lemmas import Std.Data.Array.Basic +import Std.Tactic.SeqFocus import Std.Tactic.HaveI import Std.Tactic.Simpa import Std.Util.ProofWanted @@ -41,10 +42,26 @@ attribute [simp] isEmpty uget @[simp] theorem data_length {l : Array α} : l.data.length = l.size := rfl +/-- # mem -/ + theorem mem_data {a : α} {l : Array α} : a ∈ l.data ↔ a ∈ l := (mem_def _ _).symm theorem not_mem_nil (a : α) : ¬ a ∈ #[] := fun. +/-- # set -/ + +@[simp] theorem set!_is_setD : @set! = @setD := rfl + +@[simp] theorem size_setD (a : Array α) (index : Nat) (val : α) : + (Array.setD a index val).size = a.size := by + if h : index < a.size then + simp [setD, h] + else + simp [setD, h] + + +/-- # get lemmas -/ + theorem getElem?_mem {l : Array α} {i : Fin l.size} : l[i] ∈ l := by erw [Array.mem_def, getElem_eq_data_get] apply List.get_mem @@ -131,6 +148,22 @@ theorem get_set (a : Array α) (i : Fin a.size) (j : Nat) (hj : j < a.size) (v : (a.set i v)[j]'(by simp [*]) = if i = j then v else a[j] := by if h : i.1 = j then subst j; simp [*] else simp [*] +@[simp] theorem getElem_setD (a : Array α) (i : Nat) (v : α) (h : i < (setD a i v).size) : + (setD a i v)[i]'h = v := by + simp at h + simp only [setD, h, dite_true, get_set, ite_true] + +/-- +This lemma simplifies a normal form from `get!` +-/ +@[simp] theorem getD_get?_setD (a : Array α) (i : Nat) (v d : α) : + Option.getD (setD a i v)[i]? d = if i < a.size then v else d := by + if h : i < a.size then + simp [setD, h, getElem?, get_set] + else + have p : i ≥ a.size := Nat.le_of_not_gt h + simp [setD, h, get?_len_le, p] + theorem set_set (a : Array α) (i : Fin a.size) (v v' : α) : (a.set i v).set ⟨i, by simp [i.2]⟩ v' = a.set i v' := by simp [set, List.set_set] @@ -266,8 +299,8 @@ theorem mapIdx_induction' (as : Array α) (f : Fin as.size → α → β) theorem size_eq_length_data (as : Array α) : as.size = as.data.length := rfl -@[simp] theorem size_swap! (a : Array α) (i j) (hi : i < a.size) (hj : j < a.size) : - (a.swap! i j).size = a.size := by simp [swap!, hi, hj] +@[simp] theorem size_swap! (a : Array α) (i j) : + (a.swap! i j).size = a.size := by unfold swap!; split <;> (try split) <;> simp [size_swap] @[simp] theorem size_reverse (a : Array α) : a.reverse.size = a.size := by let rec go (as : Array α) (i j) : (reverse.loop as i j).size = as.size := by @@ -276,8 +309,8 @@ theorem size_eq_length_data (as : Array α) : as.size = as.data.length := rfl have := reverse.termination h simp [(go · (i+1) ⟨j-1, ·⟩), h] else simp [h] + termination_by j - i simp only [reverse]; split <;> simp [go] -termination_by _ => j - i @[simp] theorem size_range {n : Nat} : (range n).size = n := by unfold range @@ -323,6 +356,7 @@ theorem size_modifyM [Monad m] [LawfulMonad m] (a : Array α) (i : Nat) (f : α cases Nat.le_antisymm h₂.1 h₂.2 exact (List.get?_reverse' _ _ h).symm · rfl + termination_by j - i simp only [reverse]; split · match a with | ⟨[]⟩ | ⟨[_]⟩ => rfl · have := Nat.sub_add_cancel (Nat.le_of_not_le ‹_›) @@ -330,7 +364,6 @@ theorem size_modifyM [Monad m] [LawfulMonad m] (a : Array α) (i : Nat) (f : α split; {rfl}; rename_i h simp [← show k < _ + 1 ↔ _ from Nat.lt_succ (n := a.size - 1), this] at h rw [List.get?_eq_none.2 ‹_›, List.get?_eq_none.2 (a.data.length_reverse ▸ ‹_›)] -termination_by _ => j - i @[simp] theorem size_ofFn_go {n} (f : Fin n → α) (i acc) : (ofFn.go f i acc).size = acc.size + (n - i) := by @@ -343,7 +376,7 @@ termination_by _ => j - i have : n - i = 0 := Nat.sub_eq_zero_of_le (Nat.le_of_not_lt hin) unfold ofFn.go simp [hin, this] -termination_by _ => n - i +termination_by n - i @[simp] theorem size_ofFn (f : Fin n → α) : (ofFn f).size = n := by simp [ofFn] @@ -363,7 +396,7 @@ theorem getElem_ofFn_go (f : Fin n → α) (i) {acc k} | inr hj => simp [get_push, *] else simp [hin, hacc k (Nat.lt_of_lt_of_le hki (Nat.le_of_not_lt (hi ▸ hin)))] -termination_by _ => n - i +termination_by n - i @[simp] theorem getElem_ofFn (f : Fin n → α) (i : Nat) (h) : (ofFn f)[i] = f ⟨i, size_ofFn f ▸ h⟩ := @@ -427,8 +460,8 @@ theorem zipWith_eq_zipWith_data (f : α → β → γ) (as : Array α) (bs : Arr ((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 [zipWith, loop 0 #[] (by simp) (by simp)] -termination_by loop i _ _ _ => as.size - i theorem size_zipWith (as : Array α) (bs : Array β) (f : α → β → γ) : (as.zipWith bs f).size = min as.size bs.size := by diff --git a/Std/Data/Array/Match.lean b/Std/Data/Array/Match.lean index 36432a756b..f5ff981777 100644 --- a/Std/Data/Array/Match.lean +++ b/Std/Data/Array/Match.lean @@ -45,7 +45,7 @@ def PrefixTable.step [BEq α] (t : PrefixTable α) (x : α) : Fin (t.size+1) → ⟨k+1, Nat.succ_lt_succ hsz⟩ else cont () else cont () -termination_by _ k => k.val +termination_by k => k.val /-- Extend a prefix table by one element diff --git a/Std/Data/Array/Merge.lean b/Std/Data/Array/Merge.lean index be301a9fb9..fc7727be02 100644 --- a/Std/Data/Array/Merge.lean +++ b/Std/Data/Array/Merge.lean @@ -40,7 +40,7 @@ where have : xs.size + ys.size - (i + j + 1) < xs.size + ys.size - (i + j) := Nat.sub_succ_lt_self _ _ hij go (acc.push y) i (j + 1) -termination_by go => xs.size + ys.size - (i + j) + termination_by xs.size + ys.size - (i + j) /-- Merge arrays `xs` and `ys`, which must be sorted according to `compare` and must @@ -83,7 +83,7 @@ where rw [show i + j + 2 = (i + 1) + (j + 1) by simp_arith] exact Nat.add_le_add hi hj go (acc.push (merge x y)) (i + 1) (j + 1) -termination_by go => xs.size + ys.size - (i + j) + termination_by xs.size + ys.size - (i + j) /-- Merge arrays `xs` and `ys`, which must be sorted according to `compare` and must @@ -130,7 +130,7 @@ where go (acc.push hd) (i + 1) x else acc.push hd -termination_by _ i _ => xs.size - i + termination_by xs.size - i /-- Deduplicate a sorted array. The array must be sorted with to an order which diff --git a/Std/Data/BinomialHeap/Basic.lean b/Std/Data/BinomialHeap/Basic.lean index e21a79d9bc..8f2ed693bb 100644 --- a/Std/Data/BinomialHeap/Basic.lean +++ b/Std/Data/BinomialHeap/Basic.lean @@ -140,7 +140,7 @@ by rank and `merge` maintains this invariant. else if t₂.rankGT r then merge le t₁ (.cons r a n t₂) else .cons r a n (merge le t₁ t₂) -termination_by _ s₁ s₂ => s₁.length + s₂.length +termination_by s₁ s₂ => s₁.length + s₂.length /-- `O(log n)`. Convert a `HeapNode` to a `Heap` by reversing the order of the nodes @@ -219,7 +219,7 @@ theorem Heap.realSize_merge (le) (s₁ s₂ : Heap α) : rw [combine] at eq; split at eq <;> cases eq <;> simp [Nat.add_assoc, Nat.add_left_comm, Nat.add_comm] split <;> split <;> simp [IH₁, IH₂, IH₃, this, Nat.add_assoc, Nat.add_left_comm] -termination_by _ => s₁.length + s₂.length +termination_by s₁.length + s₂.length private def FindMin.HasSize (res : FindMin α) (n : Nat) : Prop := ∃ m, @@ -280,7 +280,7 @@ by repeatedly pulling the minimum element out of the heap. | some (hd, tl) => do have : tl.realSize < s.realSize := by simp_arith [Heap.realSize_deleteMin eq] foldM le tl (← f init hd) f -termination_by _ => s.realSize +termination_by s.realSize /-- `O(n log n)`. Fold over the elements of a heap in increasing order, @@ -402,7 +402,7 @@ theorem Heap.WF.merge' (h₁ : s₁.WF le n) (h₂ : s₂.WF le n) : · let ⟨ih₁, ih₂⟩ := merge' ht₁ ht₂ exact ⟨⟨Nat.le_succ_of_le hr₁, this, ih₁.of_rankGT (ih₂ (iff_of_false hl₁ hl₂))⟩, fun _ => Nat.lt_succ_of_le hr₁⟩ -termination_by _ => s₁.length + s₂.length +termination_by s₁.length + s₂.length theorem Heap.WF.merge (h₁ : s₁.WF le n) (h₂ : s₂.WF le n) : (merge le s₁ s₂).WF le n := (merge' h₁ h₂).1 diff --git a/Std/Data/BitVec/Basic.lean b/Std/Data/BitVec/Basic.lean index 1aee53e59a..3c9e8d07bc 100644 --- a/Std/Data/BitVec/Basic.lean +++ b/Std/Data/BitVec/Basic.lean @@ -44,7 +44,9 @@ namespace BitVec /-- The `BitVec` with value `i mod 2^n`. Treated as an operation on bitvectors, this is truncation of the high bits when downcasting and zero-extension when upcasting. -/ protected def ofNat (n : Nat) (i : Nat) : BitVec n where - toFin := Fin.ofNat' i (Nat.pow_two_pos n) + toFin := Fin.ofNat' i (Nat.two_pow_pos n) + +instance : NatCast (BitVec w) := ⟨BitVec.ofNat w⟩ /-- Given a bitvector `a`, return the underlying `Nat`. This is O(1) because `BitVec` is a (zero-cost) wrapper around a `Nat`. -/ @@ -67,7 +69,7 @@ protected def toInt (a : BitVec n) : Int := if a.msb then Int.ofNat a.toNat - Int.ofNat (2^n) else a.toNat /-- Return a bitvector `0` of size `n`. This is the bitvector with all zero bits. -/ -protected def zero (n : Nat) : BitVec n := ⟨0, Nat.pow_two_pos n⟩ +protected def zero (n : Nat) : BitVec n := ⟨0, Nat.two_pow_pos n⟩ instance : Inhabited (BitVec n) where default := .zero n @@ -231,6 +233,8 @@ SMT-Lib name: `bvult`. -/ protected def ult (x y : BitVec n) : Bool := x.toFin < y.toFin instance : LT (BitVec n) where lt x y := x.toFin < y.toFin +instance (x y : BitVec n) : Decidable (x < y) := + inferInstanceAs (Decidable (x.toFin < y.toFin)) /-- Unsigned less-than-or-equal-to for bit vectors. @@ -240,6 +244,8 @@ SMT-Lib name: `bvule`. protected def ule (x y : BitVec n) : Bool := x.toFin ≤ y.toFin instance : LE (BitVec n) where le x y := x.toFin ≤ y.toFin +instance (x y : BitVec n) : Decidable (x ≤ y) := + inferInstanceAs (Decidable (x.toFin ≤ y.toFin)) /-- Signed less-than for bit vectors. @@ -307,10 +313,7 @@ Bitwise NOT for bit vectors. SMT-Lib name: `bvnot`. -/ protected def not (x : BitVec n) : BitVec n := - let ones := .ofFin ⟨(1 <<< n).pred, - n.one_shiftLeft.symm ▸ - (Nat.pred_lt <| Nat.ne_of_gt <| Nat.pos_pow_of_pos _ <| Nat.zero_lt_succ _)⟩; - ones ^^^ x + allOnes n ^^^ x instance : Complement (BitVec w) := ⟨.not⟩ /-- The `BitVec` with value `(2^n + (i mod 2^n)) mod 2^n`. -/ @@ -319,6 +322,8 @@ protected def ofInt (n : Nat) (i : Int) : BitVec n := | Int.ofNat a => .ofNat n a | Int.negSucc a => ~~~.ofNat n a +instance : IntCast (BitVec w) := ⟨BitVec.ofInt w⟩ + /-- Left shift for bit vectors. The low bits are filled with zeros. As a numeric operation, this is equivalent to `a * 2^s`, modulo `2^n`. @@ -334,7 +339,13 @@ As a numeric operation, this is equivalent to `a / 2^s`, rounding down. SMT-Lib name: `bvlshr` except this operator uses a `Nat` shift value. -/ -def ushiftRight (a : BitVec n) (s : Nat) : BitVec n := .ofNat n (a.toNat >>> s) +def ushiftRight (a : BitVec n) (s : Nat) : BitVec n := + ⟨a.toNat >>> s, by + let ⟨a, lt⟩ := a + simp only [BitVec.toNat, Nat.shiftRight_eq_div_pow, Nat.div_lt_iff_lt_mul (Nat.two_pow_pos s)] + rw [←Nat.mul_one a] + exact Nat.mul_lt_mul_of_lt_of_le' lt (Nat.two_pow_pos s) (Nat.le_refl 1)⟩ + instance : HShiftRight (BitVec w) Nat (BitVec w) := ⟨.ushiftRight⟩ /-- @@ -387,7 +398,7 @@ def shiftLeftZeroExtend (msbs : BitVec w) (m : Nat) : BitVec (w+m) := let shiftLeftLt {x : Nat} (p : x < 2^w) (m : Nat) : x <<< m < 2^(w+m) := by simp [Nat.shiftLeft_eq, Nat.pow_add] apply Nat.mul_lt_mul_of_pos_right p - exact (Nat.pow_two_pos m) + exact (Nat.two_pow_pos m) ⟨msbs.toNat <<< m, shiftLeftLt msbs.isLt m⟩ /-- diff --git a/Std/Data/BitVec/Bitblast.lean b/Std/Data/BitVec/Bitblast.lean index 09cd93255d..ef4f57c75f 100644 --- a/Std/Data/BitVec/Bitblast.lean +++ b/Std/Data/BitVec/Bitblast.lean @@ -19,11 +19,13 @@ The module is named for the bit-blasting operation in an SMT solver that convert expressions into expressions about individual bits in each vector. ## Main results -* `x + y : BitVec w` is equivalent to `adc x y false`. +* `x + y : BitVec w` is `(adc x y false).2`. + ## Future work All other operations are to be PR'ed later and are already proved in -https://github.com/mhk119/lean-smt/blob/bitvec/Smt/Data/BitVec.lean. +https://github.com/mhk119/lean-smt/blob/bitvec/Smt/Data/Bitwise.lean. + -/ open Nat Bool @@ -59,7 +61,7 @@ private theorem mod_two_pow_succ (x i : Nat) : apply Nat.eq_of_testBit_eq intro j simp only [Nat.mul_add_lt_is_or, testBit_or, testBit_mod_two_pow, testBit_shiftLeft, - Nat.testBit_bool_to_nat, Nat.sub_eq_zero_iff_le, Nat.mod_lt, Nat.pow_two_pos, + Nat.testBit_bool_to_nat, Nat.sub_eq_zero_iff_le, Nat.mod_lt, Nat.two_pow_pos, testBit_mul_pow_two] rcases Nat.lt_trichotomy i j with i_lt_j | i_eq_j | j_lt_i · have i_le_j : i ≤ j := Nat.le_of_lt i_lt_j @@ -74,46 +76,43 @@ private theorem mod_two_pow_succ (x i : Nat) : have not_j_ge_i : ¬(j ≥ i) := Nat.not_le_of_lt j_lt_i simp [j_lt_i, j_le_i, not_j_ge_i, j_le_i_succ] -private theorem mod_two_pow_lt (x i : Nat) : x % 2 ^ i < 2^i := Nat.mod_lt _ (Nat.pow_two_pos _) +private theorem mod_two_pow_lt (x i : Nat) : x % 2 ^ i < 2^i := Nat.mod_lt _ (Nat.two_pow_pos _) /-! ### Addition -/ /-- carry w x y c returns true if the `w` carry bit is true when computing `x + y + c`. -/ -def carry (w x y : Nat) (c : Bool) : Bool := decide (x % 2^w + y % 2^w + c.toNat ≥ 2^w) +def carry (w x y : Nat) (c : Bool) : Bool := decide (x % 2^w + y % 2^w + c.toNat ≥ 2^w) @[simp] theorem carry_zero : carry 0 x y c = c := by cases c <;> simp [carry, mod_one] +/-- At least two out of three booleans are true. -/ +abbrev atLeastTwo (a b c : Bool) : Bool := a && b || a && c || b && c + /-- Carry function for bitwise addition. -/ -def adcb (x y c : Bool) : Bool × Bool := (x && y || x && c || y && c, Bool.xor x (Bool.xor y c)) +def adcb (x y c : Bool) : Bool × Bool := (atLeastTwo x y c, Bool.xor x (Bool.xor y c)) /-- Bitwise addition implemented via a ripple carry adder. -/ def adc (x y : BitVec w) : Bool → Bool × BitVec w := iunfoldr fun (i : Fin w) c => adcb (x.getLsb i) (y.getLsb i) c theorem adc_overflow_limit (x y i : Nat) (c : Bool) : x % 2^i + (y % 2^i + c.toNat) < 2^(i+1) := by - apply Nat.lt_of_succ_le - simp only [←Nat.succ_add, Nat.pow_succ, Nat.mul_two] - apply Nat.add_le_add (mod_two_pow_lt _ _) - apply Nat.le_trans - exact (Nat.add_le_add_left (Bool.toNat_le_one c) _) - exact Nat.mod_lt _ (Nat.pow_two_pos i) - -theorem carry_succ (w x y : Nat) (c : Bool) : carry (succ w) x y c = - decide ((x.testBit w).toNat + (y.testBit w).toNat + (carry w x y c).toNat ≥ 2) := by - simp only [carry, mod_two_pow_succ _ w, decide_eq_decide] + have : c.toNat ≤ 1 := Bool.toNat_le_one c + rw [Nat.pow_succ] + omega + +theorem carry_succ (w x y : Nat) (c : Bool) : + carry (succ w) x y c = atLeastTwo (x.testBit w) (y.testBit w) (carry w x y c) := by + simp only [carry, mod_two_pow_succ, atLeastTwo] + simp only [Nat.pow_succ'] generalize testBit x w = xh generalize testBit y w = yh have sum_bnd : x%2^w + (y%2^w + c.toNat) < 2*2^w := by - simp [Nat.mul_comm 2 _, ←Nat.pow_succ ] + simp only [← Nat.pow_succ'] exact adc_overflow_limit x y w c - simp only [Nat.pow_succ] - cases xh <;> cases yh <;> cases Decidable.em (x%2^w + (y%2^w + toNat c) ≥ 2 ^ w) with | _ pred => - simp [Nat.one_shiftLeft, Nat.add_assoc, Nat.add_left_comm _ (2^_) _, Nat.mul_comm (2^_) _, - Nat.not_le_of_lt, Nat.add_succ, Nat.succ_le_succ, - mul_le_add_right, le_add_right, sum_bnd, pred] + cases xh <;> cases yh <;> (simp; omega) -theorem adc_value_step {i : Nat} (i_lt : i < w) (x y : BitVec w) (c : Bool) : +theorem getLsb_add_add_bool {i : Nat} (i_lt : i < w) (x y : BitVec w) (c : Bool) : getLsb (x + y + zeroExtend w (ofBool c)) i = Bool.xor (getLsb x i) (Bool.xor (getLsb y i) (carry i x.toNat y.toNat c)) := by let ⟨x, x_lt⟩ := x @@ -134,7 +133,12 @@ theorem adc_value_step {i : Nat} (i_lt : i < w) (x y : BitVec w) (c : Bool) : ] simp [testBit_to_div_mod, carry, Nat.add_assoc] -theorem adc_correct (x y : BitVec w) (c : Bool) : +theorem getLsb_add {i : Nat} (i_lt : i < w) (x y : BitVec w) : + getLsb (x + y) i = + Bool.xor (getLsb x i) (Bool.xor (getLsb y i) (carry i x.toNat y.toNat false)) := by + simpa using getLsb_add_add_bool i_lt x y false + +theorem adc_spec (x y : BitVec w) (c : Bool) : adc x y c = (carry w x.toNat y.toNat c, x + y + zeroExtend w (ofBool c)) := by simp only [adc] apply iunfoldr_replace @@ -150,14 +154,26 @@ theorem adc_correct (x y : BitVec w) (c : Bool) : apply And.intro case left => rw [testBit_toNat, testBit_toNat] - cases x.getLsb i <;> - cases y.getLsb i <;> - cases carry i x.toNat y.toNat c <;> simp [Nat.succ_le_succ_iff] case right => - simp [adc_value_step lt] + simp [getLsb_add_add_bool lt] + +theorem add_eq_adc (w : Nat) (x y : BitVec w) : x + y = (adc x y false).snd := by + simp [adc_spec] + +/-! ### add -/ theorem add_as_adc (w : Nat) (x y : BitVec w) : x + y = (adc x y false).snd := by simp [adc_correct] + + /-- Adding a bitvector to its own complement yields the all ones bitpattern -/ +@[simp] theorem add_not_self (x : BitVec w) : x + ~~~x = allOnes w := by + rw [add_eq_adc, adc, iunfoldr_replace (fun _ => false) (allOnes w)] + · rfl + · simp [adcb, atLeastTwo] + +/-- Subtracting `x` from the all ones bitvector is equivalent to taking its complement -/ +theorem allOnes_sub_eq_not (x : BitVec w) : allOnes w - x = ~~~x := by + rw [← add_not_self x, BitVec.add_comm, add_sub_cancel] /-! ### Negation -/ @@ -204,3 +220,4 @@ theorem bit_neg_eq_neg (x : BitVec w) : bit_neg x = -x := by · rw [eq_sub_of_add_eq (bit_not_add_self x), add_comm, ← add_sub_assoc] simp [add_right_neg, zero_sub] · simp [bit_not_testBit x _] + diff --git a/Std/Data/BitVec/Folds.lean b/Std/Data/BitVec/Folds.lean index 4f661bf81b..2811ee46cd 100644 --- a/Std/Data/BitVec/Folds.lean +++ b/Std/Data/BitVec/Folds.lean @@ -1,3 +1,8 @@ +/- +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 Std.Data.BitVec.Lemmas import Std.Data.Fin.Iterate import Std.Data.Nat.Lemmas diff --git a/Std/Data/BitVec/Lemmas.lean b/Std/Data/BitVec/Lemmas.lean index 05b8b98b92..890c2cd352 100644 --- a/Std/Data/BitVec/Lemmas.lean +++ b/Std/Data/BitVec/Lemmas.lean @@ -1,8 +1,17 @@ +/- +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 Std.Data.Bool import Std.Data.BitVec.Basic +import Std.Data.Fin.Lemmas import Std.Data.Nat.Lemmas +import Std.Tactic.Ext import Std.Tactic.Simpa +import Std.Tactic.Omega +import Std.Util.ProofWanted namespace Std.BitVec @@ -10,18 +19,37 @@ namespace Std.BitVec theorem eq_of_toNat_eq {n} : ∀ {i j : BitVec n}, i.toNat = j.toNat → i = j | ⟨_, _⟩, ⟨_, _⟩, rfl => rfl +/-- Replaced 2024-02-07. -/ +@[deprecated] alias zero_is_unique := eq_nil + +@[simp] theorem val_toFin (x : BitVec w) : x.toFin.val = x.toNat := rfl + theorem toNat_eq (x y : BitVec n) : x = y ↔ x.toNat = y.toNat := Iff.intro (congrArg BitVec.toNat) eq_of_toNat_eq +theorem toNat_lt (x : BitVec n) : x.toNat < 2^n := x.toFin.2 + theorem testBit_toNat (x : BitVec w) : x.toNat.testBit i = x.getLsb i := rfl +@[simp] theorem getLsb_ofFin (x : Fin (2^n)) (i : Nat) : + getLsb (BitVec.ofFin x) i = x.val.testBit i := rfl + @[simp] theorem getLsb_ge (x : BitVec w) (i : Nat) (ge : i ≥ w) : getLsb x i = false := by let ⟨x, x_lt⟩ := x - apply Nat.testBit_lt_two - apply Nat.lt_of_lt_of_le x_lt - apply Nat.pow_le_pow_of_le_right (by trivial : 0 < 2) ge + simp + apply Nat.testBit_lt_two_pow + have p : 2^w ≤ 2^i := Nat.pow_le_pow_of_le_right (by omega) ge + omega + +theorem lt_of_getLsb (x : BitVec w) (i : Nat) : getLsb x i = true → i < w := by + if h : i < w then + simp [h] + else + simp [Nat.ge_of_not_lt h] -theorem eq_of_getLsb_eq {x y : BitVec w} +-- We choose `eq_of_getLsb_eq` as the `@[ext]` theorem for `BitVec` +-- somewhat arbitrarily over `eq_of_getMsg_eq`. +@[ext] theorem eq_of_getLsb_eq {x y : BitVec w} (pred : ∀(i : Fin w), x.getLsb i.val = y.getLsb i.val) : x = y := by apply eq_of_toNat_eq apply Nat.eq_of_testBit_eq @@ -51,22 +79,35 @@ theorem eq_of_getMsb_eq {x y : BitVec w} have q := pred ⟨w - 1 - i, q_lt⟩ simpa [q_lt, Nat.sub_sub_self, r] using q +theorem eq_of_toFin_eq {x y : BitVec w} (w : x.toFin = y.toFin) : x = y := by + ext + simp only [← testBit_toNat, ← val_toFin, w] + @[simp] theorem toNat_ofBool (b : Bool) : (ofBool b).toNat = b.toNat := by cases b <;> rfl +theorem ofNat_one (n : Nat) : BitVec.ofNat 1 n = BitVec.ofBool (n % 2 = 1) := by + rcases (Nat.mod_two_eq_zero_or_one n) with h | h <;> simp [h, BitVec.ofNat, Fin.ofNat'] + @[simp] theorem toNat_ofFin (x : Fin (2^n)) : (BitVec.ofFin x).toNat = x.val := rfl -@[simp] theorem val_toFin_eq_toNat (x : BitVec w) : (toFin x).val = x.toNat := rfl -theorem toNat_ofNat (x w : Nat) : (BitVec.ofNat w x).toNat = x % 2^w := by +@[simp] theorem toNat_ofNat (x w : Nat) : (x#w).toNat = x % 2^w := by simp [BitVec.toNat, BitVec.ofNat, Fin.ofNat'] -@[simp] theorem toNat_zero (n : Nat) : (0#n).toNat = 0 := by trivial +@[simp] theorem getLsb_ofNat (n : Nat) (x : Nat) (i : Nat) : + getLsb (x#n) i = (i < n && x.testBit i) := by + simp [getLsb, BitVec.ofNat, Fin.val_ofNat'] + +@[deprecated toNat_ofNat] theorem toNat_zero (n : Nat) : (0#n).toNat = 0 := by trivial + +@[simp] theorem toNat_mod_cancel (x : BitVec n) : x.toNat % (2^n) = x.toNat := + Nat.mod_eq_of_lt x.isLt private theorem lt_two_pow_of_le {x m n : Nat} (lt : x < 2 ^ m) (le : m ≤ n) : x < 2 ^ n := Nat.lt_of_lt_of_le lt (Nat.pow_le_pow_of_le_right (by trivial : 0 < 2) le) -@[simp] theorem ofNat_toNat (m : Nat) (x : BitVec n) : (BitVec.ofNat m x.toNat) = truncate m x := by +@[simp] theorem ofNat_toNat (m : Nat) (x : BitVec n) : x.toNat#m = truncate m x := by let ⟨x, lt_n⟩ := x unfold truncate unfold zeroExtend @@ -77,31 +118,13 @@ private theorem lt_two_pow_of_le {x m n : Nat} (lt : x < 2 ^ m) (le : m ≤ n) : else simp [h] -theorem toNat_append (x : BitVec m) (y : BitVec n) : (x ++ y).toNat = x.toNat <<< n ||| y.toNat := - rfl +/-! ### cast -/ @[simp] theorem toNat_cast (e : m = n) (x : BitVec m) : (cast e x).toNat = x.toNat := rfl -theorem toNat_cons (b : Bool) (x : BitVec w) : (cons b x).toNat = (b.toNat <<< w) ||| x.toNat := by - let ⟨x, _⟩ := x - simp [cons, toNat_append, toNat_ofBool] - -/-! ### cons -/ - -theorem getLsb_cons (b : Bool) {n} (x : BitVec n) (i : Nat) : - getLsb (cons b x) i = if i = n then b else getLsb x i := by - simp only [getLsb, toNat_cons, Nat.testBit_or] - rw [Nat.testBit_shiftLeft] - rcases Nat.lt_trichotomy i n with i_lt_n | i_eq_n | n_lt_i - · have i_ge_n := Nat.not_le_of_gt i_lt_n - have not_eq := Nat.ne_of_lt i_lt_n - simp [i_ge_n, not_eq] - · simp [i_eq_n, testBit_toNat] - cases b <;> trivial - · have i_ge_n : n ≤ i := Nat.le_of_lt n_lt_i - have not_eq := Nat.ne_of_gt n_lt_i - have neq_zero : i - n ≠ 0 := Nat.sub_ne_zero_of_lt n_lt_i - simp [i_ge_n, not_eq, Nat.testBit_bool_to_nat, neq_zero] +@[simp] theorem getLsb_cast : (cast h v).getLsb i = v.getLsb i := by + cases h + simp /-! ### zeroExtend and truncate -/ @@ -134,18 +157,231 @@ theorem toNat_zeroExtend (i : Nat) (x : BitVec n) : @[simp] theorem toNat_truncate (x : BitVec n) : (truncate i x).toNat = x.toNat % 2^i := toNat_zeroExtend i x -theorem getLsb_zeroExtend' (ge : m ≥ n) (x : BitVec n) (i : Nat) : +@[simp] theorem getLsb_zeroExtend' (ge : m ≥ n) (x : BitVec n) (i : Nat) : getLsb (zeroExtend' ge x) i = getLsb x i := by simp [getLsb, toNat_zeroExtend'] -theorem getLsb_zeroExtend (m : Nat) (x : BitVec n) (i : Nat) : +@[simp] theorem getLsb_zeroExtend (m : Nat) (x : BitVec n) (i : Nat) : getLsb (zeroExtend m x) i = (decide (i < m) && getLsb x i) := by simp [getLsb, toNat_zeroExtend, Nat.testBit_mod_two_pow] -theorem getLsb_truncate (m : Nat) (x : BitVec n) (i : Nat) : +@[simp] theorem getLsb_truncate (m : Nat) (x : BitVec n) (i : Nat) : getLsb (truncate m x) i = (decide (i < m) && getLsb x i) := getLsb_zeroExtend m x i +/-! ## extractLsb -/ + +@[simp] +protected theorem extractLsb_ofFin {n} (x : Fin (2^n)) (hi lo : Nat) : + extractLsb hi lo (@BitVec.ofFin n x) = .ofNat (hi-lo+1) (x.val >>> lo) := rfl + +@[simp] +protected theorem extractLsb_ofNat (x n : Nat) (hi lo : Nat) : + extractLsb hi lo x#n = .ofNat (hi - lo + 1) ((x % 2^n) >>> lo) := by + apply eq_of_getLsb_eq + intro ⟨i, _lt⟩ + simp [BitVec.ofNat] + +@[simp] theorem extractLsb'_toNat (s m : Nat) (x : BitVec n) : + (extractLsb' s m x).toNat = (x.toNat >>> s) % 2^m := rfl + +@[simp] theorem extractLsb_toNat (hi lo : Nat) (x : BitVec n) : + (extractLsb hi lo x).toNat = (x.toNat >>> lo) % 2^(hi-lo+1) := rfl + +@[simp] theorem getLsb_extract (hi lo : Nat) (x : BitVec n) (i : Nat) : + getLsb (extractLsb hi lo x) i = (i ≤ (hi-lo) && getLsb x (lo+i)) := by + unfold getLsb + simp [Nat.lt_succ] + +/-! ### allOnes -/ + +private theorem allOnes_def : + allOnes v = .ofFin (⟨0, Nat.two_pow_pos v⟩ - ⟨1 % 2^v, Nat.mod_lt _ (Nat.two_pow_pos v)⟩) := by + rfl + +@[simp] theorem toNat_allOnes : (allOnes v).toNat = 2^v - 1 := by + simp only [allOnes_def, toNat_ofFin, Fin.coe_sub, Nat.zero_add] + by_cases h : v = 0 + · subst h + rfl + · rw [Nat.mod_eq_of_lt (Nat.one_lt_two_pow h), Nat.mod_eq_of_lt] + exact Nat.pred_lt_self (Nat.two_pow_pos v) + +@[simp] theorem getLsb_allOnes : (allOnes v).getLsb i = decide (i < v) := by + simp only [allOnes_def, getLsb_ofFin, Fin.coe_sub, Nat.zero_add, Nat.testBit_mod_two_pow] + if h : i < v then + simp only [h, decide_True, Bool.true_and] + match i, v, h with + | i, (v + 1), h => + rw [Nat.mod_eq_of_lt (by simp), Nat.testBit_two_pow_sub_one] + simp [h] + else + simp [h] + +/-! ### or -/ + +@[simp] theorem toNat_or {x y : BitVec v} : + BitVec.toNat (x ||| y) = BitVec.toNat x ||| BitVec.toNat y := rfl + +@[simp] theorem toFin_or {x y : BitVec v} : + BitVec.toFin (x ||| y) = BitVec.toFin x ||| BitVec.toFin y := by + simp only [HOr.hOr, OrOp.or, BitVec.or, Fin.lor, val_toFin, Fin.mk.injEq] + exact (Nat.mod_eq_of_lt <| Nat.or_lt_two_pow x.isLt y.isLt).symm + + +@[simp] theorem getLsb_or {x y : BitVec v} : (x ||| y).getLsb i = (x.getLsb i || y.getLsb i) := by + rw [← testBit_toNat, getLsb, getLsb] + simp + +/-! ### and -/ + +@[simp] theorem toNat_and {x y : BitVec v} : + BitVec.toNat (x &&& y) = BitVec.toNat x &&& BitVec.toNat y := rfl + +@[simp] theorem toFin_and {x y : BitVec v} : + BitVec.toFin (x &&& y) = BitVec.toFin x &&& BitVec.toFin y := by + simp only [HAnd.hAnd, AndOp.and, BitVec.and, Fin.land, val_toFin, Fin.mk.injEq] + exact (Nat.mod_eq_of_lt <| Nat.and_lt_two_pow _ y.isLt).symm + +@[simp] theorem getLsb_and {x y : BitVec v} : (x &&& y).getLsb i = (x.getLsb i && y.getLsb i) := by + rw [← testBit_toNat, getLsb, getLsb] + simp + +/-! ### xor -/ + +@[simp] theorem toNat_xor {x y : BitVec v} : + BitVec.toNat (x ^^^ y) = BitVec.toNat x ^^^ BitVec.toNat y := rfl + +@[simp] theorem toFin_xor {x y : BitVec v} : + BitVec.toFin (x ^^^ y) = BitVec.toFin x ^^^ BitVec.toFin y := by + simp only [HXor.hXor, Xor.xor, BitVec.xor, Fin.xor, val_toFin, Fin.mk.injEq] + exact (Nat.mod_eq_of_lt <| Nat.xor_lt_two_pow x.isLt y.isLt).symm + +@[simp] theorem getLsb_xor {x y : BitVec v} : + (x ^^^ y).getLsb i = (xor (x.getLsb i) (y.getLsb i)) := by + rw [← testBit_toNat, getLsb, getLsb] + simp + +/-! ### not -/ + +theorem not_def {x : BitVec v} : ~~~x = allOnes v ^^^ x := rfl + +@[simp] theorem toNat_not {x : BitVec v} : (~~~x).toNat = 2^v - 1 - x.toNat := by + rw [Nat.sub_sub, Nat.add_comm, not_def, toNat_xor] + apply Nat.eq_of_testBit_eq + intro i + simp only [toNat_allOnes, Nat.testBit_xor, Nat.testBit_two_pow_sub_one] + match h : BitVec.toNat x with + | 0 => simp + | y+1 => + rw [Nat.succ_eq_add_one] at h + rw [← h] + rw [Nat.testBit_two_pow_sub_succ (toNat_lt _)] + · cases w : decide (i < v) + · simp at w + simp [w] + rw [Nat.testBit_lt_two_pow] + calc BitVec.toNat x < 2 ^ v := toNat_lt _ + _ ≤ 2 ^ i := Nat.pow_le_pow_of_le_right Nat.zero_lt_two w + · simp + +@[simp] theorem getLsb_not {x : BitVec v} : (~~~x).getLsb i = (decide (i < v) && ! x.getLsb i) := by + by_cases h' : i < v <;> simp_all [not_def] + +/-! ### shiftLeft -/ + +@[simp] theorem toNat_shiftLeft {x : BitVec v} : + BitVec.toNat (x <<< n) = BitVec.toNat x <<< n % 2^v := + BitVec.toNat_ofNat _ _ + +@[simp] theorem getLsb_shiftLeft (x : BitVec m) (n) : + getLsb (x <<< n) i = (decide (i < m) && !decide (i < n) && getLsb x (i - n)) := by + rw [← testBit_toNat, getLsb] + simp only [toNat_shiftLeft, Nat.testBit_mod_two_pow, Nat.testBit_shiftLeft, ge_iff_le] + -- This step could be a case bashing tactic. + cases h₁ : decide (i < m) <;> cases h₂ : decide (n ≤ i) <;> cases h₃ : decide (i < n) + all_goals { simp_all <;> omega } + +theorem shiftLeftZeroExtend_eq {x : BitVec w} : + shiftLeftZeroExtend x n = zeroExtend (w+n) x <<< n := by + apply eq_of_toNat_eq + rw [shiftLeftZeroExtend, zeroExtend] + split + · simp + rw [Nat.mod_eq_of_lt] + rw [Nat.shiftLeft_eq, Nat.pow_add] + exact Nat.mul_lt_mul_of_pos_right (BitVec.toNat_lt x) (Nat.two_pow_pos _) + · omega + +@[simp] theorem getLsb_shiftLeftZeroExtend (x : BitVec m) (n : Nat) : + getLsb (shiftLeftZeroExtend x n) i = ((! decide (i < n)) && getLsb x (i - n)) := by + rw [shiftLeftZeroExtend_eq] + simp only [getLsb_shiftLeft, getLsb_zeroExtend] + cases h₁ : decide (i < n) <;> cases h₂ : decide (i - n < m + n) <;> cases h₃ : decide (i < m + n) + <;> simp_all + <;> (rw [getLsb_ge]; omega) + +/-! ### ushiftRight -/ + +@[simp] theorem toNat_ushiftRight (x : BitVec n) (i : Nat) : + (x >>> i).toNat = x.toNat >>> i := rfl + +@[simp] theorem getLsb_ushiftRight (x : BitVec n) (i j : Nat) : + getLsb (x >>> i) j = getLsb x (i+j) := by + unfold getLsb ; simp + +/-! ### append -/ + +theorem append_def (x : BitVec v) (y : BitVec w) : + x ++ y = (shiftLeftZeroExtend x w ||| zeroExtend' (Nat.le_add_left w v) y) := rfl + +@[simp] theorem toNat_append (x : BitVec m) (y : BitVec n) : + (x ++ y).toNat = x.toNat <<< n ||| y.toNat := + rfl + +@[simp] theorem getLsb_append {v : BitVec n} {w : BitVec m} : + getLsb (v ++ w) i = bif i < m then getLsb w i else getLsb v (i - m) := by + simp [append_def] + by_cases h : i < m + · simp [h] + · simp [h]; simp_all + +/-! ### rev -/ + +theorem getLsb_rev (x : BitVec w) (i : Fin w) : + x.getLsb i.rev = x.getMsb i := by + simp [getLsb, getMsb] + congr 1 + omega + +theorem getMsb_rev (x : BitVec w) (i : Fin w) : + x.getMsb i.rev = x.getLsb i := by + simp only [← getLsb_rev] + simp only [Fin.rev] + congr + omega + +/-! ### cons -/ + +@[simp] theorem toNat_cons (b : Bool) (x : BitVec w) : + (cons b x).toNat = (b.toNat <<< w) ||| x.toNat := by + let ⟨x, _⟩ := x + simp [cons, toNat_append, toNat_ofBool] + +@[simp] theorem getLsb_cons (b : Bool) {n} (x : BitVec n) (i : Nat) : + getLsb (cons b x) i = if i = n then b else getLsb x i := by + simp only [getLsb, toNat_cons, Nat.testBit_or] + rw [Nat.testBit_shiftLeft] + rcases Nat.lt_trichotomy i n with i_lt_n | i_eq_n | n_lt_i + · have p1 : ¬(n ≤ i) := by omega + have p2 : i ≠ n := by omega + simp [p1, p2] + · simp [i_eq_n, testBit_toNat] + cases b <;> trivial + · have p1 : i ≠ n := by omega + have p2 : i - n ≠ 0 := by omega + simp [p1, p2, Nat.testBit_bool_to_nat] + theorem truncate_succ (x : BitVec w) : truncate (i+1) x = cons (getLsb x i) (truncate i x) := by apply eq_of_getLsb_eq @@ -173,11 +409,68 @@ theorem toNat_xor {x y: BitVec w} : (x ^^^ y).toNat = x.toNat ^^^ y.toNat:= rfl /-! ### add -/ -theorem add_comm (x y : BitVec w) : x + y = y + x := by - apply BitVec.eq_of_toNat_eq; simp [Nat.add_comm] -@[simp] theorem add_zero (x : BitVec n) : x + (0#n) = x := by +theorem add_def {n} (x y : BitVec n) : x + y = .ofNat n (x.toNat + y.toNat) := rfl + +/-- +Definition of bitvector addition as a nat. +-/ +@[simp] theorem toNat_add (x y : BitVec w) : (x + y).toNat = (x.toNat + y.toNat) % 2^w := rfl +@[simp] theorem toFin_add (x y : BitVec w) : (x + y).toFin = toFin x + toFin y := rfl +@[simp] theorem ofFin_add (x : Fin (2^n)) (y : BitVec n) : + .ofFin x + y = .ofFin (x + y.toFin) := rfl +@[simp] theorem add_ofFin (x : BitVec n) (y : Fin (2^n)) : + x + .ofFin y = .ofFin (x.toFin + y) := rfl +@[simp] theorem ofNat_add_ofNat {n} (x y : Nat) : x#n + y#n = (x + y)#n := by + apply eq_of_toNat_eq ; simp [BitVec.ofNat] + +protected theorem add_assoc (x y z : BitVec n) : x + y + z = x + (y + z) := by + apply eq_of_toNat_eq ; simp [Nat.add_assoc] + +protected theorem add_comm (x y : BitVec n) : x + y = y + x := by + simp [add_def, Nat.add_comm] + +@[simp] protected theorem add_zero (x : BitVec n) : x + 0#n = x := by simp [add_def] + +@[simp] protected theorem zero_add (x : BitVec n) : 0#n + x = x := by simp [add_def] + + +/-! ### sub/neg -/ + +theorem sub_def {n} (x y : BitVec n) : x - y = .ofNat n (x.toNat + (2^n - y.toNat)) := by rfl + +@[simp] theorem toNat_sub {n} (x y : BitVec n) : + (x - y).toNat = ((x.toNat + (2^n - y.toNat)) % 2^n) := rfl +@[simp] theorem toFin_sub (x y : BitVec w) : (x - y).toFin = toFin x - toFin y := rfl + +/-- Replaced 2024-02-06. -/ +@[deprecated] alias sub_toNat := toNat_sub + +@[simp] theorem ofFin_sub (x : Fin (2^n)) (y : BitVec n) : .ofFin x - y = .ofFin (x - y.toFin) := + rfl +@[simp] theorem sub_ofFin (x : BitVec n) (y : Fin (2^n)) : x - .ofFin y = .ofFin (x.toFin - y) := + rfl +@[simp] theorem ofNat_sub_ofNat {n} (x y : Nat) : x#n - y#n = .ofNat n (x + (2^n - y % 2^n)) := by + apply eq_of_toNat_eq ; simp [BitVec.ofNat] + +@[simp] protected theorem sub_zero (x : BitVec n) : x - (0#n) = x := by apply eq_of_toNat_eq ; simp + +@[simp] protected theorem sub_self (x : BitVec n) : x - x = 0#n := by apply eq_of_toNat_eq + simp only [toNat_sub] + rw [Nat.add_sub_of_le] + · simp + · exact Nat.le_of_lt x.isLt + +@[simp] theorem toNat_neg (x : BitVec n) : (- x).toNat = (2^n - x.toNat) % 2^n := by + simp [Neg.neg, BitVec.neg] + +/-- Replaced 2024-02-06. -/ +@[deprecated] alias neg_toNat := toNat_neg + +theorem sub_toAdd {n} (x y : BitVec n) : x - y = x + - y := by + apply eq_of_toNat_eq + simp [toNat_add, isLt, Nat.mod_eq_of_lt] theorem add_right_neg (x : BitVec w) : x + -x = 0 := by @@ -197,3 +490,49 @@ theorem add_sub_assoc (x y z : BitVec w) : x + y - z = x + (y - z) := by apply BitVec.eq_of_toNat_eq; simp [Nat.add_sub_assoc, Nat.add_assoc] theorem zero_sub (x : BitVec w) : 0#w - x = -x := rfl + simp + +@[simp] theorem neg_zero (n:Nat) : -0#n = 0#n := by apply eq_of_toNat_eq ; simp + +theorem add_sub_cancel (x y : BitVec w) : x + y - y = x := by + apply eq_of_toNat_eq + have y_toNat_le := Nat.le_of_lt y.toNat_lt + rw [toNat_sub, toNat_add, Nat.mod_add_mod, Nat.add_assoc, ← Nat.add_sub_assoc y_toNat_le, + Nat.add_sub_cancel_left, Nat.add_mod_right, toNat_mod_cancel] + +/-! ### mul -/ + +theorem mul_def {n} {x y : BitVec n} : x * y = (ofFin <| x.toFin * y.toFin) := by rfl + +theorem toNat_mul {x y : BitVec w} : (x * y).toNat = (x.toNat * y.toNat) % 2 ^ w := rfl +@[simp] theorem toFin_mul {x y : BitVec w} : (x * y).toFin = (x.toFin * y.toFin) := rfl + +/-! ### le and lt -/ + +theorem le_def (x y : BitVec n) : + x ≤ y ↔ x.toNat ≤ y.toNat := Iff.rfl + +@[simp] theorem le_ofFin (x : BitVec n) (y : Fin (2^n)) : + x ≤ BitVec.ofFin y ↔ x.toFin ≤ y := Iff.rfl +@[simp] theorem ofFin_le (x : Fin (2^n)) (y : BitVec n) : + BitVec.ofFin x ≤ y ↔ x ≤ y.toFin := Iff.rfl +@[simp] theorem ofNat_le_ofNat {n} (x y : Nat) : (x#n) ≤ (y#n) ↔ x % 2^n ≤ y % 2^n := by + simp [le_def] + +theorem lt_def (x y : BitVec n) : + x < y ↔ x.toNat < y.toNat := Iff.rfl + +@[simp] theorem lt_ofFin (x : BitVec n) (y : Fin (2^n)) : + x < BitVec.ofFin y ↔ x.toFin < y := Iff.rfl +@[simp] theorem ofFin_lt (x : Fin (2^n)) (y : BitVec n) : + BitVec.ofFin x < y ↔ x < y.toFin := Iff.rfl +@[simp] theorem ofNat_lt_ofNat {n} (x y : Nat) : (x#n) < (y#n) ↔ x % 2^n < y % 2^n := by + simp [lt_def] + +protected theorem lt_of_le_ne (x y : BitVec n) (h1 : x <= y) (h2 : ¬ x = y) : x < y := by + revert h1 h2 + let ⟨x, lt⟩ := x + let ⟨y, lt⟩ := y + simp + exact Nat.lt_of_le_of_ne + diff --git a/Std/Data/Bool.lean b/Std/Data/Bool.lean index 539dc04077..26424373cb 100644 --- a/Std/Data/Bool.lean +++ b/Std/Data/Bool.lean @@ -38,6 +38,14 @@ instance (x y : Bool) : Decidable (x < y) := inferInstanceAs (Decidable (!x && y instance : Max Bool := ⟨or⟩ instance : Min Bool := ⟨and⟩ +theorem false_ne_true : false ≠ true := Bool.noConfusion + +theorem eq_false_or_eq_true : (b : Bool) → b = true ∨ b = false := by decide + +theorem eq_false_iff : {b : Bool} → b = false ↔ b ≠ true := by decide + +theorem ne_false_iff : {b : Bool} → b ≠ false ↔ b = true := by decide + /-! ### and -/ @[simp] theorem not_and_self : ∀ (x : Bool), (!x && x) = false := by decide diff --git a/Std/Data/ByteArray.lean b/Std/Data/ByteArray.lean index 35a9267f82..4652153d5c 100644 --- a/Std/Data/ByteArray.lean +++ b/Std/Data/ByteArray.lean @@ -23,6 +23,8 @@ 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 empty_data : empty.data = #[] := rfl @[simp] theorem size_empty : empty.size = 0 := rfl @@ -109,3 +111,34 @@ theorem get_extract_aux {a : ByteArray} {start stop} (h : i < (a.extract start s @[simp] theorem get_extract {a : ByteArray} {start stop} (h : i < (a.extract start stop).size) : (a.extract start stop)[i] = a[start+i]'(get_extract_aux h) := by simp [getElem_eq_data_getElem] + +/-! ### ofFn -/ + +/--- `ofFn f` with `f : Fin n → UInt8` returns the byte array whose `i`th element is `f i`. --/ +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 size_ofFn (f : Fin n → UInt8) : (ofFn f).size = n := by + simp [size] + +@[simp] theorem get_ofFn (f : Fin n → UInt8) (i : Fin (ofFn f).size) : + (ofFn f).get i = f (i.cast (size_ofFn f)) := by + simp [get, Fin.cast] + +@[simp] theorem getElem_ofFn (f : Fin n → UInt8) (i) (h : i < (ofFn f).size) : + (ofFn f)[i] = f ⟨i, size_ofFn f ▸ h⟩ := get_ofFn .. + +private def ofFnAux (f : Fin n → UInt8) : ByteArray := go 0 (mkEmpty n) where + go (i : Nat) (acc : ByteArray) : ByteArray := + if h : i < n then go (i+1) (acc.push (f ⟨i, h⟩)) else acc +termination_by n - i + +@[csimp] private theorem ofFn_eq_ofFnAux : @ofFn = @ofFnAux := by + funext n f; ext; simp [ofFnAux, Array.ofFn, ofFnAux_data, mkEmpty] +where + ofFnAux_data {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 + termination_by n - i diff --git a/Std/Data/Char.lean b/Std/Data/Char.lean index 3c7a790840..ed40840832 100644 --- a/Std/Data/Char.lean +++ b/Std/Data/Char.lean @@ -3,9 +3,12 @@ Copyright (c) 2022 Jannis Limperg. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jannis Limperg -/ - +import Std.Tactic.Ext.Attr import Std.Tactic.RCases +@[ext] theorem Char.ext : {a b : Char} → a.val = b.val → a = b + | ⟨_,_⟩, ⟨_,_⟩, rfl => rfl + namespace String private theorem csize_eq (c) : diff --git a/Std/Data/Fin/Basic.lean b/Std/Data/Fin/Basic.lean index f7f3f112af..109f92892f 100644 --- a/Std/Data/Fin/Basic.lean +++ b/Std/Data/Fin/Basic.lean @@ -54,7 +54,7 @@ def clamp (n m : Nat) : Fin (m + 1) := ⟨min n m, Nat.lt_succ_of_le (Nat.min_le /-- Inner loop for `Fin.foldl`. `Fin.foldl.loop n f x i = f (f (f x i) ...) (n-1)` -/ loop (x : α) (i : Nat) : α := if h : i < n then loop (f x ⟨i, h⟩) (i+1) else x -termination_by loop i => n - i + termination_by n - i /-- Folds over `Fin n` from the right: `foldr 3 f x = f 0 (f 1 (f 2 x))`. -/ @[inline] def foldr (n) (f : Fin n → α → α) (init : α) : α := loop ⟨n, Nat.le_refl n⟩ init where diff --git a/Std/Data/Fin/Iterate.lean b/Std/Data/Fin/Iterate.lean index 6c6076bddd..c21fdffd89 100644 --- a/Std/Data/Fin/Iterate.lean +++ b/Std/Data/Fin/Iterate.lean @@ -21,7 +21,7 @@ def hIterateFrom (P : Nat → Sort _) {n} (f : ∀(i : Fin n), P i.val → P (i. else have p : i = n := (or_iff_left g).mp (Nat.eq_or_lt_of_le ubnd) cast (congrArg P p) a - termination_by hIterateFrom i _ _ => n - i + termination_by n - i /-- `hIterate` is a heterogenous iterative operation that applies a diff --git a/Std/Data/Fin/Lemmas.lean b/Std/Data/Fin/Lemmas.lean index 64db605074..d1651a585b 100644 --- a/Std/Data/Fin/Lemmas.lean +++ b/Std/Data/Fin/Lemmas.lean @@ -8,6 +8,8 @@ import Std.Data.Nat.Lemmas import Std.Tactic.Ext import Std.Tactic.Simpa import Std.Tactic.NormCast.Lemmas +import Std.Tactic.Omega +import Std.Tactic.SimpTrace namespace Fin @@ -56,7 +58,10 @@ theorem eq_mk_iff_val_eq {a : Fin n} {k : Nat} {hk : k < n} : theorem mk_val (i : Fin n) : (⟨i, i.isLt⟩ : Fin n) = i := Fin.eta .. -@[simp] theorem ofNat'_zero_val : (Fin.ofNat' 0 h).val = 0 := Nat.zero_mod _ +@[simp] theorem val_ofNat' (a : Nat) (is_pos : n > 0) : + (Fin.ofNat' a is_pos).val = a % n := rfl + +@[deprecated ofNat'_zero_val] theorem ofNat'_zero_val : (Fin.ofNat' 0 h).val = 0 := Nat.zero_mod _ @[simp] theorem mod_val (a b : Fin n) : (a % b).val = a.val % b.val := rfl @@ -81,6 +86,8 @@ theorem le_def {a b : Fin n} : a ≤ b ↔ a.1 ≤ b.1 := .rfl theorem lt_def {a b : Fin n} : a < b ↔ a.1 < b.1 := .rfl +theorem lt_iff_val_lt_val {a b : Fin n} : a < b ↔ a.val < b.val := Iff.rfl + @[simp] protected theorem not_le {a b : Fin n} : ¬ a ≤ b ↔ b < a := Nat.not_le @[simp] protected theorem not_lt {a b : Fin n} : ¬ a < b ↔ b ≤ a := Nat.not_lt @@ -89,6 +96,8 @@ protected theorem ne_of_lt {a b : Fin n} (h : a < b) : a ≠ b := Fin.ne_of_val_ protected theorem ne_of_gt {a b : Fin n} (h : a < b) : b ≠ a := Fin.ne_of_val_ne (Nat.ne_of_gt h) +protected theorem le_of_lt {a b : Fin n} (h : a < b) : a ≤ b := Nat.le_of_lt h + theorem is_le (i : Fin (n + 1)) : i ≤ n := Nat.le_of_lt_succ i.is_lt @[simp] theorem is_le' {a : Fin n} : a ≤ n := Nat.le_of_lt a.is_lt @@ -673,7 +682,7 @@ and `cast` defines the inductive step using `motive i.succ`, inducting downwards else let j : Fin n := ⟨i, Nat.lt_of_le_of_ne (Nat.le_of_lt_succ i.2) fun h => hi (Fin.ext h)⟩ cast _ (reverseInduction last cast j.succ) -termination_by _ => n + 1 - i +termination_by n + 1 - i decreasing_by decreasing_with -- FIXME: we put the proof down here to avoid getting a dummy `have` in the definition exact Nat.add_sub_add_right .. ▸ Nat.sub_lt_sub_left i.2 (Nat.lt_succ_self i) @@ -723,6 +732,57 @@ theorem addCases_right {m n : Nat} {motive : Fin (m + n) → Sort _} {left right @[simp] theorem coe_clamp (n m : Nat) : (clamp n m : Nat) = min n m := rfl +/-! ### add -/ + +@[simp] theorem ofNat'_add (x : Nat) (lt : 0 < n) (y : Fin n) : + Fin.ofNat' x lt + y = Fin.ofNat' (x + y.val) lt := by + apply Fin.eq_of_val_eq + simp [Fin.ofNat', Fin.add_def] + +@[simp] theorem add_ofNat' (x : Fin n) (y : Nat) (lt : 0 < n) : + x + Fin.ofNat' y lt = Fin.ofNat' (x.val + y) lt := by + apply Fin.eq_of_val_eq + simp [Fin.ofNat', Fin.add_def] + +/-! ### sub -/ + +protected theorem coe_sub (a b : Fin n) : ((a - b : Fin n) : Nat) = (a + (n - b)) % n := by + cases a; cases b; rfl + +@[simp] theorem ofNat'_sub (x : Nat) (lt : 0 < n) (y : Fin n) : + Fin.ofNat' x lt - y = Fin.ofNat' (x + (n - y.val)) lt := by + apply Fin.eq_of_val_eq + simp [Fin.ofNat', Fin.sub_def] + +@[simp] theorem sub_ofNat' (x : Fin n) (y : Nat) (lt : 0 < n) : + x - Fin.ofNat' y lt = Fin.ofNat' (x.val + (n - y % n)) lt := by + apply Fin.eq_of_val_eq + simp [Fin.ofNat', Fin.sub_def] + +private theorem _root_.Nat.mod_eq_sub_of_lt_two_mul {x n} (h₁ : n ≤ x) (h₂ : x < 2 * n) : + x % n = x - n := by + rw [Nat.mod_eq, if_pos (by omega), Nat.mod_eq_of_lt (by omega)] + +theorem coe_sub_iff_le {a b : Fin n} : (↑(a - b) : Nat) = a - b ↔ b ≤ a := by + rw [sub_def, le_def] + dsimp only + if h : n ≤ a + (n - b) then + rw [Nat.mod_eq_sub_of_lt_two_mul h] + all_goals omega + else + rw [Nat.mod_eq_of_lt] + all_goals omega + +theorem coe_sub_iff_lt {a b : Fin n} : (↑(a - b) : Nat) = n + a - b ↔ a < b := by + rw [sub_def, lt_def] + dsimp only + if h : n ≤ a + (n - b) then + rw [Nat.mod_eq_sub_of_lt_two_mul h] + all_goals omega + else + rw [Nat.mod_eq_of_lt] + all_goals omega + /-! ### mul -/ theorem val_mul {n : Nat} : ∀ a b : Fin n, (a * b).val = a.val * b.val % n diff --git a/Std/Data/HashMap/Basic.lean b/Std/Data/HashMap/Basic.lean index 168c47c299..9635d17ce3 100644 --- a/Std/Data/HashMap/Basic.lean +++ b/Std/Data/HashMap/Basic.lean @@ -151,7 +151,7 @@ where let target := es.foldl reinsertAux target go (i+1) source target else target -termination_by _ i source _ => source.size - i + termination_by source.size - i /-- Inserts key-value pair `a, b` into the map. diff --git a/Std/Data/HashMap/WF.lean b/Std/Data/HashMap/WF.lean index d5cc2a7058..82f1f9758e 100644 --- a/Std/Data/HashMap/WF.lean +++ b/Std/Data/HashMap/WF.lean @@ -108,7 +108,7 @@ where simp 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 go i source _ _ => source.size - i + termination_by source.size - i theorem expand_WF.foldl [BEq α] [Hashable α] (rank : α → Nat) {l : List (α × β)} {i : Nat} (hl₁ : ∀ [PartialEquivBEq α] [LawfulHashable α], l.Pairwise fun a b => ¬(a.1 == b.1)) @@ -170,7 +170,7 @@ where refine ⟨Nat.le_of_lt this, fun _ h h' => Nat.ne_of_lt this ?_⟩ exact LawfulHashable.hash_eq h' ▸ hs₂ _ H _ h · exact ht.1 -termination_by go i source _ _ _ _ => source.size - i + termination_by source.size - i theorem insert_size [BEq α] [Hashable α] {m : Imp α β} {k v} (h : m.size = m.buckets.size) : diff --git a/Std/Data/Int.lean b/Std/Data/Int.lean index 7bfed88cee..ebd45712cb 100644 --- a/Std/Data/Int.lean +++ b/Std/Data/Int.lean @@ -1,4 +1,7 @@ import Std.Data.Int.Basic import Std.Data.Int.DivMod import Std.Data.Int.Gcd -import Std.Data.Int.Lemmas +import Std.Data.Int.Init.DivMod +import Std.Data.Int.Init.Lemmas +import Std.Data.Int.Init.Order +import Std.Data.Int.Order diff --git a/Std/Data/Int/DivMod.lean b/Std/Data/Int/DivMod.lean index 1c4a45fdd6..8e2ae5766c 100644 --- a/Std/Data/Int/DivMod.lean +++ b/Std/Data/Int/DivMod.lean @@ -3,7 +3,8 @@ 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 Std.Data.Int.Lemmas +import Std.Data.Int.Order +import Std.Data.Int.Init.DivMod import Std.Tactic.Change /-! @@ -23,8 +24,6 @@ theorem ofNat_fdiv : ∀ m n : Nat, ↑(m / n) = fdiv ↑m ↑n | 0, _ => by simp [fdiv] | succ _, _ => rfl -@[simp, norm_cast] theorem ofNat_ediv (m n : Nat) : (↑(m / n) : Int) = ↑m / ↑n := rfl - theorem negSucc_ediv (m : Nat) {b : Int} (H : 0 < b) : -[m+1] / b = -(div m b + 1) := match b, eq_succ_of_zero_lt H with | _, ⟨_, rfl⟩ => rfl @@ -33,20 +32,12 @@ theorem negSucc_ediv (m : Nat) {b : Int} (H : 0 < b) : -[m+1] / b = -(div m b + | ofNat _ => show ofNat _ = _ by simp | -[_+1] => show -ofNat _ = _ by simp -@[simp] theorem zero_ediv : ∀ b : Int, 0 / b = 0 - | ofNat _ => show ofNat _ = _ by simp - | -[_+1] => show -ofNat _ = _ by simp - @[simp] theorem zero_fdiv (b : Int) : fdiv 0 b = 0 := by cases b <;> rfl @[simp] protected theorem div_zero : ∀ a : Int, div a 0 = 0 | ofNat _ => show ofNat _ = _ by simp | -[_+1] => rfl -@[simp] protected theorem ediv_zero : ∀ a : Int, a / 0 = 0 - | ofNat _ => show ofNat _ = _ by simp - | -[_+1] => rfl - @[simp] protected theorem fdiv_zero : ∀ a : Int, fdiv a 0 = 0 | 0 => rfl | succ _ => rfl @@ -68,13 +59,6 @@ theorem fdiv_eq_div {a b : Int} (Ha : 0 ≤ a) (Hb : 0 ≤ b) : fdiv a b = div a | ofNat m, -[n+1] | -[m+1], succ n => (Int.neg_neg _).symm | ofNat m, succ n | -[m+1], 0 | -[m+1], -[n+1] => rfl -@[simp] protected theorem ediv_neg : ∀ a b : Int, a / (-b) = -(a / b) - | ofNat m, 0 => show ofNat (m / 0) = -↑(m / 0) by rw [Nat.div_zero]; rfl - | ofNat m, -[n+1] => (Int.neg_neg _).symm - | ofNat m, succ n | -[m+1], 0 | -[m+1], succ n | -[m+1], -[n+1] => rfl - -protected theorem div_def (a b : Int) : a / b = Int.ediv a b := rfl - -- Lean 4 core provides an instance for `Div Int` using `Int.div`. -- Even though we provide a higher priority instance in `Std.Data.Int.Basic`, -- we provide a `simp` lemma here to unfold usages of that instance back to `Int.div`. @@ -138,54 +122,10 @@ theorem ediv_eq_zero_of_lt {a b : Int} (H1 : 0 ≤ a) (H2 : a < b) : a / b = 0 : match a, b, eq_ofNat_of_zero_le H1, eq_succ_of_zero_lt (Int.lt_of_le_of_lt H1 H2) with | _, _, ⟨_, rfl⟩, ⟨_, rfl⟩ => congrArg Nat.cast <| Nat.div_eq_of_lt <| ofNat_lt.1 H2 -theorem add_mul_ediv_right (a b : Int) {c : Int} (H : c ≠ 0) : (a + b * c) / c = a / c + b := - suffices ∀ {{a b c : Int}}, 0 < c → (a + b * c).ediv c = a.ediv c + b from - match Int.lt_trichotomy c 0 with - | Or.inl hlt => by - rw [← Int.neg_inj, ← Int.ediv_neg, Int.neg_add, ← Int.ediv_neg, ← Int.neg_mul_neg] - exact this (Int.neg_pos_of_neg hlt) - | Or.inr (Or.inl HEq) => absurd HEq H - | Or.inr (Or.inr hgt) => this hgt - suffices ∀ {k n : Nat} {a : Int}, (a + n * k.succ).ediv k.succ = a.ediv k.succ + n from - fun a b c H => match c, eq_succ_of_zero_lt H, b with - | _, ⟨_, rfl⟩, ofNat _ => this - | _, ⟨k, rfl⟩, -[n+1] => show (a - n.succ * k.succ).ediv k.succ = a.ediv k.succ - n.succ by - rw [← Int.add_sub_cancel (ediv ..), ← this, Int.sub_add_cancel] - fun {k n} => @fun - | ofNat m => congrArg ofNat <| Nat.add_mul_div_right _ _ k.succ_pos - | -[m+1] => by - show ((n * k.succ : Nat) - m.succ : Int).ediv k.succ = n - (m / k.succ + 1 : Nat) - if h : m < n * k.succ then - rw [← Int.ofNat_sub h, ← Int.ofNat_sub ((Nat.div_lt_iff_lt_mul k.succ_pos).2 h)] - apply congrArg ofNat - rw [Nat.mul_comm, Nat.mul_sub_div]; rwa [Nat.mul_comm] - else - have h := Nat.not_lt.1 h - have H {a b : Nat} (h : a ≤ b) : (a : Int) + -((b : Int) + 1) = -[b - a +1] := by - rw [negSucc_eq, Int.ofNat_sub h] - simp only [Int.sub_eq_add_neg, Int.neg_add, Int.neg_neg, Int.add_left_comm, Int.add_assoc] - show ediv (↑(n * succ k) + -((m : Int) + 1)) (succ k) = n + -(↑(m / succ k) + 1 : Int) - rw [H h, H ((Nat.le_div_iff_mul_le k.succ_pos).2 h)] - apply congrArg negSucc - rw [Nat.mul_comm, Nat.sub_mul_div]; rwa [Nat.mul_comm] - theorem add_mul_ediv_left (a : Int) {b : Int} (c : Int) (H : b ≠ 0) : (a + b * c) / b = a / b + c := Int.mul_comm .. ▸ Int.add_mul_ediv_right _ _ H -theorem add_ediv_of_dvd_right {a b c : Int} (H : c ∣ b) : (a + b) / c = a / c + b / c := - if h : c = 0 then by simp [h] else by - let ⟨k, hk⟩ := H - rw [hk, Int.mul_comm c k, Int.add_mul_ediv_right _ _ h, - ← Int.zero_add (k * c), Int.add_mul_ediv_right _ _ h, Int.zero_ediv, Int.zero_add] - -theorem add_ediv_of_dvd_left {a b c : Int} (H : c ∣ a) : (a + b) / c = a / c + b / c := by - rw [Int.add_comm, Int.add_ediv_of_dvd_right H, Int.add_comm] - -@[simp] theorem mul_ediv_cancel (a : Int) {b : Int} (H : b ≠ 0) : (a * b) / b = a := by - have := Int.add_mul_ediv_right 0 a H - rwa [Int.zero_add, Int.zero_ediv, Int.zero_add] at this - @[simp] theorem mul_fdiv_cancel (a : Int) {b : Int} (H : b ≠ 0) : fdiv (a * b) b = a := if b0 : 0 ≤ b then by rw [fdiv_eq_ediv _ b0, mul_ediv_cancel _ H] @@ -216,9 +156,6 @@ theorem add_ediv_of_dvd_left {a b c : Int} (H : c ∣ a) : (a + b) / c = a / c + @[simp] theorem mul_fdiv_cancel_left (b : Int) (H : a ≠ 0) : fdiv (a * b) a = b := Int.mul_comm .. ▸ Int.mul_fdiv_cancel _ H -@[simp] theorem mul_ediv_cancel_left (b : Int) (H : a ≠ 0) : (a * b) / a = b := - Int.mul_comm .. ▸ Int.mul_ediv_cancel _ H - @[simp] protected theorem div_self {a : Int} (H : a ≠ 0) : a.div a = 1 := by have := Int.mul_div_cancel 1 H; rwa [Int.one_mul] at this @@ -228,25 +165,10 @@ theorem add_ediv_of_dvd_left {a b c : Int} (H : c ∣ a) : (a + b) / c = a / c + @[simp] protected theorem ediv_self {a : Int} (H : a ≠ 0) : a / a = 1 := by have := Int.mul_ediv_cancel 1 H; rwa [Int.one_mul] at this -theorem div_nonneg_iff_of_pos {a b : Int} (h : 0 < b) : a / b ≥ 0 ↔ a ≥ 0 := by - rw [Int.div_def] - match b, h with - | Int.ofNat (b+1), _ => - rcases a with ⟨a⟩ <;> simp [Int.ediv] - exact decide_eq_decide.mp rfl - /-! ### mod -/ -theorem mod_def' (m n : Int) : m % n = emod m n := rfl - -theorem ofNat_mod (m n : Nat) : (↑(m % n) : Int) = mod m n := rfl - -theorem ofNat_mod_ofNat (m n : Nat) : (m % n : Int) = ↑(m % n) := rfl - theorem ofNat_fmod (m n : Nat) : ↑(m % n) = fmod m n := by cases m <;> simp [fmod] -@[simp, norm_cast] theorem ofNat_emod (m n : Nat) : (↑(m % n) : Int) = m % n := rfl - theorem negSucc_emod (m : Nat) {b : Int} (bpos : 0 < b) : -[m+1] % b = b - 1 - m % b := by rw [Int.sub_sub, Int.add_comm] match b, eq_succ_of_zero_lt bpos with @@ -256,8 +178,6 @@ theorem negSucc_emod (m : Nat) {b : Int} (bpos : 0 < b) : -[m+1] % b = b - 1 - m @[simp] theorem zero_fmod (b : Int) : fmod 0 b = 0 := by cases b <;> rfl -@[simp] theorem zero_emod (b : Int) : 0 % b = 0 := by simp [mod_def', emod] - @[simp] theorem mod_zero : ∀ a : Int, mod a 0 = a | ofNat _ => congrArg ofNat <| Nat.mod_zero _ | -[_+1] => rfl @@ -267,10 +187,6 @@ theorem negSucc_emod (m : Nat) {b : Int} (bpos : 0 < b) : -[m+1] % b = b - 1 - m | succ _ => congrArg ofNat <| Nat.mod_zero _ | -[_+1] => congrArg negSucc <| Nat.mod_zero _ -@[simp] theorem emod_zero : ∀ a : Int, a % 0 = a - | ofNat _ => congrArg ofNat <| Nat.mod_zero _ - | -[_+1] => congrArg negSucc <| Nat.mod_zero _ - theorem mod_add_div : ∀ a b : Int, mod a b + b * (a.div b) = a | ofNat _, ofNat _ => congrArg ofNat (Nat.mod_add_div ..) | ofNat m, -[n+1] => by @@ -303,38 +219,18 @@ theorem fmod_add_fdiv : ∀ a b : Int, a.fmod b + b * a.fdiv b = a show -(↑(succ m % succ n) : Int) + -↑(succ n * (succ m / succ n)) = -↑(succ m) rw [← Int.neg_add]; exact congrArg (-ofNat ·) <| Nat.mod_add_div .. -theorem emod_add_ediv : ∀ a b : Int, a % b + b * (a / b) = a - | ofNat _, ofNat _ => congrArg ofNat <| Nat.mod_add_div .. - | ofNat m, -[n+1] => by - show (m % succ n + -↑(succ n) * -↑(m / succ n) : Int) = m - rw [Int.neg_mul_neg]; exact congrArg ofNat <| Nat.mod_add_div .. - | -[_+1], 0 => by rw [emod_zero]; rfl - | -[m+1], succ n => aux m n.succ - | -[m+1], -[n+1] => aux m n.succ -where - aux (m n : Nat) : n - (m % n + 1) - (n * (m / n) + n) = -[m+1] := by - rw [← ofNat_emod, ← ofNat_ediv, ← Int.sub_sub, negSucc_eq, Int.sub_sub n, - ← Int.neg_neg (_-_), Int.neg_sub, Int.sub_sub_self, Int.add_right_comm] - exact congrArg (fun x => -(ofNat x + 1)) (Nat.mod_add_div ..) - theorem div_add_mod (a b : Int) : b * a.div b + mod a b = a := (Int.add_comm ..).trans (mod_add_div ..) theorem fdiv_add_fmod (a b : Int) : b * a.fdiv b + a.fmod b = a := (Int.add_comm ..).trans (fmod_add_fdiv ..) -theorem ediv_add_emod (a b : Int) : b * (a / b) + a % b = a := - (Int.add_comm ..).trans (emod_add_ediv ..) - theorem mod_def (a b : Int) : mod a b = a - b * a.div b := by rw [← Int.add_sub_cancel (mod a b), mod_add_div] theorem fmod_def (a b : Int) : a.fmod b = a - b * a.fdiv b := by rw [← Int.add_sub_cancel (a.fmod b), fmod_add_fdiv] -theorem emod_def (a b : Int) : a % b = a - b * (a / b) := by - rw [← Int.add_sub_cancel (a % b), emod_add_ediv] - theorem fmod_eq_emod (a : Int) {b : Int} (hb : 0 ≤ b) : fmod a b = a % b := by simp [fmod_def, emod_def, fdiv_eq_ediv _ hb] @@ -376,10 +272,6 @@ theorem fmod_eq_of_lt {a b : Int} (H1 : 0 ≤ a) (H2 : a < b) : a.fmod b = a := theorem mod_nonneg : ∀ {a : Int} (b : Int), 0 ≤ a → 0 ≤ mod a b | ofNat _, -[_+1], _ | ofNat _, ofNat _, _ => ofNat_nonneg _ -theorem emod_nonneg : ∀ (a : Int) {b : Int}, b ≠ 0 → 0 ≤ a % b - | ofNat _, _, _ => ofNat_zero_le _ - | -[_+1], _, H => Int.sub_nonneg_of_le <| ofNat_le.2 <| Nat.mod_lt _ (natAbs_pos.2 H) - theorem fmod_nonneg {a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a.fmod b := fmod_eq_mod ha hb ▸ mod_nonneg _ ha @@ -392,24 +284,9 @@ theorem mod_lt_of_pos (a : Int) {b : Int} (H : 0 < b) : mod a b < b := | -[_+1], _, ⟨n, rfl⟩ => Int.lt_of_le_of_lt (Int.neg_nonpos_of_nonneg <| Int.ofNat_nonneg _) (ofNat_pos.2 n.succ_pos) -theorem emod_lt_of_pos (a : Int) {b : Int} (H : 0 < b) : a % b < b := - match a, b, eq_succ_of_zero_lt H with - | ofNat _, _, ⟨_, rfl⟩ => ofNat_lt.2 (Nat.mod_lt _ (Nat.succ_pos _)) - | -[_+1], _, ⟨_, rfl⟩ => Int.sub_lt_self _ (ofNat_lt.2 <| Nat.succ_pos _) - theorem fmod_lt_of_pos (a : Int) {b : Int} (H : 0 < b) : a.fmod b < b := fmod_eq_emod _ (Int.le_of_lt H) ▸ emod_lt_of_pos a H -theorem mul_ediv_self_le {x k : Int} (h : k ≠ 0) : k * (x / k) ≤ x := - calc k * (x / k) - _ ≤ k * (x / k) + x % k := Int.le_add_of_nonneg_right (emod_nonneg x h) - _ = x := ediv_add_emod _ _ - -theorem lt_mul_ediv_self_add {x k : Int} (h : 0 < k) : x < k * (x / k) + k := - calc x - _ = k * (x / k) + x % k := (ediv_add_emod _ _).symm - _ < k * (x / k) + k := Int.add_lt_add_left (emod_lt_of_pos x h) _ - theorem emod_two_eq (x : Int) : x % 2 = 0 ∨ x % 2 = 1 := by have h₁ : 0 ≤ x % 2 := Int.emod_nonneg x (by decide) have h₂ : x % 2 < 2 := Int.emod_lt_of_pos x (by decide) @@ -423,55 +300,13 @@ theorem mod_add_div' (m k : Int) : mod m k + m.div k * k = m := by theorem div_add_mod' (m k : Int) : m.div k * k + mod m k = m := by rw [Int.mul_comm]; apply div_add_mod -theorem emod_add_ediv' (m k : Int) : m % k + m / k * k = m := by - rw [Int.mul_comm]; apply emod_add_ediv - theorem ediv_add_emod' (m k : Int) : m / k * k + m % k = m := by rw [Int.mul_comm]; apply ediv_add_emod -@[simp] theorem add_mul_emod_self {a b c : Int} : (a + b * c) % c = a % c := - if cz : c = 0 then by - rw [cz, Int.mul_zero, Int.add_zero] - else by - rw [Int.emod_def, Int.emod_def, Int.add_mul_ediv_right _ _ cz, Int.add_comm _ b, - Int.mul_add, Int.mul_comm, ← Int.sub_sub, Int.add_sub_cancel] - -@[simp] theorem add_mul_emod_self_left (a b c : Int) : (a + b * c) % b = a % b := by - rw [Int.mul_comm, Int.add_mul_emod_self] - -@[simp] theorem add_emod_self {a b : Int} : (a + b) % b = a % b := by - have := add_mul_emod_self_left a b 1; rwa [Int.mul_one] at this - -@[simp] theorem add_emod_self_left {a b : Int} : (a + b) % a = b % a := by - rw [Int.add_comm, Int.add_emod_self] - -theorem neg_emod {a b : Int} : -a % b = (b - a) % b := by - rw [← add_emod_self_left]; rfl - -@[simp] theorem emod_add_emod (m n k : Int) : (m % n + k) % n = (m + k) % n := by - have := (add_mul_emod_self_left (m % n + k) n (m / n)).symm - rwa [Int.add_right_comm, emod_add_ediv] at this - -@[simp] theorem add_emod_emod (m n k : Int) : (m + n % k) % k = (m + n) % k := by - rw [Int.add_comm, emod_add_emod, Int.add_comm] - -theorem add_emod (a b n : Int) : (a + b) % n = (a % n + b % n) % n := by - rw [add_emod_emod, emod_add_emod] - -theorem add_emod_eq_add_emod_right {m n k : Int} (i : Int) - (H : m % n = k % n) : (m + i) % n = (k + i) % n := by - rw [← emod_add_emod, ← emod_add_emod k, H] - theorem add_emod_eq_add_emod_left {m n k : Int} (i : Int) (H : m % n = k % n) : (i + m) % n = (i + k) % n := by rw [Int.add_comm, add_emod_eq_add_emod_right _ H, Int.add_comm] -theorem emod_add_cancel_right {m n k : Int} (i) : (m + i) % n = (k + i) % n ↔ m % n = k % n := - ⟨fun H => by - have := add_emod_eq_add_emod_right (-i) H - rwa [Int.add_neg_cancel_right, Int.add_neg_cancel_right] at this, - add_emod_eq_add_emod_right _⟩ - theorem emod_add_cancel_left {m n k i : Int} : (i + m) % n = (i + k) % n ↔ m % n = k % n := by rw [Int.add_comm, Int.add_comm i, emod_add_cancel_right] @@ -489,46 +324,18 @@ theorem emod_eq_emod_iff_emod_sub_eq_zero {m n k : Int} : m % n = k % n ↔ (m - if h : b = 0 then by simp [h, Int.mul_zero] else by rw [Int.fmod_def, Int.mul_fdiv_cancel _ h, Int.mul_comm, Int.sub_self] -@[simp] theorem mul_emod_left (a b : Int) : (a * b) % b = 0 := by - rw [← Int.zero_add (a * b), Int.add_mul_emod_self, Int.zero_emod] - @[simp] theorem mul_mod_right (a b : Int) : (a * b).mod a = 0 := by rw [Int.mul_comm, mul_mod_left] @[simp] theorem mul_fmod_right (a b : Int) : (a * b).fmod a = 0 := by rw [Int.mul_comm, mul_fmod_left] -@[simp] theorem mul_emod_right (a b : Int) : (a * b) % a = 0 := by - rw [Int.mul_comm, mul_emod_left] - -theorem mul_emod (a b n : Int) : (a * b) % n = (a % n) * (b % n) % n := by - conv => lhs; rw [ - ← emod_add_ediv a n, ← emod_add_ediv' b n, Int.add_mul, Int.mul_add, Int.mul_add, - Int.mul_assoc, Int.mul_assoc, ← Int.mul_add n _ _, add_mul_emod_self_left, - ← Int.mul_assoc, add_mul_emod_self] - @[simp] theorem mod_self {a : Int} : a.mod a = 0 := by have := mul_mod_left 1 a; rwa [Int.one_mul] at this @[simp] theorem fmod_self {a : Int} : a.fmod a = 0 := by have := mul_fmod_left 1 a; rwa [Int.one_mul] at this -@[local simp] theorem emod_self {a : Int} : a % a = 0 := by - have := mul_emod_left 1 a; rwa [Int.one_mul] at this - -@[simp] theorem emod_emod_of_dvd (n : Int) {m k : Int} - (h : m ∣ k) : (n % k) % m = n % m := by - conv => rhs; rw [← emod_add_ediv n k] - match k, h with - | _, ⟨t, rfl⟩ => rw [Int.mul_assoc, add_mul_emod_self_left] - -@[simp] theorem emod_emod (a b : Int) : (a % b) % b = a % b := by - conv => rhs; rw [← emod_add_ediv a b, add_mul_emod_self_left] - -theorem sub_emod (a b n : Int) : (a - b) % n = (a % n - b % n) % n := by - apply (emod_add_cancel_right b).mp - rw [Int.sub_add_cancel, ← Int.add_emod_emod, Int.sub_add_cancel, emod_emod] - protected theorem ediv_emod_unique {a b r q : Int} (h : 0 < b) : a / b = q ∧ a % b = r ↔ r + b * q = a ∧ 0 ≤ r ∧ r < b := by constructor @@ -613,44 +420,8 @@ theorem mul_div_cancel_of_mod_eq_zero {a b : Int} (H : a.mod b = 0) : b * (a.div theorem div_mul_cancel_of_mod_eq_zero {a b : Int} (H : a.mod b = 0) : a.div b * b = a := by rw [Int.mul_comm, mul_div_cancel_of_mod_eq_zero H] -theorem mul_ediv_cancel_of_emod_eq_zero {a b : Int} (H : a % b = 0) : b * (a / b) = a := by - have := emod_add_ediv a b; rwa [H, Int.zero_add] at this - -theorem ediv_mul_cancel_of_emod_eq_zero {a b : Int} (H : a % b = 0) : a / b * b = a := by - rw [Int.mul_comm, mul_ediv_cancel_of_emod_eq_zero H] - /-! ### dvd -/ -protected theorem dvd_zero (n : Int) : n ∣ 0 := ⟨0, (Int.mul_zero _).symm⟩ - -protected theorem dvd_refl (n : Int) : n ∣ n := ⟨1, (Int.mul_one _).symm⟩ - -protected theorem one_dvd (n : Int) : 1 ∣ n := ⟨n, (Int.one_mul n).symm⟩ - -protected theorem dvd_trans : ∀ {a b c : Int}, a ∣ b → b ∣ c → a ∣ c - | _, _, _, ⟨d, rfl⟩, ⟨e, rfl⟩ => ⟨d * e, by rw [Int.mul_assoc]⟩ - -@[simp] protected theorem zero_dvd {n : Int} : 0 ∣ n ↔ n = 0 := - ⟨fun ⟨k, e⟩ => by rw [e, Int.zero_mul], fun h => h.symm ▸ Int.dvd_refl _⟩ - -protected theorem neg_dvd {a b : Int} : -a ∣ b ↔ a ∣ b := by - constructor <;> exact fun ⟨k, e⟩ => - ⟨-k, by simp [e, Int.neg_mul, Int.mul_neg, Int.neg_neg]⟩ - -protected theorem dvd_neg {a b : Int} : a ∣ -b ↔ a ∣ b := by - constructor <;> exact fun ⟨k, e⟩ => - ⟨-k, by simp [← e, Int.neg_mul, Int.mul_neg, Int.neg_neg]⟩ - -protected theorem dvd_mul_right (a b : Int) : a ∣ a * b := ⟨_, rfl⟩ - -protected theorem dvd_mul_left (a b : Int) : b ∣ a * b := ⟨_, Int.mul_comm ..⟩ - -protected theorem dvd_add : ∀ {a b c : Int}, a ∣ b → a ∣ c → a ∣ b + c - | _, _, _, ⟨d, rfl⟩, ⟨e, rfl⟩ => ⟨d + e, by rw [Int.mul_add]⟩ - -protected theorem dvd_sub : ∀ {a b c : Int}, a ∣ b → a ∣ c → a ∣ b - c - | _, _, _, ⟨d, rfl⟩, ⟨e, rfl⟩ => ⟨d - e, by rw [Int.mul_sub]⟩ - protected theorem dvd_add_left {a b c : Int} (H : a ∣ c) : a ∣ b + c ↔ a ∣ b := ⟨fun h => by have := Int.dvd_sub h H; rwa [Int.add_sub_cancel] at this, (Int.dvd_add · H)⟩ @@ -664,23 +435,6 @@ protected theorem dvd_iff_dvd_of_dvd_sub {a b c : Int} (H : a ∣ b - c) : a ∣ protected theorem dvd_iff_dvd_of_dvd_add {a b c : Int} (H : a ∣ b + c) : a ∣ b ↔ a ∣ c := by rw [← Int.sub_neg] at H; rw [Int.dvd_iff_dvd_of_dvd_sub H, Int.dvd_neg] -@[norm_cast] theorem ofNat_dvd {m n : Nat} : (↑m : Int) ∣ ↑n ↔ m ∣ n := by - refine ⟨fun ⟨a, ae⟩ => ?_, fun ⟨k, e⟩ => ⟨k, by rw [e, Int.ofNat_mul]⟩⟩ - match Int.le_total a 0 with - | .inl h => - have := ae.symm ▸ Int.mul_nonpos_of_nonneg_of_nonpos (ofNat_zero_le _) h - rw [Nat.le_antisymm (ofNat_le.1 this) (Nat.zero_le _)] - apply Nat.dvd_zero - | .inr h => match a, eq_ofNat_of_zero_le h with - | _, ⟨k, rfl⟩ => exact ⟨k, Int.ofNat.inj ae⟩ - -@[simp] theorem natAbs_dvd_natAbs {a b : Int} : natAbs a ∣ natAbs b ↔ a ∣ b := by - refine ⟨fun ⟨k, hk⟩ => ?_, fun ⟨k, hk⟩ => ⟨natAbs k, hk.symm ▸ natAbs_mul a k⟩⟩ - rw [← natAbs_ofNat k, ← natAbs_mul, natAbs_eq_natAbs_iff] at hk - cases hk <;> subst b - · apply Int.dvd_mul_right - · rw [← Int.mul_neg]; apply Int.dvd_mul_right - theorem natAbs_dvd {a b : Int} : (a.natAbs : Int) ∣ b ↔ a ∣ b := match natAbs_eq a with | .inl e => by rw [← e] @@ -699,9 +453,6 @@ theorem dvd_natAbs_self {a : Int} : a ∣ (a.natAbs : Int) := by rw [Int.dvd_natAbs] exact Int.dvd_refl a -theorem ofNat_dvd_left {n : Nat} {z : Int} : (↑n : Int) ∣ z ↔ n ∣ z.natAbs := by - rw [← natAbs_dvd_natAbs, natAbs_ofNat] - theorem ofNat_dvd_right {n : Nat} {z : Int} : z ∣ (↑n : Int) ↔ z.natAbs ∣ n := by rw [← natAbs_dvd_natAbs, natAbs_ofNat] @@ -713,33 +464,12 @@ theorem dvd_antisymm {a b : Int} (H1 : 0 ≤ a) (H2 : 0 ≤ b) : a ∣ b → b theorem dvd_of_mod_eq_zero {a b : Int} (H : mod b a = 0) : a ∣ b := ⟨b.div a, (mul_div_cancel_of_mod_eq_zero H).symm⟩ -theorem dvd_of_emod_eq_zero {a b : Int} (H : b % a = 0) : a ∣ b := - ⟨b / a, (mul_ediv_cancel_of_emod_eq_zero H).symm⟩ - -theorem dvd_emod_sub_self {x : Int} {m : Nat} : (m : Int) ∣ x % m - x := by - apply dvd_of_emod_eq_zero - simp [sub_emod] - theorem mod_eq_zero_of_dvd : ∀ {a b : Int}, a ∣ b → mod b a = 0 | _, _, ⟨_, rfl⟩ => mul_mod_right .. -theorem emod_eq_zero_of_dvd : ∀ {a b : Int}, a ∣ b → b % a = 0 - | _, _, ⟨_, rfl⟩ => mul_emod_right .. - theorem dvd_iff_mod_eq_zero (a b : Int) : a ∣ b ↔ mod b a = 0 := ⟨mod_eq_zero_of_dvd, dvd_of_mod_eq_zero⟩ -theorem dvd_iff_emod_eq_zero (a b : Int) : a ∣ b ↔ b % a = 0 := - ⟨emod_eq_zero_of_dvd, dvd_of_emod_eq_zero⟩ - -theorem emod_pos_of_not_dvd {a b : Int} (h : ¬ a ∣ b) : a = 0 ∨ 0 < b % a := by - rw [dvd_iff_emod_eq_zero] at h - if w : a = 0 then simp_all - else exact Or.inr (Int.lt_iff_le_and_ne.mpr ⟨emod_nonneg b w, Ne.symm h⟩) - -instance decidableDvd : DecidableRel (α := Int) (· ∣ ·) := fun _ _ => - decidable_of_decidable_of_iff (dvd_iff_mod_eq_zero ..).symm - /-- If `a % b = c` then `b` divides `a - c`. -/ theorem dvd_sub_of_emod_eq {a b c : Int} (h : a % b = c) : b ∣ a - c := by have hx : (a % b) % b = c % b := by @@ -750,33 +480,18 @@ theorem dvd_sub_of_emod_eq {a b c : Int} (h : a % b = c) : b ∣ a - c := by protected theorem div_mul_cancel {a b : Int} (H : b ∣ a) : a.div b * b = a := div_mul_cancel_of_mod_eq_zero (mod_eq_zero_of_dvd H) -protected theorem ediv_mul_cancel {a b : Int} (H : b ∣ a) : a / b * b = a := - ediv_mul_cancel_of_emod_eq_zero (emod_eq_zero_of_dvd H) - protected theorem mul_div_cancel' {a b : Int} (H : a ∣ b) : a * b.div a = b := by rw [Int.mul_comm, Int.div_mul_cancel H] -protected theorem mul_ediv_cancel' {a b : Int} (H : a ∣ b) : a * (b / a) = b := by - rw [Int.mul_comm, Int.ediv_mul_cancel H] - protected theorem mul_div_assoc (a : Int) : ∀ {b c : Int}, c ∣ b → (a * b).div c = a * (b.div c) | _, c, ⟨d, rfl⟩ => if cz : c = 0 then by simp [cz, Int.mul_zero] else by rw [Int.mul_left_comm, Int.mul_div_cancel_left _ cz, Int.mul_div_cancel_left _ cz] -protected theorem mul_ediv_assoc (a : Int) : ∀ {b c : Int}, c ∣ b → (a * b) / c = a * (b / c) - | _, c, ⟨d, rfl⟩ => - if cz : c = 0 then by simp [cz, Int.mul_zero] else by - rw [Int.mul_left_comm, Int.mul_ediv_cancel_left _ cz, Int.mul_ediv_cancel_left _ cz] - protected theorem mul_div_assoc' (b : Int) {a c : Int} (h : c ∣ a) : (a * b).div c = a.div c * b := by rw [Int.mul_comm, Int.mul_div_assoc _ h, Int.mul_comm] -protected theorem mul_ediv_assoc' (b : Int) {a c : Int} - (h : c ∣ a) : (a * b) / c = a / c * b := by - rw [Int.mul_comm, Int.mul_ediv_assoc _ h, Int.mul_comm] - theorem div_dvd_div : ∀ {a b c : Int}, a ∣ b → b ∣ c → b.div a ∣ c.div a | a, _, _, ⟨b, rfl⟩, ⟨c, rfl⟩ => by if az : a = 0 then simp [az] else @@ -849,15 +564,6 @@ theorem fdiv_eq_ediv_of_dvd : ∀ {a b : Int}, b ∣ a → a.fdiv b = a / b | _, b, ⟨c, rfl⟩ => by if bz : b = 0 then simp [bz] else rw [mul_fdiv_cancel_left _ bz, mul_ediv_cancel_left _ bz] -theorem neg_ediv_of_dvd : ∀ {a b : Int}, b ∣ a → (-a) / b = -(a / b) - | _, b, ⟨c, rfl⟩ => by if bz : b = 0 then simp [bz] else - rw [Int.neg_mul_eq_mul_neg, Int.mul_ediv_cancel_left _ bz, Int.mul_ediv_cancel_left _ bz] - -theorem sub_ediv_of_dvd (a : Int) {b c : Int} - (hcb : c ∣ b) : (a - b) / c = a / c - b / c := by - rw [Int.sub_eq_add_neg, Int.sub_eq_add_neg, Int.add_ediv_of_dvd_right (Int.dvd_neg.2 hcb)] - congr; exact Int.neg_ediv_of_dvd hcb - theorem sub_ediv_of_dvd_sub {a b c : Int} (hcab : c ∣ a - b) : (a - b) / c = a / c - b / c := by rw [← Int.add_sub_cancel ((a-b) / c), ← Int.add_ediv_of_dvd_left hcab, Int.sub_add_cancel] @@ -910,24 +616,8 @@ theorem eq_one_of_mul_eq_one_left {a b : Int} (H : 0 ≤ b) (H' : a * b = 1) : b /-! # `bmod` ("balanced" mod) -We use balanced mod in the omega algorithm, -to make ±1 coefficients appear in equations without them. -/ -/-- -Balanced mod, taking values in the range [- m/2, (m - 1)/2]. --/ -def bmod (x : Int) (m : Nat) : Int := - let r := x % m - if r < (m + 1) / 2 then - r - else - r - m - -@[simp] theorem bmod_emod : bmod x m % m = x % m := by - dsimp [bmod] - split <;> simp [Int.sub_emod] - @[simp] theorem emod_bmod {x : Int} {m : Nat} : bmod (x % m) m = bmod x m := by simp [bmod] diff --git a/Std/Data/Int/Init/DivMod.lean b/Std/Data/Int/Init/DivMod.lean new file mode 100644 index 0000000000..02df034871 --- /dev/null +++ b/Std/Data/Int/Init/DivMod.lean @@ -0,0 +1,342 @@ +/- +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 Std.Data.Int.Init.Order +import Std.Tactic.Change +import Std.Tactic.RCases + +/-! +# Lemmas about integer division needed to bootstrap `omega`. +-/ + + +open Nat + +namespace Int + +/-! ### `/` -/ + +@[simp, norm_cast] theorem ofNat_ediv (m n : Nat) : (↑(m / n) : Int) = ↑m / ↑n := rfl + +@[simp] theorem zero_ediv : ∀ b : Int, 0 / b = 0 + | ofNat _ => show ofNat _ = _ by simp + | -[_+1] => show -ofNat _ = _ by simp + +@[simp] protected theorem ediv_zero : ∀ a : Int, a / 0 = 0 + | ofNat _ => show ofNat _ = _ by simp + | -[_+1] => rfl + +@[simp] protected theorem ediv_neg : ∀ a b : Int, a / (-b) = -(a / b) + | ofNat m, 0 => show ofNat (m / 0) = -↑(m / 0) by rw [Nat.div_zero]; rfl + | ofNat m, -[n+1] => (Int.neg_neg _).symm + | ofNat m, succ n | -[m+1], 0 | -[m+1], succ n | -[m+1], -[n+1] => rfl + +protected theorem div_def (a b : Int) : a / b = Int.ediv a b := rfl + +theorem add_mul_ediv_right (a b : Int) {c : Int} (H : c ≠ 0) : (a + b * c) / c = a / c + b := + suffices ∀ {{a b c : Int}}, 0 < c → (a + b * c).ediv c = a.ediv c + b from + match Int.lt_trichotomy c 0 with + | Or.inl hlt => by + rw [← Int.neg_inj, ← Int.ediv_neg, Int.neg_add, ← Int.ediv_neg, ← Int.neg_mul_neg] + exact this (Int.neg_pos_of_neg hlt) + | Or.inr (Or.inl HEq) => absurd HEq H + | Or.inr (Or.inr hgt) => this hgt + suffices ∀ {k n : Nat} {a : Int}, (a + n * k.succ).ediv k.succ = a.ediv k.succ + n from + fun a b c H => match c, eq_succ_of_zero_lt H, b with + | _, ⟨_, rfl⟩, ofNat _ => this + | _, ⟨k, rfl⟩, -[n+1] => show (a - n.succ * k.succ).ediv k.succ = a.ediv k.succ - n.succ by + rw [← Int.add_sub_cancel (ediv ..), ← this, Int.sub_add_cancel] + fun {k n} => @fun + | ofNat m => congrArg ofNat <| Nat.add_mul_div_right _ _ k.succ_pos + | -[m+1] => by + show ((n * k.succ : Nat) - m.succ : Int).ediv k.succ = n - (m / k.succ + 1 : Nat) + if h : m < n * k.succ then + rw [← Int.ofNat_sub h, ← Int.ofNat_sub ((Nat.div_lt_iff_lt_mul k.succ_pos).2 h)] + apply congrArg ofNat + rw [Nat.mul_comm, Nat.mul_sub_div]; rwa [Nat.mul_comm] + else + have h := Nat.not_lt.1 h + have H {a b : Nat} (h : a ≤ b) : (a : Int) + -((b : Int) + 1) = -[b - a +1] := by + rw [negSucc_eq, Int.ofNat_sub h] + simp only [Int.sub_eq_add_neg, Int.neg_add, Int.neg_neg, Int.add_left_comm, Int.add_assoc] + show ediv (↑(n * succ k) + -((m : Int) + 1)) (succ k) = n + -(↑(m / succ k) + 1 : Int) + rw [H h, H ((Nat.le_div_iff_mul_le k.succ_pos).2 h)] + apply congrArg negSucc + rw [Nat.mul_comm, Nat.sub_mul_div]; rwa [Nat.mul_comm] + +theorem add_ediv_of_dvd_right {a b c : Int} (H : c ∣ b) : (a + b) / c = a / c + b / c := + if h : c = 0 then by simp [h] else by + let ⟨k, hk⟩ := H + rw [hk, Int.mul_comm c k, Int.add_mul_ediv_right _ _ h, + ← Int.zero_add (k * c), Int.add_mul_ediv_right _ _ h, Int.zero_ediv, Int.zero_add] + +theorem add_ediv_of_dvd_left {a b c : Int} (H : c ∣ a) : (a + b) / c = a / c + b / c := by + rw [Int.add_comm, Int.add_ediv_of_dvd_right H, Int.add_comm] + +@[simp] theorem mul_ediv_cancel (a : Int) {b : Int} (H : b ≠ 0) : (a * b) / b = a := by + have := Int.add_mul_ediv_right 0 a H + rwa [Int.zero_add, Int.zero_ediv, Int.zero_add] at this + +@[simp] theorem mul_ediv_cancel_left (b : Int) (H : a ≠ 0) : (a * b) / a = b := + Int.mul_comm .. ▸ Int.mul_ediv_cancel _ H + +theorem div_nonneg_iff_of_pos {a b : Int} (h : 0 < b) : a / b ≥ 0 ↔ a ≥ 0 := by + rw [Int.div_def] + match b, h with + | Int.ofNat (b+1), _ => + rcases a with ⟨a⟩ <;> simp [Int.ediv] + exact decide_eq_decide.mp rfl + +/-! ### mod -/ + +theorem mod_def' (m n : Int) : m % n = emod m n := rfl + +theorem ofNat_mod (m n : Nat) : (↑(m % n) : Int) = mod m n := rfl + +theorem ofNat_mod_ofNat (m n : Nat) : (m % n : Int) = ↑(m % n) := rfl + +@[simp, norm_cast] theorem ofNat_emod (m n : Nat) : (↑(m % n) : Int) = m % n := rfl + +@[simp] theorem zero_emod (b : Int) : 0 % b = 0 := by simp [mod_def', emod] + +@[simp] theorem emod_zero : ∀ a : Int, a % 0 = a + | ofNat _ => congrArg ofNat <| Nat.mod_zero _ + | -[_+1] => congrArg negSucc <| Nat.mod_zero _ + +theorem emod_add_ediv : ∀ a b : Int, a % b + b * (a / b) = a + | ofNat _, ofNat _ => congrArg ofNat <| Nat.mod_add_div .. + | ofNat m, -[n+1] => by + show (m % succ n + -↑(succ n) * -↑(m / succ n) : Int) = m + rw [Int.neg_mul_neg]; exact congrArg ofNat <| Nat.mod_add_div .. + | -[_+1], 0 => by rw [emod_zero]; rfl + | -[m+1], succ n => aux m n.succ + | -[m+1], -[n+1] => aux m n.succ +where + aux (m n : Nat) : n - (m % n + 1) - (n * (m / n) + n) = -[m+1] := by + rw [← ofNat_emod, ← ofNat_ediv, ← Int.sub_sub, negSucc_eq, Int.sub_sub n, + ← Int.neg_neg (_-_), Int.neg_sub, Int.sub_sub_self, Int.add_right_comm] + exact congrArg (fun x => -(ofNat x + 1)) (Nat.mod_add_div ..) + +theorem ediv_add_emod (a b : Int) : b * (a / b) + a % b = a := + (Int.add_comm ..).trans (emod_add_ediv ..) + +theorem emod_def (a b : Int) : a % b = a - b * (a / b) := by + rw [← Int.add_sub_cancel (a % b), emod_add_ediv] + +theorem emod_nonneg : ∀ (a : Int) {b : Int}, b ≠ 0 → 0 ≤ a % b + | ofNat _, _, _ => ofNat_zero_le _ + | -[_+1], _, H => Int.sub_nonneg_of_le <| ofNat_le.2 <| Nat.mod_lt _ (natAbs_pos.2 H) + +theorem emod_lt_of_pos (a : Int) {b : Int} (H : 0 < b) : a % b < b := + match a, b, eq_succ_of_zero_lt H with + | ofNat _, _, ⟨_, rfl⟩ => ofNat_lt.2 (Nat.mod_lt _ (Nat.succ_pos _)) + | -[_+1], _, ⟨_, rfl⟩ => Int.sub_lt_self _ (ofNat_lt.2 <| Nat.succ_pos _) + +theorem mul_ediv_self_le {x k : Int} (h : k ≠ 0) : k * (x / k) ≤ x := + calc k * (x / k) + _ ≤ k * (x / k) + x % k := Int.le_add_of_nonneg_right (emod_nonneg x h) + _ = x := ediv_add_emod _ _ + +theorem lt_mul_ediv_self_add {x k : Int} (h : 0 < k) : x < k * (x / k) + k := + calc x + _ = k * (x / k) + x % k := (ediv_add_emod _ _).symm + _ < k * (x / k) + k := Int.add_lt_add_left (emod_lt_of_pos x h) _ + +theorem emod_add_ediv' (m k : Int) : m % k + m / k * k = m := by + rw [Int.mul_comm]; apply emod_add_ediv + +@[simp] theorem add_mul_emod_self {a b c : Int} : (a + b * c) % c = a % c := + if cz : c = 0 then by + rw [cz, Int.mul_zero, Int.add_zero] + else by + rw [Int.emod_def, Int.emod_def, Int.add_mul_ediv_right _ _ cz, Int.add_comm _ b, + Int.mul_add, Int.mul_comm, ← Int.sub_sub, Int.add_sub_cancel] + +@[simp] theorem add_mul_emod_self_left (a b c : Int) : (a + b * c) % b = a % b := by + rw [Int.mul_comm, Int.add_mul_emod_self] + +@[simp] theorem add_emod_self {a b : Int} : (a + b) % b = a % b := by + have := add_mul_emod_self_left a b 1; rwa [Int.mul_one] at this + +@[simp] theorem add_emod_self_left {a b : Int} : (a + b) % a = b % a := by + rw [Int.add_comm, Int.add_emod_self] + +theorem neg_emod {a b : Int} : -a % b = (b - a) % b := by + rw [← add_emod_self_left]; rfl + +@[simp] theorem emod_add_emod (m n k : Int) : (m % n + k) % n = (m + k) % n := by + have := (add_mul_emod_self_left (m % n + k) n (m / n)).symm + rwa [Int.add_right_comm, emod_add_ediv] at this + +@[simp] theorem add_emod_emod (m n k : Int) : (m + n % k) % k = (m + n) % k := by + rw [Int.add_comm, emod_add_emod, Int.add_comm] + +theorem add_emod (a b n : Int) : (a + b) % n = (a % n + b % n) % n := by + rw [add_emod_emod, emod_add_emod] + +theorem add_emod_eq_add_emod_right {m n k : Int} (i : Int) + (H : m % n = k % n) : (m + i) % n = (k + i) % n := by + rw [← emod_add_emod, ← emod_add_emod k, H] + +theorem emod_add_cancel_right {m n k : Int} (i) : (m + i) % n = (k + i) % n ↔ m % n = k % n := + ⟨fun H => by + have := add_emod_eq_add_emod_right (-i) H + rwa [Int.add_neg_cancel_right, Int.add_neg_cancel_right] at this, + add_emod_eq_add_emod_right _⟩ + +@[simp] theorem mul_emod_left (a b : Int) : (a * b) % b = 0 := by + rw [← Int.zero_add (a * b), Int.add_mul_emod_self, Int.zero_emod] + +@[simp] theorem mul_emod_right (a b : Int) : (a * b) % a = 0 := by + rw [Int.mul_comm, mul_emod_left] + +theorem mul_emod (a b n : Int) : (a * b) % n = (a % n) * (b % n) % n := by + conv => lhs; rw [ + ← emod_add_ediv a n, ← emod_add_ediv' b n, Int.add_mul, Int.mul_add, Int.mul_add, + Int.mul_assoc, Int.mul_assoc, ← Int.mul_add n _ _, add_mul_emod_self_left, + ← Int.mul_assoc, add_mul_emod_self] + +@[local simp] theorem emod_self {a : Int} : a % a = 0 := by + have := mul_emod_left 1 a; rwa [Int.one_mul] at this + +@[simp] theorem emod_emod_of_dvd (n : Int) {m k : Int} + (h : m ∣ k) : (n % k) % m = n % m := by + conv => rhs; rw [← emod_add_ediv n k] + match k, h with + | _, ⟨t, rfl⟩ => rw [Int.mul_assoc, add_mul_emod_self_left] + +@[simp] theorem emod_emod (a b : Int) : (a % b) % b = a % b := by + conv => rhs; rw [← emod_add_ediv a b, add_mul_emod_self_left] + +theorem sub_emod (a b n : Int) : (a - b) % n = (a % n - b % n) % n := by + apply (emod_add_cancel_right b).mp + rw [Int.sub_add_cancel, ← Int.add_emod_emod, Int.sub_add_cancel, emod_emod] + +/-! ### properties of `/` and `%` -/ + +theorem mul_ediv_cancel_of_emod_eq_zero {a b : Int} (H : a % b = 0) : b * (a / b) = a := by + have := emod_add_ediv a b; rwa [H, Int.zero_add] at this + +theorem ediv_mul_cancel_of_emod_eq_zero {a b : Int} (H : a % b = 0) : a / b * b = a := by + rw [Int.mul_comm, mul_ediv_cancel_of_emod_eq_zero H] + +/-! ### dvd -/ + +protected theorem dvd_zero (n : Int) : n ∣ 0 := ⟨0, (Int.mul_zero _).symm⟩ + +protected theorem dvd_refl (n : Int) : n ∣ n := ⟨1, (Int.mul_one _).symm⟩ + +protected theorem one_dvd (n : Int) : 1 ∣ n := ⟨n, (Int.one_mul n).symm⟩ + +protected theorem dvd_trans : ∀ {a b c : Int}, a ∣ b → b ∣ c → a ∣ c + | _, _, _, ⟨d, rfl⟩, ⟨e, rfl⟩ => ⟨d * e, by rw [Int.mul_assoc]⟩ + +@[simp] protected theorem zero_dvd {n : Int} : 0 ∣ n ↔ n = 0 := + ⟨fun ⟨k, e⟩ => by rw [e, Int.zero_mul], fun h => h.symm ▸ Int.dvd_refl _⟩ + +protected theorem neg_dvd {a b : Int} : -a ∣ b ↔ a ∣ b := by + constructor <;> exact fun ⟨k, e⟩ => + ⟨-k, by simp [e, Int.neg_mul, Int.mul_neg, Int.neg_neg]⟩ + +protected theorem dvd_neg {a b : Int} : a ∣ -b ↔ a ∣ b := by + constructor <;> exact fun ⟨k, e⟩ => + ⟨-k, by simp [← e, Int.neg_mul, Int.mul_neg, Int.neg_neg]⟩ + +protected theorem dvd_mul_right (a b : Int) : a ∣ a * b := ⟨_, rfl⟩ + +protected theorem dvd_mul_left (a b : Int) : b ∣ a * b := ⟨_, Int.mul_comm ..⟩ + +protected theorem dvd_add : ∀ {a b c : Int}, a ∣ b → a ∣ c → a ∣ b + c + | _, _, _, ⟨d, rfl⟩, ⟨e, rfl⟩ => ⟨d + e, by rw [Int.mul_add]⟩ + +protected theorem dvd_sub : ∀ {a b c : Int}, a ∣ b → a ∣ c → a ∣ b - c + | _, _, _, ⟨d, rfl⟩, ⟨e, rfl⟩ => ⟨d - e, by rw [Int.mul_sub]⟩ + + +@[norm_cast] theorem ofNat_dvd {m n : Nat} : (↑m : Int) ∣ ↑n ↔ m ∣ n := by + refine ⟨fun ⟨a, ae⟩ => ?_, fun ⟨k, e⟩ => ⟨k, by rw [e, Int.ofNat_mul]⟩⟩ + match Int.le_total a 0 with + | .inl h => + have := ae.symm ▸ Int.mul_nonpos_of_nonneg_of_nonpos (ofNat_zero_le _) h + rw [Nat.le_antisymm (ofNat_le.1 this) (Nat.zero_le _)] + apply Nat.dvd_zero + | .inr h => match a, eq_ofNat_of_zero_le h with + | _, ⟨k, rfl⟩ => exact ⟨k, Int.ofNat.inj ae⟩ + +@[simp] theorem natAbs_dvd_natAbs {a b : Int} : natAbs a ∣ natAbs b ↔ a ∣ b := by + refine ⟨fun ⟨k, hk⟩ => ?_, fun ⟨k, hk⟩ => ⟨natAbs k, hk.symm ▸ natAbs_mul a k⟩⟩ + rw [← natAbs_ofNat k, ← natAbs_mul, natAbs_eq_natAbs_iff] at hk + cases hk <;> subst b + · apply Int.dvd_mul_right + · rw [← Int.mul_neg]; apply Int.dvd_mul_right + +theorem ofNat_dvd_left {n : Nat} {z : Int} : (↑n : Int) ∣ z ↔ n ∣ z.natAbs := by + rw [← natAbs_dvd_natAbs, natAbs_ofNat] + +theorem dvd_of_emod_eq_zero {a b : Int} (H : b % a = 0) : a ∣ b := + ⟨b / a, (mul_ediv_cancel_of_emod_eq_zero H).symm⟩ + +theorem dvd_emod_sub_self {x : Int} {m : Nat} : (m : Int) ∣ x % m - x := by + apply dvd_of_emod_eq_zero + simp [sub_emod] + +theorem emod_eq_zero_of_dvd : ∀ {a b : Int}, a ∣ b → b % a = 0 + | _, _, ⟨_, rfl⟩ => mul_emod_right .. + +theorem dvd_iff_emod_eq_zero (a b : Int) : a ∣ b ↔ b % a = 0 := + ⟨emod_eq_zero_of_dvd, dvd_of_emod_eq_zero⟩ + +theorem emod_pos_of_not_dvd {a b : Int} (h : ¬ a ∣ b) : a = 0 ∨ 0 < b % a := by + rw [dvd_iff_emod_eq_zero] at h + if w : a = 0 then simp_all + else exact Or.inr (Int.lt_iff_le_and_ne.mpr ⟨emod_nonneg b w, Ne.symm h⟩) + +instance decidableDvd : DecidableRel (α := Int) (· ∣ ·) := fun _ _ => + decidable_of_decidable_of_iff (dvd_iff_emod_eq_zero ..).symm + +protected theorem ediv_mul_cancel {a b : Int} (H : b ∣ a) : a / b * b = a := + ediv_mul_cancel_of_emod_eq_zero (emod_eq_zero_of_dvd H) + +protected theorem mul_ediv_cancel' {a b : Int} (H : a ∣ b) : a * (b / a) = b := by + rw [Int.mul_comm, Int.ediv_mul_cancel H] + +protected theorem mul_ediv_assoc (a : Int) : ∀ {b c : Int}, c ∣ b → (a * b) / c = a * (b / c) + | _, c, ⟨d, rfl⟩ => + if cz : c = 0 then by simp [cz, Int.mul_zero] else by + rw [Int.mul_left_comm, Int.mul_ediv_cancel_left _ cz, Int.mul_ediv_cancel_left _ cz] + +protected theorem mul_ediv_assoc' (b : Int) {a c : Int} + (h : c ∣ a) : (a * b) / c = a / c * b := by + rw [Int.mul_comm, Int.mul_ediv_assoc _ h, Int.mul_comm] + +theorem neg_ediv_of_dvd : ∀ {a b : Int}, b ∣ a → (-a) / b = -(a / b) + | _, b, ⟨c, rfl⟩ => by if bz : b = 0 then simp [bz] else + rw [Int.neg_mul_eq_mul_neg, Int.mul_ediv_cancel_left _ bz, Int.mul_ediv_cancel_left _ bz] + +theorem sub_ediv_of_dvd (a : Int) {b c : Int} + (hcb : c ∣ b) : (a - b) / c = a / c - b / c := by + rw [Int.sub_eq_add_neg, Int.sub_eq_add_neg, Int.add_ediv_of_dvd_right (Int.dvd_neg.2 hcb)] + congr; exact Int.neg_ediv_of_dvd hcb + +/-! +# `bmod` ("balanced" mod) + +We use balanced mod in the omega algorithm, +to make ±1 coefficients appear in equations without them. +-/ + +/-- +Balanced mod, taking values in the range [- m/2, (m - 1)/2]. +-/ +def bmod (x : Int) (m : Nat) : Int := + let r := x % m + if r < (m + 1) / 2 then + r + else + r - m + +@[simp] theorem bmod_emod : bmod x m % m = x % m := by + dsimp [bmod] + split <;> simp [Int.sub_emod] diff --git a/Std/Data/Int/Init/Lemmas.lean b/Std/Data/Int/Init/Lemmas.lean new file mode 100644 index 0000000000..6ef85425f9 --- /dev/null +++ b/Std/Data/Int/Init/Lemmas.lean @@ -0,0 +1,507 @@ +/- +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 Std.Classes.Cast +import Std.Data.Nat.Lemmas +import Std.Data.Int.Basic +import Std.Tactic.NormCast.Lemmas + +open Nat + +namespace Int + +@[simp] theorem ofNat_eq_coe : ofNat n = Nat.cast n := rfl + +@[simp] theorem ofNat_zero : ((0 : Nat) : Int) = 0 := rfl + +@[simp] theorem ofNat_one : ((1 : Nat) : Int) = 1 := rfl + +theorem ofNat_two : ((2 : Nat) : Int) = 2 := rfl + +@[simp] theorem default_eq_zero : default = (0 : Int) := rfl + +protected theorem zero_ne_one : (0 : Int) ≠ 1 := fun. + +/- ## Definitions of basic functions -/ + +theorem subNatNat_of_sub_eq_zero {m n : Nat} (h : n - m = 0) : subNatNat m n = ↑(m - n) := by + rw [subNatNat, h, ofNat_eq_coe] + +theorem subNatNat_of_sub_eq_succ {m n k : Nat} (h : n - m = succ k) : subNatNat m n = -[k+1] := by + rw [subNatNat, h] + +@[simp] protected theorem neg_zero : -(0:Int) = 0 := rfl + +@[norm_cast] theorem ofNat_add (n m : Nat) : (↑(n + m) : Int) = n + m := rfl +@[norm_cast] theorem ofNat_mul (n m : Nat) : (↑(n * m) : Int) = n * m := rfl +theorem ofNat_succ (n : Nat) : (succ n : Int) = n + 1 := rfl + +@[local simp] theorem neg_ofNat_zero : -((0 : Nat) : Int) = 0 := rfl +@[local simp] theorem neg_ofNat_succ (n : Nat) : -(succ n : Int) = -[n+1] := rfl +@[local simp] theorem neg_negSucc (n : Nat) : -(-[n+1]) = succ n := rfl + +theorem negSucc_coe (n : Nat) : -[n+1] = -↑(n + 1) := rfl + +theorem negOfNat_eq : negOfNat n = -ofNat n := rfl + +/- ## These are only for internal use -/ + +@[simp] theorem add_def {a b : Int} : Int.add a b = a + b := rfl + +@[local simp] theorem ofNat_add_ofNat (m n : Nat) : (↑m + ↑n : Int) = ↑(m + n) := rfl +@[local simp] theorem ofNat_add_negSucc (m n : Nat) : ↑m + -[n+1] = subNatNat m (succ n) := rfl +@[local simp] theorem negSucc_add_ofNat (m n : Nat) : -[m+1] + ↑n = subNatNat n (succ m) := rfl +@[local simp] theorem negSucc_add_negSucc (m n : Nat) : -[m+1] + -[n+1] = -[succ (m + n) +1] := rfl + +@[simp] theorem mul_def {a b : Int} : Int.mul a b = a * b := rfl + +@[local simp] theorem ofNat_mul_ofNat (m n : Nat) : (↑m * ↑n : Int) = ↑(m * n) := rfl +@[local simp] theorem ofNat_mul_negSucc' (m n : Nat) : ↑m * -[n+1] = negOfNat (m * succ n) := rfl +@[local simp] theorem negSucc_mul_ofNat' (m n : Nat) : -[m+1] * ↑n = negOfNat (succ m * n) := rfl +@[local simp] theorem negSucc_mul_negSucc' (m n : Nat) : + -[m+1] * -[n+1] = ofNat (succ m * succ n) := rfl + +/- ## some basic functions and properties -/ + +@[norm_cast] theorem ofNat_inj : ((m : Nat) : Int) = (n : Nat) ↔ m = n := ⟨ofNat.inj, congrArg _⟩ + +theorem ofNat_eq_zero : ((n : Nat) : Int) = 0 ↔ n = 0 := ofNat_inj + +theorem ofNat_ne_zero : ((n : Nat) : Int) ≠ 0 ↔ n ≠ 0 := not_congr ofNat_eq_zero + +theorem negSucc_inj : negSucc m = negSucc n ↔ m = n := ⟨negSucc.inj, fun H => by simp [H]⟩ + +theorem negSucc_eq (n : Nat) : -[n+1] = -((n : Int) + 1) := rfl + +@[simp] theorem negSucc_ne_zero (n : Nat) : -[n+1] ≠ 0 := fun. + +@[simp] theorem zero_ne_negSucc (n : Nat) : 0 ≠ -[n+1] := fun. + +@[simp, norm_cast] theorem Nat.cast_ofNat_Int : + (Nat.cast (no_index (OfNat.ofNat n)) : Int) = OfNat.ofNat n := rfl + +/- ## neg -/ + +@[simp] protected theorem neg_neg : ∀ a : Int, -(-a) = a + | 0 => rfl + | succ _ => rfl + | -[_+1] => rfl + +protected theorem neg_inj {a b : Int} : -a = -b ↔ a = b := + ⟨fun h => by rw [← Int.neg_neg a, ← Int.neg_neg b, h], congrArg _⟩ + +@[simp] protected theorem neg_eq_zero : -a = 0 ↔ a = 0 := Int.neg_inj (b := 0) + +protected theorem neg_ne_zero : -a ≠ 0 ↔ a ≠ 0 := not_congr Int.neg_eq_zero + +protected theorem sub_eq_add_neg {a b : Int} : a - b = a + -b := rfl + +theorem add_neg_one (i : Int) : i + -1 = i - 1 := rfl + +/- ## basic properties of subNatNat -/ + +-- @[elabAsElim] -- TODO(Mario): unexpected eliminator resulting type +theorem subNatNat_elim (m n : Nat) (motive : Nat → Nat → Int → Prop) + (hp : ∀ i n, motive (n + i) n i) + (hn : ∀ i m, motive m (m + i + 1) -[i+1]) : + motive m n (subNatNat m n) := by + unfold subNatNat + match h : n - m with + | 0 => + have ⟨k, h⟩ := Nat.le.dest (Nat.le_of_sub_eq_zero h) + rw [h.symm, Nat.add_sub_cancel_left]; apply hp + | succ k => + rw [Nat.sub_eq_iff_eq_add (Nat.le_of_lt (Nat.lt_of_sub_eq_succ h))] at h + rw [h, Nat.add_comm]; apply hn + +theorem subNatNat_add_left : subNatNat (m + n) m = n := by + unfold subNatNat + rw [Nat.sub_eq_zero_of_le (Nat.le_add_right ..), Nat.add_sub_cancel_left, ofNat_eq_coe] + +theorem subNatNat_add_right : subNatNat m (m + n + 1) = negSucc n := by + simp [subNatNat, Nat.add_assoc, Nat.add_sub_cancel_left] + +theorem subNatNat_add_add (m n k : Nat) : subNatNat (m + k) (n + k) = subNatNat m n := by + apply subNatNat_elim m n (fun m n i => subNatNat (m + k) (n + k) = i) + · intro i j + rw [Nat.add_assoc, Nat.add_comm i k, ← Nat.add_assoc] + exact subNatNat_add_left + · intro i j + rw [Nat.add_assoc j i 1, Nat.add_comm j (i+1), Nat.add_assoc, Nat.add_comm (i+1) (j+k)] + exact subNatNat_add_right + +theorem subNatNat_of_le {m n : Nat} (h : n ≤ m) : subNatNat m n = ↑(m - n) := + subNatNat_of_sub_eq_zero (Nat.sub_eq_zero_of_le h) + +theorem subNatNat_of_lt {m n : Nat} (h : m < n) : subNatNat m n = -[pred (n - m) +1] := + subNatNat_of_sub_eq_succ <| (Nat.succ_pred_eq_of_pos (Nat.sub_pos_of_lt h)).symm + + +/- # Additive group properties -/ + +/- addition -/ + +protected theorem add_comm : ∀ a b : Int, a + b = b + a + | ofNat n, ofNat m => by simp [Nat.add_comm] + | ofNat _, -[_+1] => rfl + | -[_+1], ofNat _ => rfl + | -[_+1], -[_+1] => by simp [Nat.add_comm] + +@[simp] protected theorem add_zero : ∀ a : Int, a + 0 = a + | ofNat _ => rfl + | -[_+1] => rfl + +@[simp] protected theorem zero_add (a : Int) : 0 + a = a := Int.add_comm .. ▸ a.add_zero + +theorem ofNat_add_negSucc_of_lt (h : m < n.succ) : ofNat m + -[n+1] = -[n - m+1] := + show subNatNat .. = _ by simp [succ_sub (le_of_lt_succ h), subNatNat] + +theorem subNatNat_sub (h : n ≤ m) (k : Nat) : subNatNat (m - n) k = subNatNat m (k + n) := by + rwa [← subNatNat_add_add _ _ n, Nat.sub_add_cancel] + +theorem subNatNat_add (m n k : Nat) : subNatNat (m + n) k = m + subNatNat n k := by + cases n.lt_or_ge k with + | inl h' => + simp [subNatNat_of_lt h', succ_pred_eq_of_pos (Nat.sub_pos_of_lt h')] + conv => lhs; rw [← Nat.sub_add_cancel (Nat.le_of_lt h')] + apply subNatNat_add_add + | inr h' => simp [subNatNat_of_le h', + subNatNat_of_le (Nat.le_trans h' (le_add_left ..)), Nat.add_sub_assoc h'] + +theorem subNatNat_add_negSucc (m n k : Nat) : + subNatNat m n + -[k+1] = subNatNat m (n + succ k) := by + have h := Nat.lt_or_ge m n + cases h with + | inr h' => + rw [subNatNat_of_le h'] + simp + rw [subNatNat_sub h', Nat.add_comm] + | inl h' => + have h₂ : m < n + succ k := Nat.lt_of_lt_of_le h' (le_add_right _ _) + have h₃ : m ≤ n + k := le_of_succ_le_succ h₂ + rw [subNatNat_of_lt h', subNatNat_of_lt h₂] + simp [Nat.add_comm] + rw [← add_succ, succ_pred_eq_of_pos (Nat.sub_pos_of_lt h'), add_succ, succ_sub h₃, + Nat.pred_succ] + rw [Nat.add_comm n, Nat.add_sub_assoc (Nat.le_of_lt h')] + +protected theorem add_assoc : ∀ a b c : Int, a + b + c = a + (b + c) + | (m:Nat), (n:Nat), c => aux1 .. + | Nat.cast m, b, Nat.cast k => by + rw [Int.add_comm, ← aux1, Int.add_comm k, aux1, Int.add_comm b] + | a, (n:Nat), (k:Nat) => by + rw [Int.add_comm, Int.add_comm a, ← aux1, Int.add_comm a, Int.add_comm k] + | -[m+1], -[n+1], (k:Nat) => aux2 .. + | -[m+1], (n:Nat), -[k+1] => by + rw [Int.add_comm, ← aux2, Int.add_comm n, ← aux2, Int.add_comm -[m+1]] + | (m:Nat), -[n+1], -[k+1] => by + rw [Int.add_comm, Int.add_comm m, Int.add_comm m, ← aux2, Int.add_comm -[k+1]] + | -[m+1], -[n+1], -[k+1] => by + simp [add_succ, Nat.add_comm, Nat.add_left_comm, neg_ofNat_succ] +where + aux1 (m n : Nat) : ∀ c : Int, m + n + c = m + (n + c) + | (k:Nat) => by simp [Nat.add_assoc] + | -[k+1] => by simp [subNatNat_add] + aux2 (m n k : Nat) : -[m+1] + -[n+1] + k = -[m+1] + (-[n+1] + k) := by + simp [add_succ] + rw [Int.add_comm, subNatNat_add_negSucc] + simp [add_succ, succ_add, Nat.add_comm] + +protected theorem add_left_comm (a b c : Int) : a + (b + c) = b + (a + c) := by + rw [← Int.add_assoc, Int.add_comm a, Int.add_assoc] + +protected theorem add_right_comm (a b c : Int) : a + b + c = a + c + b := by + rw [Int.add_assoc, Int.add_comm b, ← Int.add_assoc] + +/- ## negation -/ + +theorem subNatNat_self : ∀ n, subNatNat n n = 0 + | 0 => rfl + | succ m => by rw [subNatNat_of_sub_eq_zero (Nat.sub_self ..), Nat.sub_self, ofNat_zero] + +attribute [local simp] subNatNat_self + +@[local simp] protected theorem add_left_neg : ∀ a : Int, -a + a = 0 + | 0 => rfl + | succ m => by simp + | -[m+1] => by simp + +@[local simp] protected theorem add_right_neg (a : Int) : a + -a = 0 := by + rw [Int.add_comm, Int.add_left_neg] + +@[simp] protected theorem neg_eq_of_add_eq_zero {a b : Int} (h : a + b = 0) : -a = b := by + rw [← Int.add_zero (-a), ← h, ← Int.add_assoc, Int.add_left_neg, Int.zero_add] + +protected theorem eq_neg_of_eq_neg {a b : Int} (h : a = -b) : b = -a := by + rw [h, Int.neg_neg] + +protected theorem eq_neg_comm {a b : Int} : a = -b ↔ b = -a := + ⟨Int.eq_neg_of_eq_neg, Int.eq_neg_of_eq_neg⟩ + +protected theorem neg_eq_comm {a b : Int} : -a = b ↔ -b = a := by + rw [eq_comm, Int.eq_neg_comm, eq_comm] + +protected theorem neg_add_cancel_left (a b : Int) : -a + (a + b) = b := by + rw [← Int.add_assoc, Int.add_left_neg, Int.zero_add] + +protected theorem add_neg_cancel_left (a b : Int) : a + (-a + b) = b := by + rw [← Int.add_assoc, Int.add_right_neg, Int.zero_add] + +protected theorem add_neg_cancel_right (a b : Int) : a + b + -b = a := by + rw [Int.add_assoc, Int.add_right_neg, Int.add_zero] + +protected theorem neg_add_cancel_right (a b : Int) : a + -b + b = a := by + rw [Int.add_assoc, Int.add_left_neg, Int.add_zero] + +protected theorem add_left_cancel {a b c : Int} (h : a + b = a + c) : b = c := by + have h₁ : -a + (a + b) = -a + (a + c) := by rw [h] + simp [← Int.add_assoc, Int.add_left_neg, Int.zero_add] at h₁; exact h₁ + +@[local simp] protected theorem neg_add {a b : Int} : -(a + b) = -a + -b := by + apply Int.add_left_cancel (a := a + b) + rw [Int.add_right_neg, Int.add_comm a, ← Int.add_assoc, Int.add_assoc b, + Int.add_right_neg, Int.add_zero, Int.add_right_neg] + +/- ## subtraction -/ + +@[simp] theorem negSucc_sub_one (n : Nat) : -[n+1] - 1 = -[n + 1 +1] := rfl + +@[simp] protected theorem sub_self (a : Int) : a - a = 0 := by + rw [Int.sub_eq_add_neg, Int.add_right_neg] + +@[simp] protected theorem sub_zero (a : Int) : a - 0 = a := by simp [Int.sub_eq_add_neg] + +@[simp] protected theorem zero_sub (a : Int) : 0 - a = -a := by simp [Int.sub_eq_add_neg] + +protected theorem sub_eq_zero_of_eq {a b : Int} (h : a = b) : a - b = 0 := by + rw [h, Int.sub_self] + +protected theorem eq_of_sub_eq_zero {a b : Int} (h : a - b = 0) : a = b := by + have : 0 + b = b := by rw [Int.zero_add] + have : a - b + b = b := by rwa [h] + rwa [Int.sub_eq_add_neg, Int.neg_add_cancel_right] at this + +protected theorem sub_eq_zero {a b : Int} : a - b = 0 ↔ a = b := + ⟨Int.eq_of_sub_eq_zero, Int.sub_eq_zero_of_eq⟩ + +protected theorem sub_sub (a b c : Int) : a - b - c = a - (b + c) := by + simp [Int.sub_eq_add_neg, Int.add_assoc] + +protected theorem neg_sub (a b : Int) : -(a - b) = b - a := by + simp [Int.sub_eq_add_neg, Int.add_comm] + +protected theorem sub_sub_self (a b : Int) : a - (a - b) = b := by + simp [Int.sub_eq_add_neg, ← Int.add_assoc] + +protected theorem sub_neg (a b : Int) : a - -b = a + b := by simp [Int.sub_eq_add_neg] + +@[simp] protected theorem sub_add_cancel (a b : Int) : a - b + b = a := + Int.neg_add_cancel_right a b + +@[simp] protected theorem add_sub_cancel (a b : Int) : a + b - b = a := + Int.add_neg_cancel_right a b + +protected theorem add_sub_assoc (a b c : Int) : a + b - c = a + (b - c) := by + rw [Int.sub_eq_add_neg, Int.add_assoc, ← Int.sub_eq_add_neg] + +@[norm_cast] theorem ofNat_sub (h : m ≤ n) : ((n - m : Nat) : Int) = n - m := by + match m with + | 0 => rfl + | succ m => + show ofNat (n - succ m) = subNatNat n (succ m) + rw [subNatNat, Nat.sub_eq_zero_of_le h] + +theorem negSucc_coe' (n : Nat) : -[n+1] = -↑n - 1 := by + rw [Int.sub_eq_add_neg, ← Int.neg_add]; rfl + +protected theorem subNatNat_eq_coe {m n : Nat} : subNatNat m n = ↑m - ↑n := by + apply subNatNat_elim m n fun m n i => i = m - n + · intros i n + rw [Int.ofNat_add, Int.sub_eq_add_neg, Int.add_assoc, Int.add_left_comm, + Int.add_right_neg, Int.add_zero] + · intros i n + simp only [negSucc_coe, ofNat_add, Int.sub_eq_add_neg, Int.neg_add, ← Int.add_assoc] + rw [← @Int.sub_eq_add_neg n, ← ofNat_sub, Nat.sub_self, ofNat_zero, Int.zero_add] + apply Nat.le_refl + +theorem toNat_sub (m n : Nat) : toNat (m - n) = m - n := by + rw [← Int.subNatNat_eq_coe] + refine subNatNat_elim m n (fun m n i => toNat i = m - n) (fun i n => ?_) (fun i n => ?_) + · exact (Nat.add_sub_cancel_left ..).symm + · dsimp; rw [Nat.add_assoc, Nat.sub_eq_zero_of_le (Nat.le_add_right ..)]; rfl + +/- ## Ring properties -/ + +@[simp] theorem ofNat_mul_negSucc (m n : Nat) : (m : Int) * -[n+1] = -↑(m * succ n) := rfl + +@[simp] theorem negSucc_mul_ofNat (m n : Nat) : -[m+1] * n = -↑(succ m * n) := rfl + +@[simp] theorem negSucc_mul_negSucc (m n : Nat) : -[m+1] * -[n+1] = succ m * succ n := rfl + +protected theorem mul_comm (a b : Int) : a * b = b * a := by + cases a <;> cases b <;> simp [Nat.mul_comm] + +theorem ofNat_mul_negOfNat (m n : Nat) : (m : Nat) * negOfNat n = negOfNat (m * n) := by + cases n <;> rfl + +theorem negOfNat_mul_ofNat (m n : Nat) : negOfNat m * (n : Nat) = negOfNat (m * n) := by + rw [Int.mul_comm]; simp [ofNat_mul_negOfNat, Nat.mul_comm] + +theorem negSucc_mul_negOfNat (m n : Nat) : -[m+1] * negOfNat n = ofNat (succ m * n) := by + cases n <;> rfl + +theorem negOfNat_mul_negSucc (m n : Nat) : negOfNat n * -[m+1] = ofNat (n * succ m) := by + rw [Int.mul_comm, negSucc_mul_negOfNat, Nat.mul_comm] + +attribute [local simp] ofNat_mul_negOfNat negOfNat_mul_ofNat + negSucc_mul_negOfNat negOfNat_mul_negSucc + +protected theorem mul_assoc (a b c : Int) : a * b * c = a * (b * c) := by + cases a <;> cases b <;> cases c <;> simp [Nat.mul_assoc] + +protected theorem mul_left_comm (a b c : Int) : a * (b * c) = b * (a * c) := by + rw [← Int.mul_assoc, ← Int.mul_assoc, Int.mul_comm a] + +protected theorem mul_right_comm (a b c : Int) : a * b * c = a * c * b := by + rw [Int.mul_assoc, Int.mul_assoc, Int.mul_comm b] + +@[simp] protected theorem mul_zero (a : Int) : a * 0 = 0 := by cases a <;> rfl + +@[simp] protected theorem zero_mul (a : Int) : 0 * a = 0 := Int.mul_comm .. ▸ a.mul_zero + +theorem negOfNat_eq_subNatNat_zero (n) : negOfNat n = subNatNat 0 n := by cases n <;> rfl + +theorem ofNat_mul_subNatNat (m n k : Nat) : + m * subNatNat n k = subNatNat (m * n) (m * k) := by + cases m with + | zero => simp [ofNat_zero, Int.zero_mul, Nat.zero_mul] + | succ m => cases n.lt_or_ge k with + | inl h => + have h' : succ m * n < succ m * k := Nat.mul_lt_mul_of_pos_left h (Nat.succ_pos m) + simp [subNatNat_of_lt h, subNatNat_of_lt h'] + rw [succ_pred_eq_of_pos (Nat.sub_pos_of_lt h), ← neg_ofNat_succ, Nat.mul_sub_left_distrib, + ← succ_pred_eq_of_pos (Nat.sub_pos_of_lt h')]; rfl + | inr h => + have h' : succ m * k ≤ succ m * n := Nat.mul_le_mul_left _ h + simp [subNatNat_of_le h, subNatNat_of_le h', Nat.mul_sub_left_distrib] + +theorem negOfNat_add (m n : Nat) : negOfNat m + negOfNat n = negOfNat (m + n) := by + cases m <;> cases n <;> simp [Nat.succ_add] <;> rfl + +theorem negSucc_mul_subNatNat (m n k : Nat) : + -[m+1] * subNatNat n k = subNatNat (succ m * k) (succ m * n) := by + cases n.lt_or_ge k with + | inl h => + have h' : succ m * n < succ m * k := Nat.mul_lt_mul_of_pos_left h (Nat.succ_pos m) + rw [subNatNat_of_lt h, subNatNat_of_le (Nat.le_of_lt h')] + simp [succ_pred_eq_of_pos (Nat.sub_pos_of_lt h), Nat.mul_sub_left_distrib] + | inr h => cases Nat.lt_or_ge k n with + | inl h' => + have h₁ : succ m * n > succ m * k := Nat.mul_lt_mul_of_pos_left h' (Nat.succ_pos m) + rw [subNatNat_of_le h, subNatNat_of_lt h₁, negSucc_mul_ofNat, + Nat.mul_sub_left_distrib, ← succ_pred_eq_of_pos (Nat.sub_pos_of_lt h₁)]; rfl + | inr h' => rw [Nat.le_antisymm h h', subNatNat_self, subNatNat_self, Int.mul_zero] + +attribute [local simp] ofNat_mul_subNatNat negOfNat_add negSucc_mul_subNatNat + +protected theorem mul_add : ∀ a b c : Int, a * (b + c) = a * b + a * c + | (m:Nat), (n:Nat), (k:Nat) => by simp [Nat.left_distrib] + | (m:Nat), (n:Nat), -[k+1] => by + simp [negOfNat_eq_subNatNat_zero]; rw [← subNatNat_add]; rfl + | (m:Nat), -[n+1], (k:Nat) => by + simp [negOfNat_eq_subNatNat_zero]; rw [Int.add_comm, ← subNatNat_add]; rfl + | (m:Nat), -[n+1], -[k+1] => by simp; rw [← Nat.left_distrib, succ_add]; rfl + | -[m+1], (n:Nat), (k:Nat) => by simp [Nat.mul_comm]; rw [← Nat.right_distrib, Nat.mul_comm] + | -[m+1], (n:Nat), -[k+1] => by + simp [negOfNat_eq_subNatNat_zero]; rw [Int.add_comm, ← subNatNat_add]; rfl + | -[m+1], -[n+1], (k:Nat) => by simp [negOfNat_eq_subNatNat_zero]; rw [← subNatNat_add]; rfl + | -[m+1], -[n+1], -[k+1] => by simp; rw [← Nat.left_distrib, succ_add]; rfl + +protected theorem add_mul (a b c : Int) : (a + b) * c = a * c + b * c := by + simp [Int.mul_comm, Int.mul_add] + +protected theorem neg_mul_eq_neg_mul (a b : Int) : -(a * b) = -a * b := + Int.neg_eq_of_add_eq_zero <| by rw [← Int.add_mul, Int.add_right_neg, Int.zero_mul] + +protected theorem neg_mul_eq_mul_neg (a b : Int) : -(a * b) = a * -b := + Int.neg_eq_of_add_eq_zero <| by rw [← Int.mul_add, Int.add_right_neg, Int.mul_zero] + +@[local simp] protected theorem neg_mul (a b : Int) : -a * b = -(a * b) := + (Int.neg_mul_eq_neg_mul a b).symm + +@[local simp] protected theorem mul_neg (a b : Int) : a * -b = -(a * b) := + (Int.neg_mul_eq_mul_neg a b).symm + +protected theorem neg_mul_neg (a b : Int) : -a * -b = a * b := by simp + +protected theorem neg_mul_comm (a b : Int) : -a * b = a * -b := by simp + +protected theorem mul_sub (a b c : Int) : a * (b - c) = a * b - a * c := by + simp [Int.sub_eq_add_neg, Int.mul_add] + +protected theorem sub_mul (a b c : Int) : (a - b) * c = a * c - b * c := by + simp [Int.sub_eq_add_neg, Int.add_mul] + +@[simp] protected theorem one_mul : ∀ a : Int, 1 * a = a + | ofNat n => show ofNat (1 * n) = ofNat n by rw [Nat.one_mul] + | -[n+1] => show -[1 * n +1] = -[n+1] by rw [Nat.one_mul] + +@[simp] protected theorem mul_one (a : Int) : a * 1 = a := by rw [Int.mul_comm, Int.one_mul] + +protected theorem mul_neg_one (a : Int) : a * -1 = -a := by rw [Int.mul_neg, Int.mul_one] + +protected theorem neg_eq_neg_one_mul : ∀ a : Int, -a = -1 * a + | 0 => rfl + | succ n => show _ = -[1 * n +1] by rw [Nat.one_mul]; rfl + | -[n+1] => show _ = ofNat _ by rw [Nat.one_mul]; rfl + +protected theorem mul_eq_zero {a b : Int} : a * b = 0 ↔ a = 0 ∨ b = 0 := by + refine ⟨fun h => ?_, fun h => h.elim (by simp [·, Int.zero_mul]) (by simp [·, Int.mul_zero])⟩ + exact match a, b, h with + | .ofNat 0, _, _ => by simp + | _, .ofNat 0, _ => by simp + | .ofNat (a+1), .negSucc b, h => by cases h + +protected theorem mul_ne_zero {a b : Int} (a0 : a ≠ 0) (b0 : b ≠ 0) : a * b ≠ 0 := + mt Int.mul_eq_zero.1 <| not_or.2 ⟨a0, b0⟩ + +protected theorem eq_of_mul_eq_mul_right {a b c : Int} (ha : a ≠ 0) (h : b * a = c * a) : b = c := + have : (b - c) * a = 0 := by rwa [Int.sub_mul, Int.sub_eq_zero] + Int.sub_eq_zero.1 <| (Int.mul_eq_zero.1 this).resolve_right ha + +protected theorem eq_of_mul_eq_mul_left {a b c : Int} (ha : a ≠ 0) (h : a * b = a * c) : b = c := + have : a * b - a * c = 0 := Int.sub_eq_zero_of_eq h + have : a * (b - c) = 0 := by rw [Int.mul_sub, this] + have : b - c = 0 := (Int.mul_eq_zero.1 this).resolve_left ha + Int.eq_of_sub_eq_zero this + +theorem mul_eq_mul_left_iff {a b c : Int} (h : c ≠ 0) : c * a = c * b ↔ a = b := + ⟨Int.eq_of_mul_eq_mul_left h, fun w => congrArg (fun x => c * x) w⟩ + +theorem mul_eq_mul_right_iff {a b c : Int} (h : c ≠ 0) : a * c = b * c ↔ a = b := + ⟨Int.eq_of_mul_eq_mul_right h, fun w => congrArg (fun x => x * c) w⟩ + +theorem eq_one_of_mul_eq_self_left {a b : Int} (Hpos : a ≠ 0) (H : b * a = a) : b = 1 := + Int.eq_of_mul_eq_mul_right Hpos <| by rw [Int.one_mul, H] + +theorem eq_one_of_mul_eq_self_right {a b : Int} (Hpos : b ≠ 0) (H : b * a = b) : a = 1 := + Int.eq_of_mul_eq_mul_left Hpos <| by rw [Int.mul_one, H] + +/-! +The following lemmas are later subsumed by e.g. `Nat.cast_add` and `Nat.cast_mul` in Mathlib +but it is convenient to have these earlier, for users who only need `Nat` and `Int`. +-/ + +theorem natCast_zero : ((0 : Nat) : Int) = (0 : Int) := rfl + +theorem natCast_one : ((1 : Nat) : Int) = (1 : Int) := rfl + +@[simp] theorem natCast_add (a b : Nat) : ((a + b : Nat) : Int) = (a : Int) + (b : Int) := by + -- Note this only works because of local simp attributes in this file, + -- so it still makes sense to tag the lemmas with `@[simp]`. + simp + +@[simp] theorem natCast_mul (a b : Nat) : ((a * b : Nat) : Int) = (a : Int) * (b : Int) := by + simp diff --git a/Std/Data/Int/Init/Order.lean b/Std/Data/Int/Init/Order.lean new file mode 100644 index 0000000000..ff7c8ce21a --- /dev/null +++ b/Std/Data/Int/Init/Order.lean @@ -0,0 +1,433 @@ +/- +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 Std.Data.Int.Init.Lemmas + +/-! +# Results about the order properties of the integers, and the integers as an ordered ring. +-/ + +open Nat + +namespace Int + +/-! ## Order properties of the integers -/ + +theorem nonneg_def {a : Int} : NonNeg a ↔ ∃ n : Nat, a = n := + ⟨fun ⟨n⟩ => ⟨n, rfl⟩, fun h => match a, h with | _, ⟨n, rfl⟩ => ⟨n⟩⟩ + +theorem NonNeg.elim {a : Int} : NonNeg a → ∃ n : Nat, a = n := nonneg_def.1 + +theorem nonneg_or_nonneg_neg : ∀ (a : Int), NonNeg a ∨ NonNeg (-a) + | (_:Nat) => .inl ⟨_⟩ + | -[_+1] => .inr ⟨_⟩ + +theorem le_def (a b : Int) : a ≤ b ↔ NonNeg (b - a) := .rfl + +theorem lt_iff_add_one_le (a b : Int) : a < b ↔ a + 1 ≤ b := .rfl + +theorem le.intro_sub {a b : Int} (n : Nat) (h : b - a = n) : a ≤ b := by + simp [le_def, h]; constructor + +attribute [local simp] Int.add_left_neg Int.add_right_neg Int.neg_add + +theorem le.intro {a b : Int} (n : Nat) (h : a + n = b) : a ≤ b := + le.intro_sub n <| by rw [← h, Int.add_comm]; simp [Int.sub_eq_add_neg, Int.add_assoc] + +theorem le.dest_sub {a b : Int} (h : a ≤ b) : ∃ n : Nat, b - a = n := nonneg_def.1 h + +theorem le.dest {a b : Int} (h : a ≤ b) : ∃ n : Nat, a + n = b := + let ⟨n, h₁⟩ := le.dest_sub h + ⟨n, by rw [← h₁, Int.add_comm]; simp [Int.sub_eq_add_neg, Int.add_assoc]⟩ + +protected theorem le_total (a b : Int) : a ≤ b ∨ b ≤ a := + (nonneg_or_nonneg_neg (b - a)).imp_right fun H => by + rwa [show -(b - a) = a - b by simp [Int.add_comm, Int.sub_eq_add_neg]] at H + +@[simp, norm_cast] theorem ofNat_le {m n : Nat} : (↑m : Int) ≤ ↑n ↔ m ≤ n := + ⟨fun h => + let ⟨k, hk⟩ := le.dest h + Nat.le.intro <| Int.ofNat.inj <| (Int.ofNat_add m k).trans hk, + fun h => + let ⟨k, (hk : m + k = n)⟩ := Nat.le.dest h + le.intro k (by rw [← hk]; rfl)⟩ + +theorem ofNat_zero_le (n : Nat) : 0 ≤ (↑n : Int) := ofNat_le.2 n.zero_le + +theorem eq_ofNat_of_zero_le {a : Int} (h : 0 ≤ a) : ∃ n : Nat, a = n := by + have t := le.dest_sub h; rwa [Int.sub_zero] at t + +theorem eq_succ_of_zero_lt {a : Int} (h : 0 < a) : ∃ n : Nat, a = n.succ := + let ⟨n, (h : ↑(1 + n) = a)⟩ := le.dest h + ⟨n, by rw [Nat.add_comm] at h; exact h.symm⟩ + +theorem lt_add_succ (a : Int) (n : Nat) : a < a + Nat.succ n := + le.intro n <| by rw [Int.add_comm, Int.add_left_comm]; rfl + +theorem lt.intro {a b : Int} {n : Nat} (h : a + Nat.succ n = b) : a < b := + h ▸ lt_add_succ a n + +theorem lt.dest {a b : Int} (h : a < b) : ∃ n : Nat, a + Nat.succ n = b := + let ⟨n, h⟩ := le.dest h; ⟨n, by rwa [Int.add_comm, Int.add_left_comm] at h⟩ + +@[simp, norm_cast] theorem ofNat_lt {n m : Nat} : (↑n : Int) < ↑m ↔ n < m := by + rw [lt_iff_add_one_le, ← ofNat_succ, ofNat_le]; rfl + +@[simp, norm_cast] theorem ofNat_pos {n : Nat} : 0 < (↑n : Int) ↔ 0 < n := ofNat_lt + +theorem ofNat_nonneg (n : Nat) : 0 ≤ (n : Int) := ⟨_⟩ + +theorem ofNat_succ_pos (n : Nat) : 0 < (succ n : Int) := ofNat_lt.2 <| Nat.succ_pos _ + +@[simp] protected theorem le_refl (a : Int) : a ≤ a := + le.intro _ (Int.add_zero a) + +protected theorem le_trans {a b c : Int} (h₁ : a ≤ b) (h₂ : b ≤ c) : a ≤ c := + let ⟨n, hn⟩ := le.dest h₁; let ⟨m, hm⟩ := le.dest h₂ + le.intro (n + m) <| by rw [← hm, ← hn, Int.add_assoc, ofNat_add] + +protected theorem le_antisymm {a b : Int} (h₁ : a ≤ b) (h₂ : b ≤ a) : a = b := by + let ⟨n, hn⟩ := le.dest h₁; let ⟨m, hm⟩ := le.dest h₂ + have := hn; rw [← hm, Int.add_assoc, ← ofNat_add] at this + have := Int.ofNat.inj <| Int.add_left_cancel <| this.trans (Int.add_zero _).symm + rw [← hn, Nat.eq_zero_of_add_eq_zero_left this, ofNat_zero, Int.add_zero a] + +protected theorem lt_irrefl (a : Int) : ¬a < a := fun H => + let ⟨n, hn⟩ := lt.dest H + have : (a+Nat.succ n) = a+0 := by + rw [hn, Int.add_zero] + have : Nat.succ n = 0 := Int.ofNat.inj (Int.add_left_cancel this) + show False from Nat.succ_ne_zero _ this + +protected theorem ne_of_lt {a b : Int} (h : a < b) : a ≠ b := fun e => by + cases e; exact Int.lt_irrefl _ h + +protected theorem ne_of_gt {a b : Int} (h : b < a) : a ≠ b := (Int.ne_of_lt h).symm + +protected theorem le_of_lt {a b : Int} (h : a < b) : a ≤ b := + let ⟨_, hn⟩ := lt.dest h; le.intro _ hn + +protected theorem lt_iff_le_and_ne {a b : Int} : a < b ↔ a ≤ b ∧ a ≠ b := by + refine ⟨fun h => ⟨Int.le_of_lt h, Int.ne_of_lt h⟩, fun ⟨aleb, aneb⟩ => ?_⟩ + let ⟨n, hn⟩ := le.dest aleb + have : n ≠ 0 := aneb.imp fun eq => by rw [← hn, eq, ofNat_zero, Int.add_zero] + apply lt.intro; rwa [← Nat.succ_pred_eq_of_pos (Nat.pos_of_ne_zero this)] at hn + +theorem lt_succ (a : Int) : a < a + 1 := Int.le_refl _ + +protected theorem zero_lt_one : (0 : Int) < 1 := ⟨_⟩ + +protected theorem lt_iff_le_not_le {a b : Int} : a < b ↔ a ≤ b ∧ ¬b ≤ a := by + rw [Int.lt_iff_le_and_ne] + constructor <;> refine fun ⟨h, h'⟩ => ⟨h, h'.imp fun h' => ?_⟩ + · exact Int.le_antisymm h h' + · subst h'; apply Int.le_refl + +protected theorem not_le {a b : Int} : ¬a ≤ b ↔ b < a := + ⟨fun h => Int.lt_iff_le_not_le.2 ⟨(Int.le_total ..).resolve_right h, h⟩, + fun h => (Int.lt_iff_le_not_le.1 h).2⟩ + +protected theorem not_lt {a b : Int} : ¬a < b ↔ b ≤ a := + by rw [← Int.not_le, Decidable.not_not] + +protected theorem lt_trichotomy (a b : Int) : a < b ∨ a = b ∨ b < a := + if eq : a = b then .inr <| .inl eq else + if le : a ≤ b then .inl <| Int.lt_iff_le_and_ne.2 ⟨le, eq⟩ else + .inr <| .inr <| Int.not_le.1 le + +protected theorem ne_iff_lt_or_gt {a b : Int} : a ≠ b ↔ a < b ∨ b < a := by + constructor + · intro h + cases Int.lt_trichotomy a b + case inl lt => exact Or.inl lt + case inr h => + cases h + case inl =>simp_all + case inr gt => exact Or.inr gt + · intro h + cases h + case inl lt => exact Int.ne_of_lt lt + case inr gt => exact Int.ne_of_gt gt + +protected alias ⟨lt_or_gt_of_ne, _⟩ := Int.ne_iff_lt_or_gt + +protected theorem eq_iff_le_and_ge {x y : Int} : x = y ↔ x ≤ y ∧ y ≤ x := by + constructor + · simp_all + · intro ⟨h₁, h₂⟩ + exact Int.le_antisymm h₁ h₂ + +protected theorem lt_of_le_of_lt {a b c : Int} (h₁ : a ≤ b) (h₂ : b < c) : a < c := + Int.not_le.1 fun h => Int.not_le.2 h₂ (Int.le_trans h h₁) + +protected theorem lt_of_lt_of_le {a b c : Int} (h₁ : a < b) (h₂ : b ≤ c) : a < c := + Int.not_le.1 fun h => Int.not_le.2 h₁ (Int.le_trans h₂ h) + +protected theorem lt_trans {a b c : Int} (h₁ : a < b) (h₂ : b < c) : a < c := + Int.lt_of_le_of_lt (Int.le_of_lt h₁) h₂ + +instance : Trans (α := Int) (· ≤ ·) (· ≤ ·) (· ≤ ·) := ⟨Int.le_trans⟩ + +instance : Trans (α := Int) (· < ·) (· ≤ ·) (· < ·) := ⟨Int.lt_of_lt_of_le⟩ + +instance : Trans (α := Int) (· ≤ ·) (· < ·) (· < ·) := ⟨Int.lt_of_le_of_lt⟩ + +instance : Trans (α := Int) (· < ·) (· < ·) (· < ·) := ⟨Int.lt_trans⟩ + +protected theorem min_def (n m : Int) : min n m = if n ≤ m then n else m := rfl + +protected theorem max_def (n m : Int) : max n m = if n ≤ m then m else n := rfl + +protected theorem min_comm (a b : Int) : min a b = min b a := by + simp [Int.min_def] + by_cases h₁ : a ≤ b <;> by_cases h₂ : b ≤ a <;> simp [h₁, h₂] + · exact Int.le_antisymm h₁ h₂ + · cases not_or_intro h₁ h₂ <| Int.le_total .. + +protected theorem min_le_right (a b : Int) : min a b ≤ b := by rw [Int.min_def]; split <;> simp [*] + +protected theorem min_le_left (a b : Int) : min a b ≤ a := Int.min_comm .. ▸ Int.min_le_right .. + +protected theorem le_min {a b c : Int} : a ≤ min b c ↔ a ≤ b ∧ a ≤ c := + ⟨fun h => ⟨Int.le_trans h (Int.min_le_left ..), Int.le_trans h (Int.min_le_right ..)⟩, + fun ⟨h₁, h₂⟩ => by rw [Int.min_def]; split <;> assumption⟩ + +protected theorem max_comm (a b : Int) : max a b = max b a := by + simp only [Int.max_def] + by_cases h₁ : a ≤ b <;> by_cases h₂ : b ≤ a <;> simp [h₁, h₂] + · exact Int.le_antisymm h₂ h₁ + · cases not_or_intro h₁ h₂ <| Int.le_total .. + +protected theorem le_max_left (a b : Int) : a ≤ max a b := by rw [Int.max_def]; split <;> simp [*] + +protected theorem le_max_right (a b : Int) : b ≤ max a b := Int.max_comm .. ▸ Int.le_max_left .. + +protected theorem max_le {a b c : Int} : max a b ≤ c ↔ a ≤ c ∧ b ≤ c := + ⟨fun h => ⟨Int.le_trans (Int.le_max_left ..) h, Int.le_trans (Int.le_max_right ..) h⟩, + fun ⟨h₁, h₂⟩ => by rw [Int.max_def]; split <;> assumption⟩ + +theorem eq_natAbs_of_zero_le {a : Int} (h : 0 ≤ a) : a = natAbs a := by + let ⟨n, e⟩ := eq_ofNat_of_zero_le h + rw [e]; rfl + +theorem le_natAbs {a : Int} : a ≤ natAbs a := + match Int.le_total 0 a with + | .inl h => by rw [eq_natAbs_of_zero_le h]; apply Int.le_refl + | .inr h => Int.le_trans h (ofNat_zero_le _) + +theorem negSucc_lt_zero (n : Nat) : -[n+1] < 0 := + Int.not_le.1 fun h => let ⟨_, h⟩ := eq_ofNat_of_zero_le h; nomatch h + +@[simp] theorem negSucc_not_nonneg (n : Nat) : 0 ≤ -[n+1] ↔ False := by + simp only [Int.not_le, iff_false]; exact Int.negSucc_lt_zero n + +protected theorem add_le_add_left {a b : Int} (h : a ≤ b) (c : Int) : c + a ≤ c + b := + let ⟨n, hn⟩ := le.dest h; le.intro n <| by rw [Int.add_assoc, hn] + +protected theorem add_lt_add_left {a b : Int} (h : a < b) (c : Int) : c + a < c + b := + Int.lt_iff_le_and_ne.2 ⟨Int.add_le_add_left (Int.le_of_lt h) _, fun heq => + b.lt_irrefl <| by rwa [Int.add_left_cancel heq] at h⟩ + +protected theorem add_le_add_right {a b : Int} (h : a ≤ b) (c : Int) : a + c ≤ b + c := + Int.add_comm c a ▸ Int.add_comm c b ▸ Int.add_le_add_left h c + +protected theorem add_lt_add_right {a b : Int} (h : a < b) (c : Int) : a + c < b + c := + Int.add_comm c a ▸ Int.add_comm c b ▸ Int.add_lt_add_left h c + +protected theorem le_of_add_le_add_left {a b c : Int} (h : a + b ≤ a + c) : b ≤ c := by + have : -a + (a + b) ≤ -a + (a + c) := Int.add_le_add_left h _ + simp [Int.neg_add_cancel_left] at this + assumption + +protected theorem le_of_add_le_add_right {a b c : Int} (h : a + b ≤ c + b) : a ≤ c := + Int.le_of_add_le_add_left (a := b) <| by rwa [Int.add_comm b a, Int.add_comm b c] + +protected theorem add_le_add_iff_left (a : Int) : a + b ≤ a + c ↔ b ≤ c := + ⟨Int.le_of_add_le_add_left, (Int.add_le_add_left · _)⟩ + +protected theorem add_le_add_iff_right (c : Int) : a + c ≤ b + c ↔ a ≤ b := + ⟨Int.le_of_add_le_add_right, (Int.add_le_add_right · _)⟩ + +protected theorem add_le_add {a b c d : Int} (h₁ : a ≤ b) (h₂ : c ≤ d) : a + c ≤ b + d := + Int.le_trans (Int.add_le_add_right h₁ c) (Int.add_le_add_left h₂ b) + +protected theorem le_add_of_nonneg_right {a b : Int} (h : 0 ≤ b) : a ≤ a + b := by + have : a + b ≥ a + 0 := Int.add_le_add_left h a + rwa [Int.add_zero] at this + +protected theorem le_add_of_nonneg_left {a b : Int} (h : 0 ≤ b) : a ≤ b + a := by + have : 0 + a ≤ b + a := Int.add_le_add_right h a + rwa [Int.zero_add] at this + +protected theorem neg_le_neg {a b : Int} (h : a ≤ b) : -b ≤ -a := by + have : 0 ≤ -a + b := Int.add_left_neg a ▸ Int.add_le_add_left h (-a) + have : 0 + -b ≤ -a + b + -b := Int.add_le_add_right this (-b) + rwa [Int.add_neg_cancel_right, Int.zero_add] at this + +protected theorem le_of_neg_le_neg {a b : Int} (h : -b ≤ -a) : a ≤ b := + suffices - -a ≤ - -b by simp [Int.neg_neg] at this; assumption + Int.neg_le_neg h + +protected theorem neg_nonpos_of_nonneg {a : Int} (h : 0 ≤ a) : -a ≤ 0 := by + have : -a ≤ -0 := Int.neg_le_neg h + rwa [Int.neg_zero] at this + +protected theorem neg_nonneg_of_nonpos {a : Int} (h : a ≤ 0) : 0 ≤ -a := by + have : -0 ≤ -a := Int.neg_le_neg h + rwa [Int.neg_zero] at this + +protected theorem neg_lt_neg {a b : Int} (h : a < b) : -b < -a := by + have : 0 < -a + b := Int.add_left_neg a ▸ Int.add_lt_add_left h (-a) + have : 0 + -b < -a + b + -b := Int.add_lt_add_right this (-b) + rwa [Int.add_neg_cancel_right, Int.zero_add] at this + +protected theorem neg_neg_of_pos {a : Int} (h : 0 < a) : -a < 0 := by + have : -a < -0 := Int.neg_lt_neg h + rwa [Int.neg_zero] at this + +protected theorem neg_pos_of_neg {a : Int} (h : a < 0) : 0 < -a := by + have : -0 < -a := Int.neg_lt_neg h + rwa [Int.neg_zero] at this + +protected theorem sub_nonneg_of_le {a b : Int} (h : b ≤ a) : 0 ≤ a - b := by + have h := Int.add_le_add_right h (-b) + rwa [Int.add_right_neg] at h + +protected theorem le_of_sub_nonneg {a b : Int} (h : 0 ≤ a - b) : b ≤ a := by + have h := Int.add_le_add_right h b + rwa [Int.sub_add_cancel, Int.zero_add] at h + +protected theorem sub_pos_of_lt {a b : Int} (h : b < a) : 0 < a - b := by + have h := Int.add_lt_add_right h (-b) + rwa [Int.add_right_neg] at h + +protected theorem lt_of_sub_pos {a b : Int} (h : 0 < a - b) : b < a := by + have h := Int.add_lt_add_right h b + rwa [Int.sub_add_cancel, Int.zero_add] at h + +protected theorem sub_left_le_of_le_add {a b c : Int} (h : a ≤ b + c) : a - b ≤ c := by + have h := Int.add_le_add_right h (-b) + rwa [Int.add_comm b c, Int.add_neg_cancel_right] at h + +protected theorem sub_le_self (a : Int) {b : Int} (h : 0 ≤ b) : a - b ≤ a := + calc a + -b + _ ≤ a + 0 := Int.add_le_add_left (Int.neg_nonpos_of_nonneg h) _ + _ = a := by rw [Int.add_zero] + +protected theorem sub_lt_self (a : Int) {b : Int} (h : 0 < b) : a - b < a := + calc a + -b + _ < a + 0 := Int.add_lt_add_left (Int.neg_neg_of_pos h) _ + _ = a := by rw [Int.add_zero] + +theorem add_one_le_of_lt {a b : Int} (H : a < b) : a + 1 ≤ b := H + +/- ### Order properties and multiplication -/ + + +protected theorem mul_nonneg {a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a * b := by + let ⟨n, hn⟩ := eq_ofNat_of_zero_le ha + let ⟨m, hm⟩ := eq_ofNat_of_zero_le hb + rw [hn, hm, ← ofNat_mul]; apply ofNat_nonneg + +protected theorem mul_pos {a b : Int} (ha : 0 < a) (hb : 0 < b) : 0 < a * b := by + let ⟨n, hn⟩ := eq_succ_of_zero_lt ha + let ⟨m, hm⟩ := eq_succ_of_zero_lt hb + rw [hn, hm, ← ofNat_mul]; apply ofNat_succ_pos + +protected theorem mul_lt_mul_of_pos_left {a b c : Int} + (h₁ : a < b) (h₂ : 0 < c) : c * a < c * b := by + have : 0 < c * (b - a) := Int.mul_pos h₂ (Int.sub_pos_of_lt h₁) + rw [Int.mul_sub] at this + exact Int.lt_of_sub_pos this + +protected theorem mul_lt_mul_of_pos_right {a b c : Int} + (h₁ : a < b) (h₂ : 0 < c) : a * c < b * c := by + have : 0 < b - a := Int.sub_pos_of_lt h₁ + have : 0 < (b - a) * c := Int.mul_pos this h₂ + rw [Int.sub_mul] at this + exact Int.lt_of_sub_pos this + +protected theorem mul_le_mul_of_nonneg_left {a b c : Int} + (h₁ : a ≤ b) (h₂ : 0 ≤ c) : c * a ≤ c * b := by + if hba : b ≤ a then rw [Int.le_antisymm hba h₁]; apply Int.le_refl else + if hc0 : c ≤ 0 then simp [Int.le_antisymm hc0 h₂, Int.zero_mul] else + exact Int.le_of_lt <| Int.mul_lt_mul_of_pos_left + (Int.lt_iff_le_not_le.2 ⟨h₁, hba⟩) (Int.lt_iff_le_not_le.2 ⟨h₂, hc0⟩) + +protected theorem mul_le_mul_of_nonneg_right {a b c : Int} + (h₁ : a ≤ b) (h₂ : 0 ≤ c) : a * c ≤ b * c := by + rw [Int.mul_comm, Int.mul_comm b]; exact Int.mul_le_mul_of_nonneg_left h₁ h₂ + +protected theorem mul_le_mul {a b c d : Int} + (hac : a ≤ c) (hbd : b ≤ d) (nn_b : 0 ≤ b) (nn_c : 0 ≤ c) : a * b ≤ c * d := + Int.le_trans (Int.mul_le_mul_of_nonneg_right hac nn_b) (Int.mul_le_mul_of_nonneg_left hbd nn_c) + +protected theorem mul_nonpos_of_nonneg_of_nonpos {a b : Int} + (ha : 0 ≤ a) (hb : b ≤ 0) : a * b ≤ 0 := by + have h : a * b ≤ a * 0 := Int.mul_le_mul_of_nonneg_left hb ha + rwa [Int.mul_zero] at h + +protected theorem mul_nonpos_of_nonpos_of_nonneg {a b : Int} + (ha : a ≤ 0) (hb : 0 ≤ b) : a * b ≤ 0 := by + have h : a * b ≤ 0 * b := Int.mul_le_mul_of_nonneg_right ha hb + rwa [Int.zero_mul] at h + +protected theorem mul_le_mul_of_nonpos_right {a b c : Int} + (h : b ≤ a) (hc : c ≤ 0) : a * c ≤ b * c := + have : -c ≥ 0 := Int.neg_nonneg_of_nonpos hc + have : b * -c ≤ a * -c := Int.mul_le_mul_of_nonneg_right h this + Int.le_of_neg_le_neg <| by rwa [← Int.neg_mul_eq_mul_neg, ← Int.neg_mul_eq_mul_neg] at this + +protected theorem mul_le_mul_of_nonpos_left {a b c : Int} + (ha : a ≤ 0) (h : c ≤ b) : a * b ≤ a * c := by + rw [Int.mul_comm a b, Int.mul_comm a c] + apply Int.mul_le_mul_of_nonpos_right h ha + +/- ## natAbs -/ + +@[simp] theorem natAbs_ofNat (n : Nat) : natAbs ↑n = n := rfl +@[simp] theorem natAbs_negSucc (n : Nat) : natAbs -[n+1] = n.succ := rfl +@[simp] theorem natAbs_zero : natAbs (0 : Int) = (0 : Nat) := rfl +@[simp] theorem natAbs_one : natAbs (1 : Int) = (1 : Nat) := rfl + +@[simp] theorem natAbs_eq_zero : natAbs a = 0 ↔ a = 0 := + ⟨fun H => match a with + | ofNat _ => congrArg ofNat H + | -[_+1] => absurd H (succ_ne_zero _), + fun e => e ▸ rfl⟩ + +theorem natAbs_pos : 0 < natAbs a ↔ a ≠ 0 := by rw [Nat.pos_iff_ne_zero, Ne, natAbs_eq_zero] + +@[simp] theorem natAbs_neg : ∀ (a : Int), natAbs (-a) = natAbs a + | 0 => rfl + | succ _ => rfl + | -[_+1] => rfl + +theorem natAbs_eq : ∀ (a : Int), a = natAbs a ∨ a = -↑(natAbs a) + | ofNat _ => Or.inl rfl + | -[_+1] => Or.inr rfl + +theorem natAbs_negOfNat (n : Nat) : natAbs (negOfNat n) = n := by + cases n <;> rfl + +theorem natAbs_mul (a b : Int) : natAbs (a * b) = natAbs a * natAbs b := by + cases a <;> cases b <;> + simp only [← Int.mul_def, Int.mul, natAbs_negOfNat] <;> simp only [natAbs] + +theorem natAbs_eq_natAbs_iff {a b : Int} : a.natAbs = b.natAbs ↔ a = b ∨ a = -b := by + constructor <;> intro h + · cases Int.natAbs_eq a with + | inl h₁ | inr h₁ => + cases Int.natAbs_eq b with + | inl h₂ | inr h₂ => rw [h₁, h₂]; simp [h] + · cases h with (subst a; try rfl) + | inr h => rw [Int.natAbs_neg] + +theorem natAbs_of_nonneg {a : Int} (H : 0 ≤ a) : (natAbs a : Int) = a := + match a, eq_ofNat_of_zero_le H with + | _, ⟨_, rfl⟩ => rfl + +theorem ofNat_natAbs_of_nonpos {a : Int} (H : a ≤ 0) : (natAbs a : Int) = -a := by + rw [← natAbs_neg, natAbs_of_nonneg (Int.neg_nonneg_of_nonpos H)] diff --git a/Std/Data/Int/Lemmas.lean b/Std/Data/Int/Lemmas.lean deleted file mode 100644 index 85ac931b92..0000000000 --- a/Std/Data/Int/Lemmas.lean +++ /dev/null @@ -1,1492 +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 Std.Classes.Cast -import Std.Data.Nat.Lemmas -import Std.Data.Int.Basic -import Std.Data.Option.Basic -import Std.Tactic.NormCast.Lemmas -import Std.Tactic.RCases - -open Nat - -namespace Int - -@[simp] theorem ofNat_eq_coe : ofNat n = Nat.cast n := rfl - -@[simp] theorem ofNat_zero : ((0 : Nat) : Int) = 0 := rfl - -@[simp] theorem ofNat_one : ((1 : Nat) : Int) = 1 := rfl - -theorem ofNat_two : ((2 : Nat) : Int) = 2 := rfl - -@[simp] theorem default_eq_zero : default = (0 : Int) := rfl - -/- ## Definitions of basic functions -/ - -theorem subNatNat_of_sub_eq_zero {m n : Nat} (h : n - m = 0) : subNatNat m n = ↑(m - n) := by - rw [subNatNat, h, ofNat_eq_coe] - -theorem subNatNat_of_sub_eq_succ {m n k : Nat} (h : n - m = succ k) : subNatNat m n = -[k+1] := by - rw [subNatNat, h] - -@[simp] protected theorem neg_zero : -(0:Int) = 0 := rfl - -@[norm_cast] theorem ofNat_add (n m : Nat) : (↑(n + m) : Int) = n + m := rfl -@[norm_cast] theorem ofNat_mul (n m : Nat) : (↑(n * m) : Int) = n * m := rfl -theorem ofNat_succ (n : Nat) : (succ n : Int) = n + 1 := rfl - -@[local simp] theorem neg_ofNat_zero : -((0 : Nat) : Int) = 0 := rfl -@[local simp] theorem neg_ofNat_succ (n : Nat) : -(succ n : Int) = -[n+1] := rfl -@[local simp] theorem neg_negSucc (n : Nat) : -(-[n+1]) = succ n := rfl - -theorem negSucc_coe (n : Nat) : -[n+1] = -↑(n + 1) := rfl - -theorem negOfNat_eq : negOfNat n = -ofNat n := rfl - -/- ## These are only for internal use -/ - -@[simp] theorem add_def {a b : Int} : Int.add a b = a + b := rfl - -@[local simp] theorem ofNat_add_ofNat (m n : Nat) : (↑m + ↑n : Int) = ↑(m + n) := rfl -@[local simp] theorem ofNat_add_negSucc (m n : Nat) : ↑m + -[n+1] = subNatNat m (succ n) := rfl -@[local simp] theorem negSucc_add_ofNat (m n : Nat) : -[m+1] + ↑n = subNatNat n (succ m) := rfl -@[local simp] theorem negSucc_add_negSucc (m n : Nat) : -[m+1] + -[n+1] = -[succ (m + n) +1] := rfl - -@[simp] theorem mul_def {a b : Int} : Int.mul a b = a * b := rfl - -@[local simp] theorem ofNat_mul_ofNat (m n : Nat) : (↑m * ↑n : Int) = ↑(m * n) := rfl -@[local simp] theorem ofNat_mul_negSucc' (m n : Nat) : ↑m * -[n+1] = negOfNat (m * succ n) := rfl -@[local simp] theorem negSucc_mul_ofNat' (m n : Nat) : -[m+1] * ↑n = negOfNat (succ m * n) := rfl -@[local simp] theorem negSucc_mul_negSucc' (m n : Nat) : - -[m+1] * -[n+1] = ofNat (succ m * succ n) := rfl - -/- ## some basic functions and properties -/ - -@[norm_cast] theorem ofNat_inj : ((m : Nat) : Int) = (n : Nat) ↔ m = n := ⟨ofNat.inj, congrArg _⟩ - -theorem ofNat_eq_zero : ((n : Nat) : Int) = 0 ↔ n = 0 := ofNat_inj - -theorem ofNat_ne_zero : ((n : Nat) : Int) ≠ 0 ↔ n ≠ 0 := not_congr ofNat_eq_zero - -theorem negSucc_inj : negSucc m = negSucc n ↔ m = n := ⟨negSucc.inj, fun H => by simp [H]⟩ - -theorem negSucc_eq (n : Nat) : -[n+1] = -((n : Int) + 1) := rfl - -@[simp] theorem negSucc_ne_zero (n : Nat) : -[n+1] ≠ 0 := fun. - -@[simp] theorem zero_ne_negSucc (n : Nat) : 0 ≠ -[n+1] := fun. - -@[simp, norm_cast] theorem Nat.cast_ofNat_Int : - (Nat.cast (no_index (OfNat.ofNat n)) : Int) = OfNat.ofNat n := rfl - -/- ## neg -/ - -@[simp] protected theorem neg_neg : ∀ a : Int, -(-a) = a - | 0 => rfl - | succ _ => rfl - | -[_+1] => rfl - -protected theorem neg_inj {a b : Int} : -a = -b ↔ a = b := - ⟨fun h => by rw [← Int.neg_neg a, ← Int.neg_neg b, h], congrArg _⟩ - -@[simp] protected theorem neg_eq_zero : -a = 0 ↔ a = 0 := Int.neg_inj (b := 0) - -protected theorem neg_ne_zero : -a ≠ 0 ↔ a ≠ 0 := not_congr Int.neg_eq_zero - -protected theorem sub_eq_add_neg {a b : Int} : a - b = a + -b := rfl - -theorem add_neg_one (i : Int) : i + -1 = i - 1 := rfl - -/- ## basic properties of subNatNat -/ - --- @[elabAsElim] -- TODO(Mario): unexpected eliminator resulting type -theorem subNatNat_elim (m n : Nat) (motive : Nat → Nat → Int → Prop) - (hp : ∀ i n, motive (n + i) n i) - (hn : ∀ i m, motive m (m + i + 1) -[i+1]) : - motive m n (subNatNat m n) := by - unfold subNatNat - match h : n - m with - | 0 => - have ⟨k, h⟩ := Nat.le.dest (Nat.le_of_sub_eq_zero h) - rw [h.symm, Nat.add_sub_cancel_left]; apply hp - | succ k => - rw [Nat.sub_eq_iff_eq_add (Nat.le_of_lt (Nat.lt_of_sub_eq_succ h))] at h - rw [h, Nat.add_comm]; apply hn - -theorem subNatNat_add_left : subNatNat (m + n) m = n := by - unfold subNatNat - rw [Nat.sub_eq_zero_of_le (Nat.le_add_right ..), Nat.add_sub_cancel_left, ofNat_eq_coe] - -theorem subNatNat_add_right : subNatNat m (m + n + 1) = negSucc n := by - simp [subNatNat, Nat.add_assoc, Nat.add_sub_cancel_left] - -theorem subNatNat_add_add (m n k : Nat) : subNatNat (m + k) (n + k) = subNatNat m n := by - apply subNatNat_elim m n (fun m n i => subNatNat (m + k) (n + k) = i) - · intro i j - rw [Nat.add_assoc, Nat.add_comm i k, ← Nat.add_assoc] - exact subNatNat_add_left - · intro i j - rw [Nat.add_assoc j i 1, Nat.add_comm j (i+1), Nat.add_assoc, Nat.add_comm (i+1) (j+k)] - exact subNatNat_add_right - -theorem subNatNat_of_le {m n : Nat} (h : n ≤ m) : subNatNat m n = ↑(m - n) := - subNatNat_of_sub_eq_zero (Nat.sub_eq_zero_of_le h) - -theorem subNatNat_of_lt {m n : Nat} (h : m < n) : subNatNat m n = -[pred (n - m) +1] := - subNatNat_of_sub_eq_succ <| (Nat.succ_pred_eq_of_pos (Nat.sub_pos_of_lt h)).symm - -/- ## natAbs -/ - -@[simp] theorem natAbs_ofNat (n : Nat) : natAbs ↑n = n := rfl -@[simp] theorem natAbs_negSucc (n : Nat) : natAbs -[n+1] = n.succ := rfl -@[simp] theorem natAbs_zero : natAbs (0 : Int) = (0 : Nat) := rfl -@[simp] theorem natAbs_one : natAbs (1 : Int) = (1 : Nat) := rfl - -@[simp] theorem natAbs_eq_zero : natAbs a = 0 ↔ a = 0 := - ⟨fun H => match a with - | ofNat _ => congrArg ofNat H - | -[_+1] => absurd H (succ_ne_zero _), - fun e => e ▸ rfl⟩ - -theorem natAbs_ne_zero {a : Int} : a.natAbs ≠ 0 ↔ a ≠ 0 := not_congr Int.natAbs_eq_zero - -theorem natAbs_pos : 0 < natAbs a ↔ a ≠ 0 := by rw [Nat.pos_iff_ne_zero, Ne, natAbs_eq_zero] - -theorem natAbs_mul_self : ∀ {a : Int}, ↑(natAbs a * natAbs a) = a * a - | ofNat _ => rfl - | -[_+1] => rfl - -@[simp] theorem natAbs_neg : ∀ (a : Int), natAbs (-a) = natAbs a - | 0 => rfl - | succ _ => rfl - | -[_+1] => rfl - -theorem natAbs_eq : ∀ (a : Int), a = natAbs a ∨ a = -↑(natAbs a) - | ofNat _ => Or.inl rfl - | -[_+1] => Or.inr rfl - -theorem eq_nat_or_neg (a : Int) : ∃ n : Nat, a = n ∨ a = -↑n := ⟨_, natAbs_eq a⟩ - -theorem natAbs_negOfNat (n : Nat) : natAbs (negOfNat n) = n := by - cases n <;> rfl - -theorem natAbs_mul (a b : Int) : natAbs (a * b) = natAbs a * natAbs b := by - cases a <;> cases b <;> - simp only [← Int.mul_def, Int.mul, natAbs_negOfNat] <;> simp only [natAbs] - -theorem natAbs_mul_natAbs_eq {a b : Int} {c : Nat} - (h : a * b = (c : Int)) : a.natAbs * b.natAbs = c := by rw [← natAbs_mul, h, natAbs] - -@[simp] theorem natAbs_mul_self' (a : Int) : (natAbs a * natAbs a : Int) = a * a := by - rw [← Int.ofNat_mul, natAbs_mul_self] - -theorem natAbs_eq_natAbs_iff {a b : Int} : a.natAbs = b.natAbs ↔ a = b ∨ a = -b := by - constructor <;> intro h - · cases Int.natAbs_eq a with - | inl h₁ | inr h₁ => - cases Int.natAbs_eq b with - | inl h₂ | inr h₂ => rw [h₁, h₂]; simp [h] - · cases h with (subst a; try rfl) - | inr h => rw [Int.natAbs_neg] - -theorem natAbs_eq_iff {a : Int} {n : Nat} : a.natAbs = n ↔ a = n ∨ a = -↑n := by - rw [← Int.natAbs_eq_natAbs_iff, Int.natAbs_ofNat] - - -/- # ring properties -/ - -/- addition -/ - -protected theorem add_comm : ∀ a b : Int, a + b = b + a - | ofNat n, ofNat m => by simp [Nat.add_comm] - | ofNat _, -[_+1] => rfl - | -[_+1], ofNat _ => rfl - | -[_+1], -[_+1] => by simp [Nat.add_comm] - -@[simp] protected theorem add_zero : ∀ a : Int, a + 0 = a - | ofNat _ => rfl - | -[_+1] => rfl - -@[simp] protected theorem zero_add (a : Int) : 0 + a = a := Int.add_comm .. ▸ a.add_zero - -theorem ofNat_add_negSucc_of_lt (h : m < n.succ) : ofNat m + -[n+1] = -[n - m+1] := - show subNatNat .. = _ by simp [succ_sub (le_of_lt_succ h), subNatNat] - -theorem subNatNat_sub (h : n ≤ m) (k : Nat) : subNatNat (m - n) k = subNatNat m (k + n) := by - rwa [← subNatNat_add_add _ _ n, Nat.sub_add_cancel] - -theorem subNatNat_add (m n k : Nat) : subNatNat (m + n) k = m + subNatNat n k := by - cases n.lt_or_ge k with - | inl h' => - simp [subNatNat_of_lt h', succ_pred_eq_of_pos (Nat.sub_pos_of_lt h')] - conv => lhs; rw [← Nat.sub_add_cancel (Nat.le_of_lt h')] - apply subNatNat_add_add - | inr h' => simp [subNatNat_of_le h', - subNatNat_of_le (Nat.le_trans h' (le_add_left ..)), Nat.add_sub_assoc h'] - -theorem subNatNat_add_negSucc (m n k : Nat) : - subNatNat m n + -[k+1] = subNatNat m (n + succ k) := by - have h := Nat.lt_or_ge m n - cases h with - | inr h' => - rw [subNatNat_of_le h'] - simp - rw [subNatNat_sub h', Nat.add_comm] - | inl h' => - have h₂ : m < n + succ k := Nat.lt_of_lt_of_le h' (le_add_right _ _) - have h₃ : m ≤ n + k := le_of_succ_le_succ h₂ - rw [subNatNat_of_lt h', subNatNat_of_lt h₂] - simp [Nat.add_comm] - rw [← add_succ, succ_pred_eq_of_pos (Nat.sub_pos_of_lt h'), add_succ, succ_sub h₃, - Nat.pred_succ] - rw [Nat.add_comm n, Nat.add_sub_assoc (Nat.le_of_lt h')] - -protected theorem add_assoc : ∀ a b c : Int, a + b + c = a + (b + c) - | (m:Nat), (n:Nat), c => aux1 .. - | Nat.cast m, b, Nat.cast k => by - rw [Int.add_comm, ← aux1, Int.add_comm k, aux1, Int.add_comm b] - | a, (n:Nat), (k:Nat) => by - rw [Int.add_comm, Int.add_comm a, ← aux1, Int.add_comm a, Int.add_comm k] - | -[m+1], -[n+1], (k:Nat) => aux2 .. - | -[m+1], (n:Nat), -[k+1] => by - rw [Int.add_comm, ← aux2, Int.add_comm n, ← aux2, Int.add_comm -[m+1]] - | (m:Nat), -[n+1], -[k+1] => by - rw [Int.add_comm, Int.add_comm m, Int.add_comm m, ← aux2, Int.add_comm -[k+1]] - | -[m+1], -[n+1], -[k+1] => by - simp [add_succ, Nat.add_comm, Nat.add_left_comm, neg_ofNat_succ] -where - aux1 (m n : Nat) : ∀ c : Int, m + n + c = m + (n + c) - | (k:Nat) => by simp [Nat.add_assoc] - | -[k+1] => by simp [subNatNat_add] - aux2 (m n k : Nat) : -[m+1] + -[n+1] + k = -[m+1] + (-[n+1] + k) := by - simp [add_succ] - rw [Int.add_comm, subNatNat_add_negSucc] - simp [add_succ, succ_add, Nat.add_comm] - -protected theorem add_left_comm (a b c : Int) : a + (b + c) = b + (a + c) := by - rw [← Int.add_assoc, Int.add_comm a, Int.add_assoc] - -protected theorem add_right_comm (a b c : Int) : a + b + c = a + c + b := by - rw [Int.add_assoc, Int.add_comm b, ← Int.add_assoc] - -/- ## negation -/ - -theorem subNatNat_self : ∀ n, subNatNat n n = 0 - | 0 => rfl - | succ m => by rw [subNatNat_of_sub_eq_zero (Nat.sub_self ..), Nat.sub_self, ofNat_zero] - -attribute [local simp] subNatNat_self - -@[local simp] protected theorem add_left_neg : ∀ a : Int, -a + a = 0 - | 0 => rfl - | succ m => by simp - | -[m+1] => by simp - -@[local simp] protected theorem add_right_neg (a : Int) : a + -a = 0 := by - rw [Int.add_comm, Int.add_left_neg] - -@[simp] protected theorem neg_eq_of_add_eq_zero {a b : Int} (h : a + b = 0) : -a = b := by - rw [← Int.add_zero (-a), ← h, ← Int.add_assoc, Int.add_left_neg, Int.zero_add] - -protected theorem eq_neg_of_eq_neg {a b : Int} (h : a = -b) : b = -a := by - rw [h, Int.neg_neg] - -protected theorem eq_neg_comm {a b : Int} : a = -b ↔ b = -a := - ⟨Int.eq_neg_of_eq_neg, Int.eq_neg_of_eq_neg⟩ - -protected theorem neg_eq_comm {a b : Int} : -a = b ↔ -b = a := by - rw [eq_comm, Int.eq_neg_comm, eq_comm] - -protected theorem neg_add_cancel_left (a b : Int) : -a + (a + b) = b := by - rw [← Int.add_assoc, Int.add_left_neg, Int.zero_add] - -protected theorem add_neg_cancel_left (a b : Int) : a + (-a + b) = b := by - rw [← Int.add_assoc, Int.add_right_neg, Int.zero_add] - -protected theorem add_neg_cancel_right (a b : Int) : a + b + -b = a := by - rw [Int.add_assoc, Int.add_right_neg, Int.add_zero] - -protected theorem neg_add_cancel_right (a b : Int) : a + -b + b = a := by - rw [Int.add_assoc, Int.add_left_neg, Int.add_zero] - -protected theorem add_left_cancel {a b c : Int} (h : a + b = a + c) : b = c := by - have h₁ : -a + (a + b) = -a + (a + c) := by rw [h] - simp [← Int.add_assoc, Int.add_left_neg, Int.zero_add] at h₁; exact h₁ - -@[local simp] protected theorem neg_add {a b : Int} : -(a + b) = -a + -b := by - apply Int.add_left_cancel (a := a + b) - rw [Int.add_right_neg, Int.add_comm a, ← Int.add_assoc, Int.add_assoc b, - Int.add_right_neg, Int.add_zero, Int.add_right_neg] - -/- ## subtraction -/ - -@[simp] theorem negSucc_sub_one (n : Nat) : -[n+1] - 1 = -[n + 1 +1] := rfl - -@[simp] protected theorem sub_self (a : Int) : a - a = 0 := by - rw [Int.sub_eq_add_neg, Int.add_right_neg] - -@[simp] protected theorem sub_zero (a : Int) : a - 0 = a := by simp [Int.sub_eq_add_neg] - -@[simp] protected theorem zero_sub (a : Int) : 0 - a = -a := by simp [Int.sub_eq_add_neg] - -protected theorem sub_eq_zero_of_eq {a b : Int} (h : a = b) : a - b = 0 := by - rw [h, Int.sub_self] - -protected theorem eq_of_sub_eq_zero {a b : Int} (h : a - b = 0) : a = b := by - have : 0 + b = b := by rw [Int.zero_add] - have : a - b + b = b := by rwa [h] - rwa [Int.sub_eq_add_neg, Int.neg_add_cancel_right] at this - -protected theorem sub_eq_zero {a b : Int} : a - b = 0 ↔ a = b := - ⟨Int.eq_of_sub_eq_zero, Int.sub_eq_zero_of_eq⟩ - -protected theorem sub_sub (a b c : Int) : a - b - c = a - (b + c) := by - simp [Int.sub_eq_add_neg, Int.add_assoc] - -protected theorem neg_sub (a b : Int) : -(a - b) = b - a := by - simp [Int.sub_eq_add_neg, Int.add_comm] - -protected theorem sub_sub_self (a b : Int) : a - (a - b) = b := by - simp [Int.sub_eq_add_neg, ← Int.add_assoc] - -protected theorem sub_neg (a b : Int) : a - -b = a + b := by simp [Int.sub_eq_add_neg] - -/- ## multiplication -/ - -@[simp] theorem ofNat_mul_negSucc (m n : Nat) : (m : Int) * -[n+1] = -↑(m * succ n) := rfl - -@[simp] theorem negSucc_mul_ofNat (m n : Nat) : -[m+1] * n = -↑(succ m * n) := rfl - -@[simp] theorem negSucc_mul_negSucc (m n : Nat) : -[m+1] * -[n+1] = succ m * succ n := rfl - -protected theorem mul_comm (a b : Int) : a * b = b * a := by - cases a <;> cases b <;> simp [Nat.mul_comm] - -theorem ofNat_mul_negOfNat (m n : Nat) : (m : Nat) * negOfNat n = negOfNat (m * n) := by - cases n <;> rfl - -theorem negOfNat_mul_ofNat (m n : Nat) : negOfNat m * (n : Nat) = negOfNat (m * n) := by - rw [Int.mul_comm]; simp [ofNat_mul_negOfNat, Nat.mul_comm] - -theorem negSucc_mul_negOfNat (m n : Nat) : -[m+1] * negOfNat n = ofNat (succ m * n) := by - cases n <;> rfl - -theorem negOfNat_mul_negSucc (m n : Nat) : negOfNat n * -[m+1] = ofNat (n * succ m) := by - rw [Int.mul_comm, negSucc_mul_negOfNat, Nat.mul_comm] - -attribute [local simp] ofNat_mul_negOfNat negOfNat_mul_ofNat - negSucc_mul_negOfNat negOfNat_mul_negSucc - -protected theorem mul_assoc (a b c : Int) : a * b * c = a * (b * c) := by - cases a <;> cases b <;> cases c <;> simp [Nat.mul_assoc] - -protected theorem mul_left_comm (a b c : Int) : a * (b * c) = b * (a * c) := by - rw [← Int.mul_assoc, ← Int.mul_assoc, Int.mul_comm a] - -protected theorem mul_right_comm (a b c : Int) : a * b * c = a * c * b := by - rw [Int.mul_assoc, Int.mul_assoc, Int.mul_comm b] - -@[simp] protected theorem mul_zero (a : Int) : a * 0 = 0 := by cases a <;> rfl - -@[simp] protected theorem zero_mul (a : Int) : 0 * a = 0 := Int.mul_comm .. ▸ a.mul_zero - -theorem negOfNat_eq_subNatNat_zero (n) : negOfNat n = subNatNat 0 n := by cases n <;> rfl - -theorem ofNat_mul_subNatNat (m n k : Nat) : - m * subNatNat n k = subNatNat (m * n) (m * k) := by - cases m with - | zero => simp [ofNat_zero, Int.zero_mul, Nat.zero_mul] - | succ m => cases n.lt_or_ge k with - | inl h => - have h' : succ m * n < succ m * k := Nat.mul_lt_mul_of_pos_left h (Nat.succ_pos m) - simp [subNatNat_of_lt h, subNatNat_of_lt h'] - rw [succ_pred_eq_of_pos (Nat.sub_pos_of_lt h), ← neg_ofNat_succ, Nat.mul_sub_left_distrib, - ← succ_pred_eq_of_pos (Nat.sub_pos_of_lt h')]; rfl - | inr h => - have h' : succ m * k ≤ succ m * n := Nat.mul_le_mul_left _ h - simp [subNatNat_of_le h, subNatNat_of_le h', Nat.mul_sub_left_distrib] - -theorem negOfNat_add (m n : Nat) : negOfNat m + negOfNat n = negOfNat (m + n) := by - cases m <;> cases n <;> simp [Nat.succ_add] <;> rfl - -theorem negSucc_mul_subNatNat (m n k : Nat) : - -[m+1] * subNatNat n k = subNatNat (succ m * k) (succ m * n) := by - cases n.lt_or_ge k with - | inl h => - have h' : succ m * n < succ m * k := Nat.mul_lt_mul_of_pos_left h (Nat.succ_pos m) - rw [subNatNat_of_lt h, subNatNat_of_le (Nat.le_of_lt h')] - simp [succ_pred_eq_of_pos (Nat.sub_pos_of_lt h), Nat.mul_sub_left_distrib] - | inr h => cases Nat.lt_or_ge k n with - | inl h' => - have h₁ : succ m * n > succ m * k := Nat.mul_lt_mul_of_pos_left h' (Nat.succ_pos m) - rw [subNatNat_of_le h, subNatNat_of_lt h₁, negSucc_mul_ofNat, - Nat.mul_sub_left_distrib, ← succ_pred_eq_of_pos (Nat.sub_pos_of_lt h₁)]; rfl - | inr h' => rw [Nat.le_antisymm h h', subNatNat_self, subNatNat_self, Int.mul_zero] - -attribute [local simp] ofNat_mul_subNatNat negOfNat_add negSucc_mul_subNatNat - -protected theorem mul_add : ∀ a b c : Int, a * (b + c) = a * b + a * c - | (m:Nat), (n:Nat), (k:Nat) => by simp [Nat.left_distrib] - | (m:Nat), (n:Nat), -[k+1] => by - simp [negOfNat_eq_subNatNat_zero]; rw [← subNatNat_add]; rfl - | (m:Nat), -[n+1], (k:Nat) => by - simp [negOfNat_eq_subNatNat_zero]; rw [Int.add_comm, ← subNatNat_add]; rfl - | (m:Nat), -[n+1], -[k+1] => by simp; rw [← Nat.left_distrib, succ_add]; rfl - | -[m+1], (n:Nat), (k:Nat) => by simp [Nat.mul_comm]; rw [← Nat.right_distrib, Nat.mul_comm] - | -[m+1], (n:Nat), -[k+1] => by - simp [negOfNat_eq_subNatNat_zero]; rw [Int.add_comm, ← subNatNat_add]; rfl - | -[m+1], -[n+1], (k:Nat) => by simp [negOfNat_eq_subNatNat_zero]; rw [← subNatNat_add]; rfl - | -[m+1], -[n+1], -[k+1] => by simp; rw [← Nat.left_distrib, succ_add]; rfl - -protected theorem add_mul (a b c : Int) : (a + b) * c = a * c + b * c := by - simp [Int.mul_comm, Int.mul_add] - -protected theorem neg_mul_eq_neg_mul (a b : Int) : -(a * b) = -a * b := - Int.neg_eq_of_add_eq_zero <| by rw [← Int.add_mul, Int.add_right_neg, Int.zero_mul] - -protected theorem neg_mul_eq_mul_neg (a b : Int) : -(a * b) = a * -b := - Int.neg_eq_of_add_eq_zero <| by rw [← Int.mul_add, Int.add_right_neg, Int.mul_zero] - -@[local simp] protected theorem neg_mul (a b : Int) : -a * b = -(a * b) := - (Int.neg_mul_eq_neg_mul a b).symm - -@[local simp] protected theorem mul_neg (a b : Int) : a * -b = -(a * b) := - (Int.neg_mul_eq_mul_neg a b).symm - -protected theorem neg_mul_neg (a b : Int) : -a * -b = a * b := by simp - -protected theorem neg_mul_comm (a b : Int) : -a * b = a * -b := by simp - -protected theorem mul_sub (a b c : Int) : a * (b - c) = a * b - a * c := by - simp [Int.sub_eq_add_neg, Int.mul_add] - -protected theorem sub_mul (a b c : Int) : (a - b) * c = a * c - b * c := by - simp [Int.sub_eq_add_neg, Int.add_mul] - -protected theorem zero_ne_one : (0 : Int) ≠ 1 := fun. - -@[simp] protected theorem sub_add_cancel (a b : Int) : a - b + b = a := - Int.neg_add_cancel_right a b - -@[simp] protected theorem add_sub_cancel (a b : Int) : a + b - b = a := - Int.add_neg_cancel_right a b - -protected theorem add_sub_assoc (a b c : Int) : a + b - c = a + (b - c) := by - rw [Int.sub_eq_add_neg, Int.add_assoc, ← Int.sub_eq_add_neg] - -@[norm_cast] theorem ofNat_sub (h : m ≤ n) : ((n - m : Nat) : Int) = n - m := by - match m with - | 0 => rfl - | succ m => - show ofNat (n - succ m) = subNatNat n (succ m) - rw [subNatNat, Nat.sub_eq_zero_of_le h] - -theorem negSucc_coe' (n : Nat) : -[n+1] = -↑n - 1 := by - rw [Int.sub_eq_add_neg, ← Int.neg_add]; rfl - -protected theorem subNatNat_eq_coe {m n : Nat} : subNatNat m n = ↑m - ↑n := by - apply subNatNat_elim m n fun m n i => i = m - n - · intros i n - rw [Int.ofNat_add, Int.sub_eq_add_neg, Int.add_assoc, Int.add_left_comm, - Int.add_right_neg, Int.add_zero] - · intros i n - simp only [negSucc_coe, ofNat_add, Int.sub_eq_add_neg, Int.neg_add, ← Int.add_assoc] - rw [← @Int.sub_eq_add_neg n, ← ofNat_sub, Nat.sub_self, ofNat_zero, Int.zero_add] - apply Nat.le_refl - -theorem toNat_sub (m n : Nat) : toNat (m - n) = m - n := by - rw [← Int.subNatNat_eq_coe] - refine subNatNat_elim m n (fun m n i => toNat i = m - n) (fun i n => ?_) (fun i n => ?_) - · exact (Nat.add_sub_cancel_left ..).symm - · dsimp; rw [Nat.add_assoc, Nat.sub_eq_zero_of_le (Nat.le_add_right ..)]; rfl - -@[simp] protected theorem one_mul : ∀ a : Int, 1 * a = a - | ofNat n => show ofNat (1 * n) = ofNat n by rw [Nat.one_mul] - | -[n+1] => show -[1 * n +1] = -[n+1] by rw [Nat.one_mul] - -@[simp] protected theorem mul_one (a : Int) : a * 1 = a := by rw [Int.mul_comm, Int.one_mul] - -protected theorem mul_neg_one (a : Int) : a * -1 = -a := by rw [Int.mul_neg, Int.mul_one] - -protected theorem neg_eq_neg_one_mul : ∀ a : Int, -a = -1 * a - | 0 => rfl - | succ n => show _ = -[1 * n +1] by rw [Nat.one_mul]; rfl - | -[n+1] => show _ = ofNat _ by rw [Nat.one_mul]; rfl - -/-! ## Order properties of the integers -/ - -theorem nonneg_def {a : Int} : NonNeg a ↔ ∃ n : Nat, a = n := - ⟨fun ⟨n⟩ => ⟨n, rfl⟩, fun h => match a, h with | _, ⟨n, rfl⟩ => ⟨n⟩⟩ - -theorem NonNeg.elim {a : Int} : NonNeg a → ∃ n : Nat, a = n := nonneg_def.1 - -theorem nonneg_or_nonneg_neg : ∀ (a : Int), NonNeg a ∨ NonNeg (-a) - | (_:Nat) => .inl ⟨_⟩ - | -[_+1] => .inr ⟨_⟩ - -theorem le_def (a b : Int) : a ≤ b ↔ NonNeg (b - a) := .rfl - -theorem lt_iff_add_one_le (a b : Int) : a < b ↔ a + 1 ≤ b := .rfl - -theorem le.intro_sub {a b : Int} (n : Nat) (h : b - a = n) : a ≤ b := by - simp [le_def, h]; constructor - -theorem le.intro {a b : Int} (n : Nat) (h : a + n = b) : a ≤ b := - le.intro_sub n <| by rw [← h, Int.add_comm]; simp [Int.sub_eq_add_neg, Int.add_assoc] - -theorem le.dest_sub {a b : Int} (h : a ≤ b) : ∃ n : Nat, b - a = n := nonneg_def.1 h - -theorem le.dest {a b : Int} (h : a ≤ b) : ∃ n : Nat, a + n = b := - let ⟨n, h₁⟩ := le.dest_sub h - ⟨n, by rw [← h₁, Int.add_comm]; simp [Int.sub_eq_add_neg, Int.add_assoc]⟩ - -protected theorem le_total (a b : Int) : a ≤ b ∨ b ≤ a := - (nonneg_or_nonneg_neg (b - a)).imp_right fun H => by - rwa [show -(b - a) = a - b by simp [Int.add_comm, Int.sub_eq_add_neg]] at H - -@[simp, norm_cast] theorem ofNat_le {m n : Nat} : (↑m : Int) ≤ ↑n ↔ m ≤ n := - ⟨fun h => - let ⟨k, hk⟩ := le.dest h - Nat.le.intro <| Int.ofNat.inj <| (Int.ofNat_add m k).trans hk, - fun h => - let ⟨k, (hk : m + k = n)⟩ := Nat.le.dest h - le.intro k (by rw [← hk]; rfl)⟩ - -theorem ofNat_zero_le (n : Nat) : 0 ≤ (↑n : Int) := ofNat_le.2 n.zero_le - -theorem eq_ofNat_of_zero_le {a : Int} (h : 0 ≤ a) : ∃ n : Nat, a = n := by - have t := le.dest_sub h; rwa [Int.sub_zero] at t - -theorem eq_succ_of_zero_lt {a : Int} (h : 0 < a) : ∃ n : Nat, a = n.succ := - let ⟨n, (h : ↑(1 + n) = a)⟩ := le.dest h - ⟨n, by rw [Nat.add_comm] at h; exact h.symm⟩ - -theorem lt_add_succ (a : Int) (n : Nat) : a < a + Nat.succ n := - le.intro n <| by rw [Int.add_comm, Int.add_left_comm]; rfl - -theorem lt.intro {a b : Int} {n : Nat} (h : a + Nat.succ n = b) : a < b := - h ▸ lt_add_succ a n - -theorem lt.dest {a b : Int} (h : a < b) : ∃ n : Nat, a + Nat.succ n = b := - let ⟨n, h⟩ := le.dest h; ⟨n, by rwa [Int.add_comm, Int.add_left_comm] at h⟩ - -@[simp, norm_cast] theorem ofNat_lt {n m : Nat} : (↑n : Int) < ↑m ↔ n < m := by - rw [lt_iff_add_one_le, ← ofNat_succ, ofNat_le]; rfl - -@[simp, norm_cast] theorem ofNat_pos {n : Nat} : 0 < (↑n : Int) ↔ 0 < n := ofNat_lt - -theorem ofNat_nonneg (n : Nat) : 0 ≤ (n : Int) := ⟨_⟩ - -theorem ofNat_succ_pos (n : Nat) : 0 < (succ n : Int) := ofNat_lt.2 <| Nat.succ_pos _ - -@[simp] protected theorem le_refl (a : Int) : a ≤ a := - le.intro _ (Int.add_zero a) - -protected theorem le_trans {a b c : Int} (h₁ : a ≤ b) (h₂ : b ≤ c) : a ≤ c := - let ⟨n, hn⟩ := le.dest h₁; let ⟨m, hm⟩ := le.dest h₂ - le.intro (n + m) <| by rw [← hm, ← hn, Int.add_assoc, ofNat_add] - -protected theorem le_antisymm {a b : Int} (h₁ : a ≤ b) (h₂ : b ≤ a) : a = b := by - let ⟨n, hn⟩ := le.dest h₁; let ⟨m, hm⟩ := le.dest h₂ - have := hn; rw [← hm, Int.add_assoc, ← ofNat_add] at this - have := Int.ofNat.inj <| Int.add_left_cancel <| this.trans (Int.add_zero _).symm - rw [← hn, Nat.eq_zero_of_add_eq_zero_left this, ofNat_zero, Int.add_zero a] - -protected theorem lt_irrefl (a : Int) : ¬a < a := fun H => - let ⟨n, hn⟩ := lt.dest H - have : (a+Nat.succ n) = a+0 := by - rw [hn, Int.add_zero] - have : Nat.succ n = 0 := Int.ofNat.inj (Int.add_left_cancel this) - show False from Nat.succ_ne_zero _ this - -protected theorem ne_of_lt {a b : Int} (h : a < b) : a ≠ b := fun e => by - cases e; exact Int.lt_irrefl _ h - -protected theorem ne_of_gt {a b : Int} (h : b < a) : a ≠ b := (Int.ne_of_lt h).symm - -protected theorem le_of_lt {a b : Int} (h : a < b) : a ≤ b := - let ⟨_, hn⟩ := lt.dest h; le.intro _ hn - -protected theorem lt_iff_le_and_ne {a b : Int} : a < b ↔ a ≤ b ∧ a ≠ b := by - refine ⟨fun h => ⟨Int.le_of_lt h, Int.ne_of_lt h⟩, fun ⟨aleb, aneb⟩ => ?_⟩ - let ⟨n, hn⟩ := le.dest aleb - have : n ≠ 0 := aneb.imp fun eq => by rw [← hn, eq, ofNat_zero, Int.add_zero] - apply lt.intro; rwa [← Nat.succ_pred_eq_of_pos (Nat.pos_of_ne_zero this)] at hn - -theorem lt_succ (a : Int) : a < a + 1 := Int.le_refl _ - -protected theorem mul_nonneg {a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a * b := by - let ⟨n, hn⟩ := eq_ofNat_of_zero_le ha - let ⟨m, hm⟩ := eq_ofNat_of_zero_le hb - rw [hn, hm, ← ofNat_mul]; apply ofNat_nonneg - -protected theorem mul_pos {a b : Int} (ha : 0 < a) (hb : 0 < b) : 0 < a * b := by - let ⟨n, hn⟩ := eq_succ_of_zero_lt ha - let ⟨m, hm⟩ := eq_succ_of_zero_lt hb - rw [hn, hm, ← ofNat_mul]; apply ofNat_succ_pos - -protected theorem zero_lt_one : (0 : Int) < 1 := ⟨_⟩ - -protected theorem lt_iff_le_not_le {a b : Int} : a < b ↔ a ≤ b ∧ ¬b ≤ a := by - rw [Int.lt_iff_le_and_ne] - constructor <;> refine fun ⟨h, h'⟩ => ⟨h, h'.imp fun h' => ?_⟩ - · exact Int.le_antisymm h h' - · subst h'; apply Int.le_refl - -protected theorem not_le {a b : Int} : ¬a ≤ b ↔ b < a := - ⟨fun h => Int.lt_iff_le_not_le.2 ⟨(Int.le_total ..).resolve_right h, h⟩, - fun h => (Int.lt_iff_le_not_le.1 h).2⟩ - -protected theorem not_lt {a b : Int} : ¬a < b ↔ b ≤ a := - by rw [← Int.not_le, Decidable.not_not] - -protected theorem lt_trichotomy (a b : Int) : a < b ∨ a = b ∨ b < a := - if eq : a = b then .inr <| .inl eq else - if le : a ≤ b then .inl <| Int.lt_iff_le_and_ne.2 ⟨le, eq⟩ else - .inr <| .inr <| Int.not_le.1 le - -protected theorem ne_iff_lt_or_gt {a b : Int} : a ≠ b ↔ a < b ∨ b < a := by - constructor - · intro h - rcases Int.lt_trichotomy a b with lt | rfl | gt - · exact Or.inl lt - · simp_all - · exact Or.inr gt - · rintro (lt | gt) - · exact Int.ne_of_lt lt - · exact Int.ne_of_gt gt - -protected alias ⟨lt_or_gt_of_ne, _⟩ := Int.ne_iff_lt_or_gt - -protected theorem eq_iff_le_and_ge {x y : Int} : x = y ↔ x ≤ y ∧ y ≤ x := by - constructor - · simp_all - · rintro ⟨h₁, h₂⟩ - exact Int.le_antisymm h₁ h₂ - -protected theorem lt_of_le_of_lt {a b c : Int} (h₁ : a ≤ b) (h₂ : b < c) : a < c := - Int.not_le.1 fun h => Int.not_le.2 h₂ (Int.le_trans h h₁) - -protected theorem lt_of_lt_of_le {a b c : Int} (h₁ : a < b) (h₂ : b ≤ c) : a < c := - Int.not_le.1 fun h => Int.not_le.2 h₁ (Int.le_trans h₂ h) - -protected theorem lt_trans {a b c : Int} (h₁ : a < b) (h₂ : b < c) : a < c := - Int.lt_of_le_of_lt (Int.le_of_lt h₁) h₂ - -instance : Trans (α := Int) (· ≤ ·) (· ≤ ·) (· ≤ ·) := ⟨Int.le_trans⟩ - -instance : Trans (α := Int) (· < ·) (· ≤ ·) (· < ·) := ⟨Int.lt_of_lt_of_le⟩ - -instance : Trans (α := Int) (· ≤ ·) (· < ·) (· < ·) := ⟨Int.lt_of_le_of_lt⟩ - -instance : Trans (α := Int) (· < ·) (· < ·) (· < ·) := ⟨Int.lt_trans⟩ - -protected theorem le_of_not_le {a b : Int} : ¬ a ≤ b → b ≤ a := (Int.le_total a b).resolve_left - -protected theorem min_def (n m : Int) : min n m = if n ≤ m then n else m := rfl - -protected theorem max_def (n m : Int) : max n m = if n ≤ m then m else n := rfl - -protected theorem min_comm (a b : Int) : min a b = min b a := by - simp [Int.min_def] - by_cases h₁ : a ≤ b <;> by_cases h₂ : b ≤ a <;> simp [h₁, h₂] - · exact Int.le_antisymm h₁ h₂ - · cases not_or_intro h₁ h₂ <| Int.le_total .. - -protected theorem min_le_right (a b : Int) : min a b ≤ b := by rw [Int.min_def]; split <;> simp [*] - -protected theorem min_le_left (a b : Int) : min a b ≤ a := Int.min_comm .. ▸ Int.min_le_right .. - -protected theorem le_min {a b c : Int} : a ≤ min b c ↔ a ≤ b ∧ a ≤ c := - ⟨fun h => ⟨Int.le_trans h (Int.min_le_left ..), Int.le_trans h (Int.min_le_right ..)⟩, - fun ⟨h₁, h₂⟩ => by rw [Int.min_def]; split <;> assumption⟩ - -protected theorem min_eq_left {a b : Int} (h : a ≤ b) : min a b = a := by simp [Int.min_def, h] - -protected theorem min_eq_right {a b : Int} (h : b ≤ a) : min a b = b := by - rw [Int.min_comm a b]; exact Int.min_eq_left h - -protected theorem max_comm (a b : Int) : max a b = max b a := by - simp only [Int.max_def] - by_cases h₁ : a ≤ b <;> by_cases h₂ : b ≤ a <;> simp [h₁, h₂] - · exact Int.le_antisymm h₂ h₁ - · cases not_or_intro h₁ h₂ <| Int.le_total .. - -protected theorem le_max_left (a b : Int) : a ≤ max a b := by rw [Int.max_def]; split <;> simp [*] - -protected theorem le_max_right (a b : Int) : b ≤ max a b := Int.max_comm .. ▸ Int.le_max_left .. - -protected theorem max_le {a b c : Int} : max a b ≤ c ↔ a ≤ c ∧ b ≤ c := - ⟨fun h => ⟨Int.le_trans (Int.le_max_left ..) h, Int.le_trans (Int.le_max_right ..) h⟩, - fun ⟨h₁, h₂⟩ => by rw [Int.max_def]; split <;> assumption⟩ - -protected theorem max_eq_right {a b : Int} (h : a ≤ b) : max a b = b := by - simp [Int.max_def, h, Int.not_lt.2 h] - -protected theorem max_eq_left {a b : Int} (h : b ≤ a) : max a b = a := by - rw [← Int.max_comm b a]; exact Int.max_eq_right h - -theorem eq_natAbs_of_zero_le {a : Int} (h : 0 ≤ a) : a = natAbs a := by - let ⟨n, e⟩ := eq_ofNat_of_zero_le h - rw [e]; rfl - -theorem le_natAbs {a : Int} : a ≤ natAbs a := - match Int.le_total 0 a with - | .inl h => by rw [eq_natAbs_of_zero_le h]; apply Int.le_refl - | .inr h => Int.le_trans h (ofNat_zero_le _) - -theorem negSucc_lt_zero (n : Nat) : -[n+1] < 0 := - Int.not_le.1 fun h => let ⟨_, h⟩ := eq_ofNat_of_zero_le h; nomatch h - -@[simp] theorem negSucc_not_nonneg (n : Nat) : 0 ≤ -[n+1] ↔ False := by - simp only [Int.not_le, iff_false]; exact Int.negSucc_lt_zero n - -@[simp] theorem negSucc_not_pos (n : Nat) : 0 < -[n+1] ↔ False := by - simp only [Int.not_lt, iff_false]; constructor - -theorem eq_negSucc_of_lt_zero : ∀ {a : Int}, a < 0 → ∃ n : Nat, a = -[n+1] - | ofNat _, h => absurd h (Int.not_lt.2 (ofNat_zero_le _)) - | -[n+1], _ => ⟨n, rfl⟩ - -protected theorem add_le_add_left {a b : Int} (h : a ≤ b) (c : Int) : c + a ≤ c + b := - let ⟨n, hn⟩ := le.dest h; le.intro n <| by rw [Int.add_assoc, hn] - -protected theorem add_lt_add_left {a b : Int} (h : a < b) (c : Int) : c + a < c + b := - Int.lt_iff_le_and_ne.2 ⟨Int.add_le_add_left (Int.le_of_lt h) _, fun heq => - b.lt_irrefl <| by rwa [Int.add_left_cancel heq] at h⟩ - -protected theorem add_le_add_right {a b : Int} (h : a ≤ b) (c : Int) : a + c ≤ b + c := - Int.add_comm c a ▸ Int.add_comm c b ▸ Int.add_le_add_left h c - -protected theorem add_lt_add_right {a b : Int} (h : a < b) (c : Int) : a + c < b + c := - Int.add_comm c a ▸ Int.add_comm c b ▸ Int.add_lt_add_left h c - -protected theorem le_of_add_le_add_left {a b c : Int} (h : a + b ≤ a + c) : b ≤ c := by - have : -a + (a + b) ≤ -a + (a + c) := Int.add_le_add_left h _ - simp [Int.neg_add_cancel_left] at this - assumption - -protected theorem lt_of_add_lt_add_left {a b c : Int} (h : a + b < a + c) : b < c := by - have : -a + (a + b) < -a + (a + c) := Int.add_lt_add_left h _ - simp [Int.neg_add_cancel_left] at this - assumption - -protected theorem le_of_add_le_add_right {a b c : Int} (h : a + b ≤ c + b) : a ≤ c := - Int.le_of_add_le_add_left (a := b) <| by rwa [Int.add_comm b a, Int.add_comm b c] - -protected theorem lt_of_add_lt_add_right {a b c : Int} (h : a + b < c + b) : a < c := - Int.lt_of_add_lt_add_left (a := b) <| by rwa [Int.add_comm b a, Int.add_comm b c] - -protected theorem add_le_add_iff_left (a : Int) : a + b ≤ a + c ↔ b ≤ c := - ⟨Int.le_of_add_le_add_left, (Int.add_le_add_left · _)⟩ - -protected theorem add_lt_add_iff_left (a : Int) : a + b < a + c ↔ b < c := - ⟨Int.lt_of_add_lt_add_left, (Int.add_lt_add_left · _)⟩ - -protected theorem add_le_add_iff_right (c : Int) : a + c ≤ b + c ↔ a ≤ b := - ⟨Int.le_of_add_le_add_right, (Int.add_le_add_right · _)⟩ - -protected theorem add_lt_add_iff_right (c : Int) : a + c < b + c ↔ a < b := - ⟨Int.lt_of_add_lt_add_right, (Int.add_lt_add_right · _)⟩ - -protected theorem add_le_add {a b c d : Int} (h₁ : a ≤ b) (h₂ : c ≤ d) : a + c ≤ b + d := - Int.le_trans (Int.add_le_add_right h₁ c) (Int.add_le_add_left h₂ b) - -protected theorem le_add_of_nonneg_right {a b : Int} (h : 0 ≤ b) : a ≤ a + b := by - have : a + b ≥ a + 0 := Int.add_le_add_left h a - rwa [Int.add_zero] at this - -protected theorem le_add_of_nonneg_left {a b : Int} (h : 0 ≤ b) : a ≤ b + a := by - have : 0 + a ≤ b + a := Int.add_le_add_right h a - rwa [Int.zero_add] at this - -protected theorem add_lt_add {a b c d : Int} (h₁ : a < b) (h₂ : c < d) : a + c < b + d := - Int.lt_trans (Int.add_lt_add_right h₁ c) (Int.add_lt_add_left h₂ b) - -protected theorem add_lt_add_of_le_of_lt {a b c d : Int} (h₁ : a ≤ b) (h₂ : c < d) : - a + c < b + d := - Int.lt_of_le_of_lt (Int.add_le_add_right h₁ c) (Int.add_lt_add_left h₂ b) - -protected theorem add_lt_add_of_lt_of_le {a b c d : Int} (h₁ : a < b) (h₂ : c ≤ d) : - a + c < b + d := - Int.lt_of_lt_of_le (Int.add_lt_add_right h₁ c) (Int.add_le_add_left h₂ b) - -protected theorem lt_add_of_pos_right (a : Int) {b : Int} (h : 0 < b) : a < a + b := by - have : a + 0 < a + b := Int.add_lt_add_left h a - rwa [Int.add_zero] at this - -protected theorem lt_add_of_pos_left (a : Int) {b : Int} (h : 0 < b) : a < b + a := by - have : 0 + a < b + a := Int.add_lt_add_right h a - rwa [Int.zero_add] at this - -protected theorem add_nonneg {a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a + b := - Int.zero_add 0 ▸ Int.add_le_add ha hb - -protected theorem add_pos {a b : Int} (ha : 0 < a) (hb : 0 < b) : 0 < a + b := - Int.zero_add 0 ▸ Int.add_lt_add ha hb - -protected theorem add_pos_of_pos_of_nonneg {a b : Int} (ha : 0 < a) (hb : 0 ≤ b) : 0 < a + b := - Int.zero_add 0 ▸ Int.add_lt_add_of_lt_of_le ha hb - -protected theorem add_pos_of_nonneg_of_pos {a b : Int} (ha : 0 ≤ a) (hb : 0 < b) : 0 < a + b := - Int.zero_add 0 ▸ Int.add_lt_add_of_le_of_lt ha hb - -protected theorem add_nonpos {a b : Int} (ha : a ≤ 0) (hb : b ≤ 0) : a + b ≤ 0 := - Int.zero_add 0 ▸ Int.add_le_add ha hb - -protected theorem add_neg {a b : Int} (ha : a < 0) (hb : b < 0) : a + b < 0 := - Int.zero_add 0 ▸ Int.add_lt_add ha hb - -protected theorem add_neg_of_neg_of_nonpos {a b : Int} (ha : a < 0) (hb : b ≤ 0) : a + b < 0 := - Int.zero_add 0 ▸ Int.add_lt_add_of_lt_of_le ha hb - -protected theorem add_neg_of_nonpos_of_neg {a b : Int} (ha : a ≤ 0) (hb : b < 0) : a + b < 0 := - Int.zero_add 0 ▸ Int.add_lt_add_of_le_of_lt ha hb - -protected theorem lt_add_of_le_of_pos {a b c : Int} (hbc : b ≤ c) (ha : 0 < a) : b < c + a := - Int.add_zero b ▸ Int.add_lt_add_of_le_of_lt hbc ha - -theorem add_one_le_iff {a b : Int} : a + 1 ≤ b ↔ a < b := .rfl - -theorem lt_add_one_iff {a b : Int} : a < b + 1 ↔ a ≤ b := Int.add_le_add_iff_right _ - -@[simp] theorem succ_ofNat_pos (n : Nat) : 0 < (n : Int) + 1 := - lt_add_one_iff.2 (ofNat_zero_le _) - -theorem le_add_one {a b : Int} (h : a ≤ b) : a ≤ b + 1 := - Int.le_of_lt (Int.lt_add_one_iff.2 h) - -protected theorem neg_le_neg {a b : Int} (h : a ≤ b) : -b ≤ -a := by - have : 0 ≤ -a + b := Int.add_left_neg a ▸ Int.add_le_add_left h (-a) - have : 0 + -b ≤ -a + b + -b := Int.add_le_add_right this (-b) - rwa [Int.add_neg_cancel_right, Int.zero_add] at this - -protected theorem le_of_neg_le_neg {a b : Int} (h : -b ≤ -a) : a ≤ b := - suffices - -a ≤ - -b by simp [Int.neg_neg] at this; assumption - Int.neg_le_neg h - -protected theorem nonneg_of_neg_nonpos {a : Int} (h : -a ≤ 0) : 0 ≤ a := - Int.le_of_neg_le_neg <| by rwa [Int.neg_zero] - -protected theorem neg_nonpos_of_nonneg {a : Int} (h : 0 ≤ a) : -a ≤ 0 := by - have : -a ≤ -0 := Int.neg_le_neg h - rwa [Int.neg_zero] at this - -protected theorem nonpos_of_neg_nonneg {a : Int} (h : 0 ≤ -a) : a ≤ 0 := - Int.le_of_neg_le_neg <| by rwa [Int.neg_zero] - -protected theorem neg_nonneg_of_nonpos {a : Int} (h : a ≤ 0) : 0 ≤ -a := by - have : -0 ≤ -a := Int.neg_le_neg h - rwa [Int.neg_zero] at this - -protected theorem neg_lt_neg {a b : Int} (h : a < b) : -b < -a := by - have : 0 < -a + b := Int.add_left_neg a ▸ Int.add_lt_add_left h (-a) - have : 0 + -b < -a + b + -b := Int.add_lt_add_right this (-b) - rwa [Int.add_neg_cancel_right, Int.zero_add] at this - -protected theorem lt_of_neg_lt_neg {a b : Int} (h : -b < -a) : a < b := - Int.neg_neg a ▸ Int.neg_neg b ▸ Int.neg_lt_neg h - -protected theorem pos_of_neg_neg {a : Int} (h : -a < 0) : 0 < a := - Int.lt_of_neg_lt_neg <| by rwa [Int.neg_zero] - -protected theorem neg_neg_of_pos {a : Int} (h : 0 < a) : -a < 0 := by - have : -a < -0 := Int.neg_lt_neg h - rwa [Int.neg_zero] at this - -protected theorem neg_of_neg_pos {a : Int} (h : 0 < -a) : a < 0 := - have : -0 < -a := by rwa [Int.neg_zero] - Int.lt_of_neg_lt_neg this - -protected theorem neg_pos_of_neg {a : Int} (h : a < 0) : 0 < -a := by - have : -0 < -a := Int.neg_lt_neg h - rwa [Int.neg_zero] at this - -protected theorem le_neg_of_le_neg {a b : Int} (h : a ≤ -b) : b ≤ -a := by - have h := Int.neg_le_neg h - rwa [Int.neg_neg] at h - -protected theorem neg_le_of_neg_le {a b : Int} (h : -a ≤ b) : -b ≤ a := by - have h := Int.neg_le_neg h - rwa [Int.neg_neg] at h - -protected theorem lt_neg_of_lt_neg {a b : Int} (h : a < -b) : b < -a := by - have h := Int.neg_lt_neg h - rwa [Int.neg_neg] at h - -protected theorem neg_lt_of_neg_lt {a b : Int} (h : -a < b) : -b < a := by - have h := Int.neg_lt_neg h - rwa [Int.neg_neg] at h - -protected theorem sub_nonneg_of_le {a b : Int} (h : b ≤ a) : 0 ≤ a - b := by - have h := Int.add_le_add_right h (-b) - rwa [Int.add_right_neg] at h - -protected theorem le_of_sub_nonneg {a b : Int} (h : 0 ≤ a - b) : b ≤ a := by - have h := Int.add_le_add_right h b - rwa [Int.sub_add_cancel, Int.zero_add] at h - -protected theorem sub_nonpos_of_le {a b : Int} (h : a ≤ b) : a - b ≤ 0 := by - have h := Int.add_le_add_right h (-b) - rwa [Int.add_right_neg] at h - -protected theorem le_of_sub_nonpos {a b : Int} (h : a - b ≤ 0) : a ≤ b := by - have h := Int.add_le_add_right h b - rwa [Int.sub_add_cancel, Int.zero_add] at h - -protected theorem sub_pos_of_lt {a b : Int} (h : b < a) : 0 < a - b := by - have h := Int.add_lt_add_right h (-b) - rwa [Int.add_right_neg] at h - -protected theorem lt_of_sub_pos {a b : Int} (h : 0 < a - b) : b < a := by - have h := Int.add_lt_add_right h b - rwa [Int.sub_add_cancel, Int.zero_add] at h - -protected theorem sub_neg_of_lt {a b : Int} (h : a < b) : a - b < 0 := by - have h := Int.add_lt_add_right h (-b) - rwa [Int.add_right_neg] at h - -protected theorem lt_of_sub_neg {a b : Int} (h : a - b < 0) : a < b := by - have h := Int.add_lt_add_right h b - rwa [Int.sub_add_cancel, Int.zero_add] at h - -protected theorem add_le_of_le_neg_add {a b c : Int} (h : b ≤ -a + c) : a + b ≤ c := by - have h := Int.add_le_add_left h a - rwa [Int.add_neg_cancel_left] at h - -protected theorem le_neg_add_of_add_le {a b c : Int} (h : a + b ≤ c) : b ≤ -a + c := by - have h := Int.add_le_add_left h (-a) - rwa [Int.neg_add_cancel_left] at h - -protected theorem add_le_of_le_sub_left {a b c : Int} (h : b ≤ c - a) : a + b ≤ c := by - have h := Int.add_le_add_left h a - rwa [← Int.add_sub_assoc, Int.add_comm a c, Int.add_sub_cancel] at h - -protected theorem le_sub_left_of_add_le {a b c : Int} (h : a + b ≤ c) : b ≤ c - a := by - have h := Int.add_le_add_right h (-a) - rwa [Int.add_comm a b, Int.add_neg_cancel_right] at h - -protected theorem add_le_of_le_sub_right {a b c : Int} (h : a ≤ c - b) : a + b ≤ c := by - have h := Int.add_le_add_right h b - rwa [Int.sub_add_cancel] at h - -protected theorem le_sub_right_of_add_le {a b c : Int} (h : a + b ≤ c) : a ≤ c - b := by - have h := Int.add_le_add_right h (-b) - rwa [Int.add_neg_cancel_right] at h - -protected theorem le_add_of_neg_add_le {a b c : Int} (h : -b + a ≤ c) : a ≤ b + c := by - have h := Int.add_le_add_left h b - rwa [Int.add_neg_cancel_left] at h - -protected theorem neg_add_le_of_le_add {a b c : Int} (h : a ≤ b + c) : -b + a ≤ c := by - have h := Int.add_le_add_left h (-b) - rwa [Int.neg_add_cancel_left] at h - -protected theorem le_add_of_sub_left_le {a b c : Int} (h : a - b ≤ c) : a ≤ b + c := by - have h := Int.add_le_add_right h b - rwa [Int.sub_add_cancel, Int.add_comm] at h - -protected theorem sub_left_le_of_le_add {a b c : Int} (h : a ≤ b + c) : a - b ≤ c := by - have h := Int.add_le_add_right h (-b) - rwa [Int.add_comm b c, Int.add_neg_cancel_right] at h - -protected theorem le_add_of_sub_right_le {a b c : Int} (h : a - c ≤ b) : a ≤ b + c := by - have h := Int.add_le_add_right h c - rwa [Int.sub_add_cancel] at h - -protected theorem sub_right_le_of_le_add {a b c : Int} (h : a ≤ b + c) : a - c ≤ b := by - have h := Int.add_le_add_right h (-c) - rwa [Int.add_neg_cancel_right] at h - -protected theorem le_add_of_neg_add_le_left {a b c : Int} (h : -b + a ≤ c) : a ≤ b + c := by - rw [Int.add_comm] at h - exact Int.le_add_of_sub_left_le h - -protected theorem neg_add_le_left_of_le_add {a b c : Int} (h : a ≤ b + c) : -b + a ≤ c := by - rw [Int.add_comm] - exact Int.sub_left_le_of_le_add h - -protected theorem le_add_of_neg_add_le_right {a b c : Int} (h : -c + a ≤ b) : a ≤ b + c := by - rw [Int.add_comm] at h - exact Int.le_add_of_sub_right_le h - -protected theorem neg_add_le_right_of_le_add {a b c : Int} (h : a ≤ b + c) : -c + a ≤ b := by - rw [Int.add_comm] at h - exact Int.neg_add_le_left_of_le_add h - -protected theorem le_add_of_neg_le_sub_left {a b c : Int} (h : -a ≤ b - c) : c ≤ a + b := - Int.le_add_of_neg_add_le_left (Int.add_le_of_le_sub_right h) - -protected theorem neg_le_sub_left_of_le_add {a b c : Int} (h : c ≤ a + b) : -a ≤ b - c := by - have h := Int.le_neg_add_of_add_le (Int.sub_left_le_of_le_add h) - rwa [Int.add_comm] at h - -protected theorem le_add_of_neg_le_sub_right {a b c : Int} (h : -b ≤ a - c) : c ≤ a + b := - Int.le_add_of_sub_right_le (Int.add_le_of_le_sub_left h) - -protected theorem neg_le_sub_right_of_le_add {a b c : Int} (h : c ≤ a + b) : -b ≤ a - c := - Int.le_sub_left_of_add_le (Int.sub_right_le_of_le_add h) - -protected theorem sub_le_of_sub_le {a b c : Int} (h : a - b ≤ c) : a - c ≤ b := - Int.sub_left_le_of_le_add (Int.le_add_of_sub_right_le h) - -protected theorem sub_le_sub_left {a b : Int} (h : a ≤ b) (c : Int) : c - b ≤ c - a := - Int.add_le_add_left (Int.neg_le_neg h) c - -protected theorem sub_le_sub_right {a b : Int} (h : a ≤ b) (c : Int) : a - c ≤ b - c := - Int.add_le_add_right h (-c) - -protected theorem sub_le_sub {a b c d : Int} (hab : a ≤ b) (hcd : c ≤ d) : a - d ≤ b - c := - Int.add_le_add hab (Int.neg_le_neg hcd) - -protected theorem add_lt_of_lt_neg_add {a b c : Int} (h : b < -a + c) : a + b < c := by - have h := Int.add_lt_add_left h a - rwa [Int.add_neg_cancel_left] at h - -protected theorem lt_neg_add_of_add_lt {a b c : Int} (h : a + b < c) : b < -a + c := by - have h := Int.add_lt_add_left h (-a) - rwa [Int.neg_add_cancel_left] at h - -protected theorem add_lt_of_lt_sub_left {a b c : Int} (h : b < c - a) : a + b < c := by - have h := Int.add_lt_add_left h a - rwa [← Int.add_sub_assoc, Int.add_comm a c, Int.add_sub_cancel] at h - -protected theorem lt_sub_left_of_add_lt {a b c : Int} (h : a + b < c) : b < c - a := by - have h := Int.add_lt_add_right h (-a) - rwa [Int.add_comm a b, Int.add_neg_cancel_right] at h - -protected theorem add_lt_of_lt_sub_right {a b c : Int} (h : a < c - b) : a + b < c := by - have h := Int.add_lt_add_right h b - rwa [Int.sub_add_cancel] at h - -protected theorem lt_sub_right_of_add_lt {a b c : Int} (h : a + b < c) : a < c - b := by - have h := Int.add_lt_add_right h (-b) - rwa [Int.add_neg_cancel_right] at h - -protected theorem lt_add_of_neg_add_lt {a b c : Int} (h : -b + a < c) : a < b + c := by - have h := Int.add_lt_add_left h b - rwa [Int.add_neg_cancel_left] at h - -protected theorem neg_add_lt_of_lt_add {a b c : Int} (h : a < b + c) : -b + a < c := by - have h := Int.add_lt_add_left h (-b) - rwa [Int.neg_add_cancel_left] at h - -protected theorem lt_add_of_sub_left_lt {a b c : Int} (h : a - b < c) : a < b + c := by - have h := Int.add_lt_add_right h b - rwa [Int.sub_add_cancel, Int.add_comm] at h - -protected theorem sub_left_lt_of_lt_add {a b c : Int} (h : a < b + c) : a - b < c := by - have h := Int.add_lt_add_right h (-b) - rwa [Int.add_comm b c, Int.add_neg_cancel_right] at h - -protected theorem lt_add_of_sub_right_lt {a b c : Int} (h : a - c < b) : a < b + c := by - have h := Int.add_lt_add_right h c - rwa [Int.sub_add_cancel] at h - -protected theorem sub_right_lt_of_lt_add {a b c : Int} (h : a < b + c) : a - c < b := by - have h := Int.add_lt_add_right h (-c) - rwa [Int.add_neg_cancel_right] at h - -protected theorem lt_add_of_neg_add_lt_left {a b c : Int} (h : -b + a < c) : a < b + c := by - rw [Int.add_comm] at h - exact Int.lt_add_of_sub_left_lt h - -protected theorem neg_add_lt_left_of_lt_add {a b c : Int} (h : a < b + c) : -b + a < c := by - rw [Int.add_comm] - exact Int.sub_left_lt_of_lt_add h - -protected theorem lt_add_of_neg_add_lt_right {a b c : Int} (h : -c + a < b) : a < b + c := by - rw [Int.add_comm] at h - exact Int.lt_add_of_sub_right_lt h - -protected theorem neg_add_lt_right_of_lt_add {a b c : Int} (h : a < b + c) : -c + a < b := by - rw [Int.add_comm] at h - exact Int.neg_add_lt_left_of_lt_add h - -protected theorem lt_add_of_neg_lt_sub_left {a b c : Int} (h : -a < b - c) : c < a + b := - Int.lt_add_of_neg_add_lt_left (Int.add_lt_of_lt_sub_right h) - -protected theorem neg_lt_sub_left_of_lt_add {a b c : Int} (h : c < a + b) : -a < b - c := by - have h := Int.lt_neg_add_of_add_lt (Int.sub_left_lt_of_lt_add h) - rwa [Int.add_comm] at h - -protected theorem lt_add_of_neg_lt_sub_right {a b c : Int} (h : -b < a - c) : c < a + b := - Int.lt_add_of_sub_right_lt (Int.add_lt_of_lt_sub_left h) - -protected theorem neg_lt_sub_right_of_lt_add {a b c : Int} (h : c < a + b) : -b < a - c := - Int.lt_sub_left_of_add_lt (Int.sub_right_lt_of_lt_add h) - -protected theorem sub_lt_of_sub_lt {a b c : Int} (h : a - b < c) : a - c < b := - Int.sub_left_lt_of_lt_add (Int.lt_add_of_sub_right_lt h) - -protected theorem sub_lt_sub_left {a b : Int} (h : a < b) (c : Int) : c - b < c - a := - Int.add_lt_add_left (Int.neg_lt_neg h) c - -protected theorem sub_lt_sub_right {a b : Int} (h : a < b) (c : Int) : a - c < b - c := - Int.add_lt_add_right h (-c) - -protected theorem sub_lt_sub {a b c d : Int} (hab : a < b) (hcd : c < d) : a - d < b - c := - Int.add_lt_add hab (Int.neg_lt_neg hcd) - -protected theorem sub_lt_sub_of_le_of_lt {a b c d : Int} - (hab : a ≤ b) (hcd : c < d) : a - d < b - c := - Int.add_lt_add_of_le_of_lt hab (Int.neg_lt_neg hcd) - -protected theorem sub_lt_sub_of_lt_of_le {a b c d : Int} - (hab : a < b) (hcd : c ≤ d) : a - d < b - c := - Int.add_lt_add_of_lt_of_le hab (Int.neg_le_neg hcd) - -protected theorem sub_le_self (a : Int) {b : Int} (h : 0 ≤ b) : a - b ≤ a := - calc a + -b - _ ≤ a + 0 := Int.add_le_add_left (Int.neg_nonpos_of_nonneg h) _ - _ = a := by rw [Int.add_zero] - -protected theorem sub_lt_self (a : Int) {b : Int} (h : 0 < b) : a - b < a := - calc a + -b - _ < a + 0 := Int.add_lt_add_left (Int.neg_neg_of_pos h) _ - _ = a := by rw [Int.add_zero] - -protected theorem add_le_add_three {a b c d e f : Int} - (h₁ : a ≤ d) (h₂ : b ≤ e) (h₃ : c ≤ f) : a + b + c ≤ d + e + f := - Int.add_le_add (Int.add_le_add h₁ h₂) h₃ - -protected theorem mul_lt_mul_of_pos_left {a b c : Int} - (h₁ : a < b) (h₂ : 0 < c) : c * a < c * b := by - have : 0 < c * (b - a) := Int.mul_pos h₂ (Int.sub_pos_of_lt h₁) - rw [Int.mul_sub] at this - exact Int.lt_of_sub_pos this - -protected theorem mul_lt_mul_of_pos_right {a b c : Int} - (h₁ : a < b) (h₂ : 0 < c) : a * c < b * c := by - have : 0 < b - a := Int.sub_pos_of_lt h₁ - have : 0 < (b - a) * c := Int.mul_pos this h₂ - rw [Int.sub_mul] at this - exact Int.lt_of_sub_pos this - -protected theorem mul_le_mul_of_nonneg_left {a b c : Int} - (h₁ : a ≤ b) (h₂ : 0 ≤ c) : c * a ≤ c * b := by - if hba : b ≤ a then rw [Int.le_antisymm hba h₁]; apply Int.le_refl else - if hc0 : c ≤ 0 then simp [Int.le_antisymm hc0 h₂, Int.zero_mul] else - exact Int.le_of_lt <| Int.mul_lt_mul_of_pos_left - (Int.lt_iff_le_not_le.2 ⟨h₁, hba⟩) (Int.lt_iff_le_not_le.2 ⟨h₂, hc0⟩) - -protected theorem mul_le_mul_of_nonneg_right {a b c : Int} - (h₁ : a ≤ b) (h₂ : 0 ≤ c) : a * c ≤ b * c := by - rw [Int.mul_comm, Int.mul_comm b]; exact Int.mul_le_mul_of_nonneg_left h₁ h₂ - -protected theorem mul_le_mul {a b c d : Int} - (hac : a ≤ c) (hbd : b ≤ d) (nn_b : 0 ≤ b) (nn_c : 0 ≤ c) : a * b ≤ c * d := - Int.le_trans (Int.mul_le_mul_of_nonneg_right hac nn_b) (Int.mul_le_mul_of_nonneg_left hbd nn_c) - -protected theorem mul_nonpos_of_nonneg_of_nonpos {a b : Int} - (ha : 0 ≤ a) (hb : b ≤ 0) : a * b ≤ 0 := by - have h : a * b ≤ a * 0 := Int.mul_le_mul_of_nonneg_left hb ha - rwa [Int.mul_zero] at h - -protected theorem mul_nonpos_of_nonpos_of_nonneg {a b : Int} - (ha : a ≤ 0) (hb : 0 ≤ b) : a * b ≤ 0 := by - have h : a * b ≤ 0 * b := Int.mul_le_mul_of_nonneg_right ha hb - rwa [Int.zero_mul] at h - -protected theorem mul_lt_mul {a b c d : Int} - (h₁ : a < c) (h₂ : b ≤ d) (h₃ : 0 < b) (h₄ : 0 ≤ c) : a * b < c * d := - Int.lt_of_lt_of_le (Int.mul_lt_mul_of_pos_right h₁ h₃) (Int.mul_le_mul_of_nonneg_left h₂ h₄) - -protected theorem mul_lt_mul' {a b c d : Int} - (h₁ : a ≤ c) (h₂ : b < d) (h₃ : 0 ≤ b) (h₄ : 0 < c) : a * b < c * d := - Int.lt_of_le_of_lt (Int.mul_le_mul_of_nonneg_right h₁ h₃) (Int.mul_lt_mul_of_pos_left h₂ h₄) - -protected theorem mul_neg_of_pos_of_neg {a b : Int} (ha : 0 < a) (hb : b < 0) : a * b < 0 := by - have h : a * b < a * 0 := Int.mul_lt_mul_of_pos_left hb ha - rwa [Int.mul_zero] at h - -protected theorem mul_neg_of_neg_of_pos {a b : Int} (ha : a < 0) (hb : 0 < b) : a * b < 0 := by - have h : a * b < 0 * b := Int.mul_lt_mul_of_pos_right ha hb - rwa [Int.zero_mul] at h - -protected theorem mul_le_mul_of_nonpos_right {a b c : Int} - (h : b ≤ a) (hc : c ≤ 0) : a * c ≤ b * c := - have : -c ≥ 0 := Int.neg_nonneg_of_nonpos hc - have : b * -c ≤ a * -c := Int.mul_le_mul_of_nonneg_right h this - Int.le_of_neg_le_neg <| by rwa [← Int.neg_mul_eq_mul_neg, ← Int.neg_mul_eq_mul_neg] at this - -protected theorem mul_le_mul_of_nonpos_left {a b c : Int} - (ha : a ≤ 0) (h : c ≤ b) : a * b ≤ a * c := by - rw [Int.mul_comm a b, Int.mul_comm a c] - apply Int.mul_le_mul_of_nonpos_right h ha - -protected theorem mul_nonneg_of_nonpos_of_nonpos {a b : Int} - (ha : a ≤ 0) (hb : b ≤ 0) : 0 ≤ a * b := by - have : 0 * b ≤ a * b := Int.mul_le_mul_of_nonpos_right ha hb - rwa [Int.zero_mul] at this - -protected theorem mul_lt_mul_of_neg_left {a b c : Int} (h : b < a) (hc : c < 0) : c * a < c * b := - have : -c > 0 := Int.neg_pos_of_neg hc - have : -c * b < -c * a := Int.mul_lt_mul_of_pos_left h this - have : -(c * b) < -(c * a) := by - rwa [← Int.neg_mul_eq_neg_mul, ← Int.neg_mul_eq_neg_mul] at this - Int.lt_of_neg_lt_neg this - -protected theorem mul_lt_mul_of_neg_right {a b c : Int} (h : b < a) (hc : c < 0) : a * c < b * c := - have : -c > 0 := Int.neg_pos_of_neg hc - have : b * -c < a * -c := Int.mul_lt_mul_of_pos_right h this - have : -(b * c) < -(a * c) := by - rwa [← Int.neg_mul_eq_mul_neg, ← Int.neg_mul_eq_mul_neg] at this - Int.lt_of_neg_lt_neg this - -protected theorem mul_pos_of_neg_of_neg {a b : Int} (ha : a < 0) (hb : b < 0) : 0 < a * b := by - have : 0 * b < a * b := Int.mul_lt_mul_of_neg_right ha hb - rwa [Int.zero_mul] at this - -protected theorem mul_self_le_mul_self {a b : Int} (h1 : 0 ≤ a) (h2 : a ≤ b) : a * a ≤ b * b := - Int.mul_le_mul h2 h2 h1 (Int.le_trans h1 h2) - -protected theorem mul_self_lt_mul_self {a b : Int} (h1 : 0 ≤ a) (h2 : a < b) : a * a < b * b := - Int.mul_lt_mul' (Int.le_of_lt h2) h2 h1 (Int.lt_of_le_of_lt h1 h2) - -theorem exists_eq_neg_ofNat {a : Int} (H : a ≤ 0) : ∃ n : Nat, a = -(n : Int) := - let ⟨n, h⟩ := eq_ofNat_of_zero_le (Int.neg_nonneg_of_nonpos H) - ⟨n, Int.eq_neg_of_eq_neg h.symm⟩ - -theorem natAbs_of_nonneg {a : Int} (H : 0 ≤ a) : (natAbs a : Int) = a := - match a, eq_ofNat_of_zero_le H with - | _, ⟨_, rfl⟩ => rfl - -theorem ofNat_natAbs_of_nonpos {a : Int} (H : a ≤ 0) : (natAbs a : Int) = -a := by - rw [← natAbs_neg, natAbs_of_nonneg (Int.neg_nonneg_of_nonpos H)] - -theorem lt_of_add_one_le {a b : Int} (H : a + 1 ≤ b) : a < b := H - -theorem add_one_le_of_lt {a b : Int} (H : a < b) : a + 1 ≤ b := H - -theorem lt_add_one_of_le {a b : Int} (H : a ≤ b) : a < b + 1 := Int.add_le_add_right H 1 - -theorem le_of_lt_add_one {a b : Int} (H : a < b + 1) : a ≤ b := Int.le_of_add_le_add_right H - -theorem sub_one_lt_of_le {a b : Int} (H : a ≤ b) : a - 1 < b := - Int.sub_right_lt_of_lt_add <| lt_add_one_of_le H - -theorem le_of_sub_one_lt {a b : Int} (H : a - 1 < b) : a ≤ b := - le_of_lt_add_one <| Int.lt_add_of_sub_right_lt H - -theorem le_sub_one_of_lt {a b : Int} (H : a < b) : a ≤ b - 1 := Int.le_sub_right_of_add_le H - -theorem lt_of_le_sub_one {a b : Int} (H : a ≤ b - 1) : a < b := Int.add_le_of_le_sub_right H - -protected theorem mul_eq_zero {a b : Int} : a * b = 0 ↔ a = 0 ∨ b = 0 := by - refine ⟨fun h => ?_, fun h => h.elim (by simp [·, Int.zero_mul]) (by simp [·, Int.mul_zero])⟩ - exact match Int.lt_trichotomy a 0, Int.lt_trichotomy b 0 with - | .inr (.inl heq₁), _ => .inl heq₁ - | _, .inr (.inl heq₂) => .inr heq₂ - | .inl hlt₁, .inl hlt₂ => absurd h <| Int.ne_of_gt <| Int.mul_pos_of_neg_of_neg hlt₁ hlt₂ - | .inl hlt₁, .inr (.inr hgt₂) => absurd h <| Int.ne_of_lt <| Int.mul_neg_of_neg_of_pos hlt₁ hgt₂ - | .inr (.inr hgt₁), .inl hlt₂ => absurd h <| Int.ne_of_lt <| Int.mul_neg_of_pos_of_neg hgt₁ hlt₂ - | .inr (.inr hgt₁), .inr (.inr hgt₂) => absurd h <| Int.ne_of_gt <| Int.mul_pos hgt₁ hgt₂ - -protected theorem mul_ne_zero {a b : Int} (a0 : a ≠ 0) (b0 : b ≠ 0) : a * b ≠ 0 := - mt Int.mul_eq_zero.1 <| not_or.2 ⟨a0, b0⟩ - -protected theorem eq_of_mul_eq_mul_right {a b c : Int} (ha : a ≠ 0) (h : b * a = c * a) : b = c := - have : (b - c) * a = 0 := by rwa [Int.sub_mul, Int.sub_eq_zero] - Int.sub_eq_zero.1 <| (Int.mul_eq_zero.1 this).resolve_right ha - -protected theorem eq_of_mul_eq_mul_left {a b c : Int} (ha : a ≠ 0) (h : a * b = a * c) : b = c := - have : a * b - a * c = 0 := Int.sub_eq_zero_of_eq h - have : a * (b - c) = 0 := by rw [Int.mul_sub, this] - have : b - c = 0 := (Int.mul_eq_zero.1 this).resolve_left ha - Int.eq_of_sub_eq_zero this - -theorem mul_eq_mul_left_iff {a b c : Int} (h : c ≠ 0) : c * a = c * b ↔ a = b := - ⟨Int.eq_of_mul_eq_mul_left h, fun w => congrArg (fun x => c * x) w⟩ - -theorem mul_eq_mul_right_iff {a b c : Int} (h : c ≠ 0) : a * c = b * c ↔ a = b := - ⟨Int.eq_of_mul_eq_mul_right h, fun w => congrArg (fun x => x * c) w⟩ - -theorem eq_one_of_mul_eq_self_left {a b : Int} (Hpos : a ≠ 0) (H : b * a = a) : b = 1 := - Int.eq_of_mul_eq_mul_right Hpos <| by rw [Int.one_mul, H] - -theorem eq_one_of_mul_eq_self_right {a b : Int} (Hpos : b ≠ 0) (H : b * a = b) : a = 1 := - Int.eq_of_mul_eq_mul_left Hpos <| by rw [Int.mul_one, H] - -/- ## sign -/ - -@[simp] theorem sign_zero : sign 0 = 0 := rfl -@[simp] theorem sign_one : sign 1 = 1 := rfl -theorem sign_neg_one : sign (-1) = -1 := rfl - -@[simp] theorem sign_of_add_one (x : Nat) : Int.sign (x + 1) = 1 := rfl -@[simp] theorem sign_negSucc (x : Nat) : Int.sign (Int.negSucc x) = -1 := rfl - -theorem natAbs_sign (z : Int) : z.sign.natAbs = if z = 0 then 0 else 1 := - match z with | 0 | succ _ | -[_+1] => rfl - -theorem natAbs_sign_of_nonzero {z : Int} (hz : z ≠ 0) : z.sign.natAbs = 1 := by - rw [Int.natAbs_sign, if_neg hz] - -theorem sign_ofNat_of_nonzero {n : Nat} (hn : n ≠ 0) : Int.sign n = 1 := - match n, Nat.exists_eq_succ_of_ne_zero hn with - | _, ⟨n, rfl⟩ => Int.sign_of_add_one n - -@[simp] theorem sign_neg (z : Int) : Int.sign (-z) = -Int.sign z := by - match z with | 0 | succ _ | -[_+1] => rfl - -theorem sign_mul_natAbs : ∀ a : Int, sign a * natAbs a = a - | 0 => rfl - | succ _ => Int.one_mul _ - | -[_+1] => (Int.neg_eq_neg_one_mul _).symm - -@[simp] theorem sign_mul : ∀ a b, sign (a * b) = sign a * sign b - | a, 0 | 0, b => by simp [Int.mul_zero, Int.zero_mul] - | succ _, succ _ | succ _, -[_+1] | -[_+1], succ _ | -[_+1], -[_+1] => rfl - -theorem sign_eq_one_of_pos {a : Int} (h : 0 < a) : sign a = 1 := - match a, eq_succ_of_zero_lt h with - | _, ⟨_, rfl⟩ => rfl - -theorem sign_eq_neg_one_of_neg {a : Int} (h : a < 0) : sign a = -1 := - match a, eq_negSucc_of_lt_zero h with - | _, ⟨_, rfl⟩ => rfl - -theorem eq_zero_of_sign_eq_zero : ∀ {a : Int}, sign a = 0 → a = 0 - | 0, _ => rfl - -theorem pos_of_sign_eq_one : ∀ {a : Int}, sign a = 1 → 0 < a - | (_ + 1 : Nat), _ => ofNat_lt.2 (Nat.succ_pos _) - -theorem neg_of_sign_eq_neg_one : ∀ {a : Int}, sign a = -1 → a < 0 - | (_ + 1 : Nat), h => nomatch h - | 0, h => nomatch h - | -[_+1], _ => negSucc_lt_zero _ - -theorem sign_eq_one_iff_pos (a : Int) : sign a = 1 ↔ 0 < a := - ⟨pos_of_sign_eq_one, sign_eq_one_of_pos⟩ - -theorem sign_eq_neg_one_iff_neg (a : Int) : sign a = -1 ↔ a < 0 := - ⟨neg_of_sign_eq_neg_one, sign_eq_neg_one_of_neg⟩ - -@[simp] theorem sign_eq_zero_iff_zero (a : Int) : sign a = 0 ↔ a = 0 := - ⟨eq_zero_of_sign_eq_zero, fun h => by rw [h, sign_zero]⟩ - -@[simp] theorem sign_sign : sign (sign x) = sign x := by - match x with - | 0 => rfl - | .ofNat (_ + 1) => rfl - | .negSucc _ => rfl - -@[simp] theorem sign_nonneg : 0 ≤ sign x ↔ 0 ≤ x := by - match x with - | 0 => rfl - | .ofNat (_ + 1) => - simp (config := { decide := true }) only [sign, true_iff] - exact Int.le_add_one (ofNat_nonneg _) - | .negSucc _ => simp (config := { decide := true }) [sign] - - -/-! ### nat abs -/ - -@[deprecated] alias ofNat_natAbs_eq_of_nonneg := natAbs_of_nonneg - -theorem eq_natAbs_iff_mul_eq_zero : natAbs a = n ↔ (a - n) * (a + n) = 0 := by - rw [natAbs_eq_iff, Int.mul_eq_zero, ← Int.sub_neg, Int.sub_eq_zero, Int.sub_eq_zero] - -theorem natAbs_add_le (a b : Int) : natAbs (a + b) ≤ natAbs a + natAbs b := by - suffices ∀ a b : Nat, natAbs (subNatNat a b.succ) ≤ (a + b).succ by - match a, b with - | (a:Nat), (b:Nat) => simp - | (a:Nat), -[b+1] => simp; apply this - | -[a+1], (b:Nat) => simp [Nat.succ_add, Nat.add_comm a b]; apply this - | -[a+1], -[b+1] => simp [Nat.succ_add]; apply Nat.le_refl - refine fun a b => subNatNat_elim a b.succ - (fun m n i => n = b.succ → natAbs i ≤ (m + b).succ) ?_ - (fun i n (e : (n + i).succ = _) => ?_) rfl - · rintro i n rfl - rw [Nat.add_comm _ i, Nat.add_assoc] - exact Nat.le_add_right i (b.succ + b).succ - · apply succ_le_succ - rw [← succ.inj e, ← Nat.add_assoc, Nat.add_comm] - apply Nat.le_add_right - -theorem natAbs_sub_le (a b : Int) : natAbs (a - b) ≤ natAbs a + natAbs b := by - rw [← Int.natAbs_neg b]; apply natAbs_add_le - -theorem negSucc_eq' (m : Nat) : -[m+1] = -m - 1 := by simp only [negSucc_eq, Int.neg_add]; rfl - -theorem natAbs_lt_natAbs_of_nonneg_of_lt {a b : Int} - (w₁ : 0 ≤ a) (w₂ : a < b) : a.natAbs < b.natAbs := - match a, b, eq_ofNat_of_zero_le w₁, eq_ofNat_of_zero_le (Int.le_trans w₁ (Int.le_of_lt w₂)) with - | _, _, ⟨_, rfl⟩, ⟨_, rfl⟩ => ofNat_lt.1 w₂ - -/-! ### toNat -/ - -theorem toNat_eq_max : ∀ a : Int, (toNat a : Int) = max a 0 - | (n : Nat) => (Int.max_eq_left (ofNat_zero_le n)).symm - | -[n+1] => (Int.max_eq_right (Int.le_of_lt (negSucc_lt_zero n))).symm - -@[simp] theorem toNat_zero : (0 : Int).toNat = 0 := rfl - -@[simp] theorem toNat_one : (1 : Int).toNat = 1 := rfl - -@[simp] theorem toNat_of_nonneg {a : Int} (h : 0 ≤ a) : (toNat a : Int) = a := by - rw [toNat_eq_max, Int.max_eq_left h] - -@[simp] theorem toNat_ofNat (n : Nat) : toNat ↑n = n := rfl - -@[simp] theorem toNat_ofNat_add_one {n : Nat} : ((n : Int) + 1).toNat = n + 1 := rfl - -theorem self_le_toNat (a : Int) : a ≤ toNat a := by rw [toNat_eq_max]; apply Int.le_max_left - -@[simp] theorem le_toNat {n : Nat} {z : Int} (h : 0 ≤ z) : n ≤ z.toNat ↔ (n : Int) ≤ z := by - rw [← Int.ofNat_le, Int.toNat_of_nonneg h] - -@[simp] theorem toNat_lt {n : Nat} {z : Int} (h : 0 ≤ z) : z.toNat < n ↔ z < (n : Int) := by - rw [← Int.not_le, ← Nat.not_le, Int.le_toNat h] - -theorem toNat_add {a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : (a + b).toNat = a.toNat + b.toNat := - match a, b, eq_ofNat_of_zero_le ha, eq_ofNat_of_zero_le hb with - | _, _, ⟨_, rfl⟩, ⟨_, rfl⟩ => rfl - -theorem toNat_add_nat {a : Int} (ha : 0 ≤ a) (n : Nat) : (a + n).toNat = a.toNat + n := - match a, eq_ofNat_of_zero_le ha with | _, ⟨_, rfl⟩ => rfl - -@[simp] theorem pred_toNat : ∀ i : Int, (i - 1).toNat = i.toNat - 1 - | 0 => rfl - | (n+1:Nat) => by simp [ofNat_add] - | -[n+1] => rfl - -@[simp] theorem toNat_sub_toNat_neg : ∀ n : Int, ↑n.toNat - ↑(-n).toNat = n - | 0 => rfl - | (_+1:Nat) => Int.sub_zero _ - | -[_+1] => Int.zero_sub _ - -@[simp] theorem toNat_add_toNat_neg_eq_natAbs : ∀ n : Int, n.toNat + (-n).toNat = n.natAbs - | 0 => rfl - | (_+1:Nat) => Nat.add_zero _ - | -[_+1] => Nat.zero_add _ - -theorem mem_toNat' : ∀ (a : Int) (n : Nat), toNat' a = some n ↔ a = n - | (m : Nat), n => Option.some_inj.trans ofNat_inj.symm - | -[m+1], n => by constructor <;> intro. - -@[simp] theorem toNat_neg_nat : ∀ n : Nat, (-(n : Int)).toNat = 0 - | 0 => rfl - | _+1 => rfl - -/-! -The following lemmas are later subsumed by e.g. `Nat.cast_add` and `Nat.cast_mul` in Mathlib -but it is convenient to have these earlier, for users who only need `Nat` and `Int`. --/ - -theorem natCast_zero : ((0 : Nat) : Int) = (0 : Int) := rfl - -theorem natCast_one : ((1 : Nat) : Int) = (1 : Int) := rfl - -@[simp] theorem natCast_add (a b : Nat) : ((a + b : Nat) : Int) = (a : Int) + (b : Int) := by - -- Note this only works because of local simp attributes in this file, - -- so it still makes sense to tag the lemmas with `@[simp]`. - simp - -@[simp] theorem natCast_mul (a b : Nat) : ((a * b : Nat) : Int) = (a : Int) * (b : Int) := by - simp diff --git a/Std/Data/Int/Order.lean b/Std/Data/Int/Order.lean new file mode 100644 index 0000000000..15912c01df --- /dev/null +++ b/Std/Data/Int/Order.lean @@ -0,0 +1,593 @@ +/- +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 Std.Data.Int.Init.Order +import Std.Data.Option.Basic +import Std.Tactic.RCases +import Std.Tactic.ByCases +import Std.Tactic.Omega + +/-! +# Results about the order properties of the integers, and the integers as an ordered ring. +-/ + +open Nat + +namespace Int + +/-! ## Order properties of the integers -/ + +protected theorem le_of_not_le {a b : Int} : ¬ a ≤ b → b ≤ a := (Int.le_total a b).resolve_left + +protected theorem min_eq_left {a b : Int} (h : a ≤ b) : min a b = a := by simp [Int.min_def, h] + +protected theorem min_eq_right {a b : Int} (h : b ≤ a) : min a b = b := by + rw [Int.min_comm a b]; exact Int.min_eq_left h + +protected theorem max_eq_right {a b : Int} (h : a ≤ b) : max a b = b := by + simp [Int.max_def, h, Int.not_lt.2 h] + +protected theorem max_eq_left {a b : Int} (h : b ≤ a) : max a b = a := by + rw [← Int.max_comm b a]; exact Int.max_eq_right h + +@[simp] theorem negSucc_not_pos (n : Nat) : 0 < -[n+1] ↔ False := by + simp only [Int.not_lt, iff_false]; constructor + +theorem eq_negSucc_of_lt_zero : ∀ {a : Int}, a < 0 → ∃ n : Nat, a = -[n+1] + | ofNat _, h => absurd h (Int.not_lt.2 (ofNat_zero_le _)) + | -[n+1], _ => ⟨n, rfl⟩ + +protected theorem lt_of_add_lt_add_left {a b c : Int} (h : a + b < a + c) : b < c := by + have : -a + (a + b) < -a + (a + c) := Int.add_lt_add_left h _ + simp [Int.neg_add_cancel_left] at this + assumption + +protected theorem lt_of_add_lt_add_right {a b c : Int} (h : a + b < c + b) : a < c := + Int.lt_of_add_lt_add_left (a := b) <| by rwa [Int.add_comm b a, Int.add_comm b c] + +protected theorem add_lt_add_iff_left (a : Int) : a + b < a + c ↔ b < c := + ⟨Int.lt_of_add_lt_add_left, (Int.add_lt_add_left · _)⟩ + +protected theorem add_lt_add_iff_right (c : Int) : a + c < b + c ↔ a < b := + ⟨Int.lt_of_add_lt_add_right, (Int.add_lt_add_right · _)⟩ + +protected theorem add_lt_add {a b c d : Int} (h₁ : a < b) (h₂ : c < d) : a + c < b + d := + Int.lt_trans (Int.add_lt_add_right h₁ c) (Int.add_lt_add_left h₂ b) + +protected theorem add_lt_add_of_le_of_lt {a b c d : Int} (h₁ : a ≤ b) (h₂ : c < d) : + a + c < b + d := + Int.lt_of_le_of_lt (Int.add_le_add_right h₁ c) (Int.add_lt_add_left h₂ b) + +protected theorem add_lt_add_of_lt_of_le {a b c d : Int} (h₁ : a < b) (h₂ : c ≤ d) : + a + c < b + d := + Int.lt_of_lt_of_le (Int.add_lt_add_right h₁ c) (Int.add_le_add_left h₂ b) + +protected theorem lt_add_of_pos_right (a : Int) {b : Int} (h : 0 < b) : a < a + b := by + have : a + 0 < a + b := Int.add_lt_add_left h a + rwa [Int.add_zero] at this + +protected theorem lt_add_of_pos_left (a : Int) {b : Int} (h : 0 < b) : a < b + a := by + have : 0 + a < b + a := Int.add_lt_add_right h a + rwa [Int.zero_add] at this + +protected theorem add_nonneg {a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a + b := + Int.zero_add 0 ▸ Int.add_le_add ha hb + +protected theorem add_pos {a b : Int} (ha : 0 < a) (hb : 0 < b) : 0 < a + b := + Int.zero_add 0 ▸ Int.add_lt_add ha hb + +protected theorem add_pos_of_pos_of_nonneg {a b : Int} (ha : 0 < a) (hb : 0 ≤ b) : 0 < a + b := + Int.zero_add 0 ▸ Int.add_lt_add_of_lt_of_le ha hb + +protected theorem add_pos_of_nonneg_of_pos {a b : Int} (ha : 0 ≤ a) (hb : 0 < b) : 0 < a + b := + Int.zero_add 0 ▸ Int.add_lt_add_of_le_of_lt ha hb + +protected theorem add_nonpos {a b : Int} (ha : a ≤ 0) (hb : b ≤ 0) : a + b ≤ 0 := + Int.zero_add 0 ▸ Int.add_le_add ha hb + +protected theorem add_neg {a b : Int} (ha : a < 0) (hb : b < 0) : a + b < 0 := + Int.zero_add 0 ▸ Int.add_lt_add ha hb + +protected theorem add_neg_of_neg_of_nonpos {a b : Int} (ha : a < 0) (hb : b ≤ 0) : a + b < 0 := + Int.zero_add 0 ▸ Int.add_lt_add_of_lt_of_le ha hb + +protected theorem add_neg_of_nonpos_of_neg {a b : Int} (ha : a ≤ 0) (hb : b < 0) : a + b < 0 := + Int.zero_add 0 ▸ Int.add_lt_add_of_le_of_lt ha hb + +protected theorem lt_add_of_le_of_pos {a b c : Int} (hbc : b ≤ c) (ha : 0 < a) : b < c + a := + Int.add_zero b ▸ Int.add_lt_add_of_le_of_lt hbc ha + +theorem add_one_le_iff {a b : Int} : a + 1 ≤ b ↔ a < b := .rfl + +theorem lt_add_one_iff {a b : Int} : a < b + 1 ↔ a ≤ b := Int.add_le_add_iff_right _ + +@[simp] theorem succ_ofNat_pos (n : Nat) : 0 < (n : Int) + 1 := + lt_add_one_iff.2 (ofNat_zero_le _) + +theorem le_add_one {a b : Int} (h : a ≤ b) : a ≤ b + 1 := + Int.le_of_lt (Int.lt_add_one_iff.2 h) + +protected theorem nonneg_of_neg_nonpos {a : Int} (h : -a ≤ 0) : 0 ≤ a := + Int.le_of_neg_le_neg <| by rwa [Int.neg_zero] + +protected theorem nonpos_of_neg_nonneg {a : Int} (h : 0 ≤ -a) : a ≤ 0 := + Int.le_of_neg_le_neg <| by rwa [Int.neg_zero] + +protected theorem lt_of_neg_lt_neg {a b : Int} (h : -b < -a) : a < b := + Int.neg_neg a ▸ Int.neg_neg b ▸ Int.neg_lt_neg h + +protected theorem pos_of_neg_neg {a : Int} (h : -a < 0) : 0 < a := + Int.lt_of_neg_lt_neg <| by rwa [Int.neg_zero] + +protected theorem neg_of_neg_pos {a : Int} (h : 0 < -a) : a < 0 := + have : -0 < -a := by rwa [Int.neg_zero] + Int.lt_of_neg_lt_neg this + +protected theorem le_neg_of_le_neg {a b : Int} (h : a ≤ -b) : b ≤ -a := by + have h := Int.neg_le_neg h + rwa [Int.neg_neg] at h + +protected theorem neg_le_of_neg_le {a b : Int} (h : -a ≤ b) : -b ≤ a := by + have h := Int.neg_le_neg h + rwa [Int.neg_neg] at h + +protected theorem lt_neg_of_lt_neg {a b : Int} (h : a < -b) : b < -a := by + have h := Int.neg_lt_neg h + rwa [Int.neg_neg] at h + +protected theorem neg_lt_of_neg_lt {a b : Int} (h : -a < b) : -b < a := by + have h := Int.neg_lt_neg h + rwa [Int.neg_neg] at h + +protected theorem sub_nonpos_of_le {a b : Int} (h : a ≤ b) : a - b ≤ 0 := by + have h := Int.add_le_add_right h (-b) + rwa [Int.add_right_neg] at h + +protected theorem le_of_sub_nonpos {a b : Int} (h : a - b ≤ 0) : a ≤ b := by + have h := Int.add_le_add_right h b + rwa [Int.sub_add_cancel, Int.zero_add] at h + +protected theorem sub_neg_of_lt {a b : Int} (h : a < b) : a - b < 0 := by + have h := Int.add_lt_add_right h (-b) + rwa [Int.add_right_neg] at h + +protected theorem lt_of_sub_neg {a b : Int} (h : a - b < 0) : a < b := by + have h := Int.add_lt_add_right h b + rwa [Int.sub_add_cancel, Int.zero_add] at h + +protected theorem add_le_of_le_neg_add {a b c : Int} (h : b ≤ -a + c) : a + b ≤ c := by + have h := Int.add_le_add_left h a + rwa [Int.add_neg_cancel_left] at h + +protected theorem le_neg_add_of_add_le {a b c : Int} (h : a + b ≤ c) : b ≤ -a + c := by + have h := Int.add_le_add_left h (-a) + rwa [Int.neg_add_cancel_left] at h + +protected theorem add_le_of_le_sub_left {a b c : Int} (h : b ≤ c - a) : a + b ≤ c := by + have h := Int.add_le_add_left h a + rwa [← Int.add_sub_assoc, Int.add_comm a c, Int.add_sub_cancel] at h + +protected theorem le_sub_left_of_add_le {a b c : Int} (h : a + b ≤ c) : b ≤ c - a := by + have h := Int.add_le_add_right h (-a) + rwa [Int.add_comm a b, Int.add_neg_cancel_right] at h + +protected theorem add_le_of_le_sub_right {a b c : Int} (h : a ≤ c - b) : a + b ≤ c := by + have h := Int.add_le_add_right h b + rwa [Int.sub_add_cancel] at h + +protected theorem le_sub_right_of_add_le {a b c : Int} (h : a + b ≤ c) : a ≤ c - b := by + have h := Int.add_le_add_right h (-b) + rwa [Int.add_neg_cancel_right] at h + +protected theorem le_add_of_neg_add_le {a b c : Int} (h : -b + a ≤ c) : a ≤ b + c := by + have h := Int.add_le_add_left h b + rwa [Int.add_neg_cancel_left] at h + +protected theorem neg_add_le_of_le_add {a b c : Int} (h : a ≤ b + c) : -b + a ≤ c := by + have h := Int.add_le_add_left h (-b) + rwa [Int.neg_add_cancel_left] at h + +protected theorem le_add_of_sub_left_le {a b c : Int} (h : a - b ≤ c) : a ≤ b + c := by + have h := Int.add_le_add_right h b + rwa [Int.sub_add_cancel, Int.add_comm] at h + +protected theorem le_add_of_sub_right_le {a b c : Int} (h : a - c ≤ b) : a ≤ b + c := by + have h := Int.add_le_add_right h c + rwa [Int.sub_add_cancel] at h + +protected theorem sub_right_le_of_le_add {a b c : Int} (h : a ≤ b + c) : a - c ≤ b := by + have h := Int.add_le_add_right h (-c) + rwa [Int.add_neg_cancel_right] at h + +protected theorem le_add_of_neg_add_le_left {a b c : Int} (h : -b + a ≤ c) : a ≤ b + c := by + rw [Int.add_comm] at h + exact Int.le_add_of_sub_left_le h + +protected theorem neg_add_le_left_of_le_add {a b c : Int} (h : a ≤ b + c) : -b + a ≤ c := by + rw [Int.add_comm] + exact Int.sub_left_le_of_le_add h + +protected theorem le_add_of_neg_add_le_right {a b c : Int} (h : -c + a ≤ b) : a ≤ b + c := by + rw [Int.add_comm] at h + exact Int.le_add_of_sub_right_le h + +protected theorem neg_add_le_right_of_le_add {a b c : Int} (h : a ≤ b + c) : -c + a ≤ b := by + rw [Int.add_comm] at h + exact Int.neg_add_le_left_of_le_add h + +protected theorem le_add_of_neg_le_sub_left {a b c : Int} (h : -a ≤ b - c) : c ≤ a + b := + Int.le_add_of_neg_add_le_left (Int.add_le_of_le_sub_right h) + +protected theorem neg_le_sub_left_of_le_add {a b c : Int} (h : c ≤ a + b) : -a ≤ b - c := by + have h := Int.le_neg_add_of_add_le (Int.sub_left_le_of_le_add h) + rwa [Int.add_comm] at h + +protected theorem le_add_of_neg_le_sub_right {a b c : Int} (h : -b ≤ a - c) : c ≤ a + b := + Int.le_add_of_sub_right_le (Int.add_le_of_le_sub_left h) + +protected theorem neg_le_sub_right_of_le_add {a b c : Int} (h : c ≤ a + b) : -b ≤ a - c := + Int.le_sub_left_of_add_le (Int.sub_right_le_of_le_add h) + +protected theorem sub_le_of_sub_le {a b c : Int} (h : a - b ≤ c) : a - c ≤ b := + Int.sub_left_le_of_le_add (Int.le_add_of_sub_right_le h) + +protected theorem sub_le_sub_left {a b : Int} (h : a ≤ b) (c : Int) : c - b ≤ c - a := + Int.add_le_add_left (Int.neg_le_neg h) c + +protected theorem sub_le_sub_right {a b : Int} (h : a ≤ b) (c : Int) : a - c ≤ b - c := + Int.add_le_add_right h (-c) + +protected theorem sub_le_sub {a b c d : Int} (hab : a ≤ b) (hcd : c ≤ d) : a - d ≤ b - c := + Int.add_le_add hab (Int.neg_le_neg hcd) + +protected theorem add_lt_of_lt_neg_add {a b c : Int} (h : b < -a + c) : a + b < c := by + have h := Int.add_lt_add_left h a + rwa [Int.add_neg_cancel_left] at h + +protected theorem lt_neg_add_of_add_lt {a b c : Int} (h : a + b < c) : b < -a + c := by + have h := Int.add_lt_add_left h (-a) + rwa [Int.neg_add_cancel_left] at h + +protected theorem add_lt_of_lt_sub_left {a b c : Int} (h : b < c - a) : a + b < c := by + have h := Int.add_lt_add_left h a + rwa [← Int.add_sub_assoc, Int.add_comm a c, Int.add_sub_cancel] at h + +protected theorem lt_sub_left_of_add_lt {a b c : Int} (h : a + b < c) : b < c - a := by + have h := Int.add_lt_add_right h (-a) + rwa [Int.add_comm a b, Int.add_neg_cancel_right] at h + +protected theorem add_lt_of_lt_sub_right {a b c : Int} (h : a < c - b) : a + b < c := by + have h := Int.add_lt_add_right h b + rwa [Int.sub_add_cancel] at h + +protected theorem lt_sub_right_of_add_lt {a b c : Int} (h : a + b < c) : a < c - b := by + have h := Int.add_lt_add_right h (-b) + rwa [Int.add_neg_cancel_right] at h + +protected theorem lt_add_of_neg_add_lt {a b c : Int} (h : -b + a < c) : a < b + c := by + have h := Int.add_lt_add_left h b + rwa [Int.add_neg_cancel_left] at h + +protected theorem neg_add_lt_of_lt_add {a b c : Int} (h : a < b + c) : -b + a < c := by + have h := Int.add_lt_add_left h (-b) + rwa [Int.neg_add_cancel_left] at h + +protected theorem lt_add_of_sub_left_lt {a b c : Int} (h : a - b < c) : a < b + c := by + have h := Int.add_lt_add_right h b + rwa [Int.sub_add_cancel, Int.add_comm] at h + +protected theorem sub_left_lt_of_lt_add {a b c : Int} (h : a < b + c) : a - b < c := by + have h := Int.add_lt_add_right h (-b) + rwa [Int.add_comm b c, Int.add_neg_cancel_right] at h + +protected theorem lt_add_of_sub_right_lt {a b c : Int} (h : a - c < b) : a < b + c := by + have h := Int.add_lt_add_right h c + rwa [Int.sub_add_cancel] at h + +protected theorem sub_right_lt_of_lt_add {a b c : Int} (h : a < b + c) : a - c < b := by + have h := Int.add_lt_add_right h (-c) + rwa [Int.add_neg_cancel_right] at h + +protected theorem lt_add_of_neg_add_lt_left {a b c : Int} (h : -b + a < c) : a < b + c := by + rw [Int.add_comm] at h + exact Int.lt_add_of_sub_left_lt h + +protected theorem neg_add_lt_left_of_lt_add {a b c : Int} (h : a < b + c) : -b + a < c := by + rw [Int.add_comm] + exact Int.sub_left_lt_of_lt_add h + +protected theorem lt_add_of_neg_add_lt_right {a b c : Int} (h : -c + a < b) : a < b + c := by + rw [Int.add_comm] at h + exact Int.lt_add_of_sub_right_lt h + +protected theorem neg_add_lt_right_of_lt_add {a b c : Int} (h : a < b + c) : -c + a < b := by + rw [Int.add_comm] at h + exact Int.neg_add_lt_left_of_lt_add h + +protected theorem lt_add_of_neg_lt_sub_left {a b c : Int} (h : -a < b - c) : c < a + b := + Int.lt_add_of_neg_add_lt_left (Int.add_lt_of_lt_sub_right h) + +protected theorem neg_lt_sub_left_of_lt_add {a b c : Int} (h : c < a + b) : -a < b - c := by + have h := Int.lt_neg_add_of_add_lt (Int.sub_left_lt_of_lt_add h) + rwa [Int.add_comm] at h + +protected theorem lt_add_of_neg_lt_sub_right {a b c : Int} (h : -b < a - c) : c < a + b := + Int.lt_add_of_sub_right_lt (Int.add_lt_of_lt_sub_left h) + +protected theorem neg_lt_sub_right_of_lt_add {a b c : Int} (h : c < a + b) : -b < a - c := + Int.lt_sub_left_of_add_lt (Int.sub_right_lt_of_lt_add h) + +protected theorem sub_lt_of_sub_lt {a b c : Int} (h : a - b < c) : a - c < b := + Int.sub_left_lt_of_lt_add (Int.lt_add_of_sub_right_lt h) + +protected theorem sub_lt_sub_left {a b : Int} (h : a < b) (c : Int) : c - b < c - a := + Int.add_lt_add_left (Int.neg_lt_neg h) c + +protected theorem sub_lt_sub_right {a b : Int} (h : a < b) (c : Int) : a - c < b - c := + Int.add_lt_add_right h (-c) + +protected theorem sub_lt_sub {a b c d : Int} (hab : a < b) (hcd : c < d) : a - d < b - c := + Int.add_lt_add hab (Int.neg_lt_neg hcd) + +protected theorem sub_lt_sub_of_le_of_lt {a b c d : Int} + (hab : a ≤ b) (hcd : c < d) : a - d < b - c := + Int.add_lt_add_of_le_of_lt hab (Int.neg_lt_neg hcd) + +protected theorem sub_lt_sub_of_lt_of_le {a b c d : Int} + (hab : a < b) (hcd : c ≤ d) : a - d < b - c := + Int.add_lt_add_of_lt_of_le hab (Int.neg_le_neg hcd) + +protected theorem add_le_add_three {a b c d e f : Int} + (h₁ : a ≤ d) (h₂ : b ≤ e) (h₃ : c ≤ f) : a + b + c ≤ d + e + f := + Int.add_le_add (Int.add_le_add h₁ h₂) h₃ + +theorem exists_eq_neg_ofNat {a : Int} (H : a ≤ 0) : ∃ n : Nat, a = -(n : Int) := + let ⟨n, h⟩ := eq_ofNat_of_zero_le (Int.neg_nonneg_of_nonpos H) + ⟨n, Int.eq_neg_of_eq_neg h.symm⟩ + +theorem lt_of_add_one_le {a b : Int} (H : a + 1 ≤ b) : a < b := H + +theorem lt_add_one_of_le {a b : Int} (H : a ≤ b) : a < b + 1 := Int.add_le_add_right H 1 + +theorem le_of_lt_add_one {a b : Int} (H : a < b + 1) : a ≤ b := Int.le_of_add_le_add_right H + +theorem sub_one_lt_of_le {a b : Int} (H : a ≤ b) : a - 1 < b := + Int.sub_right_lt_of_lt_add <| lt_add_one_of_le H + +theorem le_of_sub_one_lt {a b : Int} (H : a - 1 < b) : a ≤ b := + le_of_lt_add_one <| Int.lt_add_of_sub_right_lt H + +theorem le_sub_one_of_lt {a b : Int} (H : a < b) : a ≤ b - 1 := Int.le_sub_right_of_add_le H + +theorem lt_of_le_sub_one {a b : Int} (H : a ≤ b - 1) : a < b := Int.add_le_of_le_sub_right H + +/- ### Order properties and multiplication -/ + +protected theorem mul_lt_mul {a b c d : Int} + (h₁ : a < c) (h₂ : b ≤ d) (h₃ : 0 < b) (h₄ : 0 ≤ c) : a * b < c * d := + Int.lt_of_lt_of_le (Int.mul_lt_mul_of_pos_right h₁ h₃) (Int.mul_le_mul_of_nonneg_left h₂ h₄) + +protected theorem mul_lt_mul' {a b c d : Int} + (h₁ : a ≤ c) (h₂ : b < d) (h₃ : 0 ≤ b) (h₄ : 0 < c) : a * b < c * d := + Int.lt_of_le_of_lt (Int.mul_le_mul_of_nonneg_right h₁ h₃) (Int.mul_lt_mul_of_pos_left h₂ h₄) + +protected theorem mul_neg_of_pos_of_neg {a b : Int} (ha : 0 < a) (hb : b < 0) : a * b < 0 := by + have h : a * b < a * 0 := Int.mul_lt_mul_of_pos_left hb ha + rwa [Int.mul_zero] at h + +protected theorem mul_neg_of_neg_of_pos {a b : Int} (ha : a < 0) (hb : 0 < b) : a * b < 0 := by + have h : a * b < 0 * b := Int.mul_lt_mul_of_pos_right ha hb + rwa [Int.zero_mul] at h + +protected theorem mul_nonneg_of_nonpos_of_nonpos {a b : Int} + (ha : a ≤ 0) (hb : b ≤ 0) : 0 ≤ a * b := by + have : 0 * b ≤ a * b := Int.mul_le_mul_of_nonpos_right ha hb + rwa [Int.zero_mul] at this + +protected theorem mul_lt_mul_of_neg_left {a b c : Int} (h : b < a) (hc : c < 0) : c * a < c * b := + have : -c > 0 := Int.neg_pos_of_neg hc + have : -c * b < -c * a := Int.mul_lt_mul_of_pos_left h this + have : -(c * b) < -(c * a) := by + rwa [← Int.neg_mul_eq_neg_mul, ← Int.neg_mul_eq_neg_mul] at this + Int.lt_of_neg_lt_neg this + +protected theorem mul_lt_mul_of_neg_right {a b c : Int} (h : b < a) (hc : c < 0) : a * c < b * c := + have : -c > 0 := Int.neg_pos_of_neg hc + have : b * -c < a * -c := Int.mul_lt_mul_of_pos_right h this + have : -(b * c) < -(a * c) := by + rwa [← Int.neg_mul_eq_mul_neg, ← Int.neg_mul_eq_mul_neg] at this + Int.lt_of_neg_lt_neg this + +protected theorem mul_pos_of_neg_of_neg {a b : Int} (ha : a < 0) (hb : b < 0) : 0 < a * b := by + have : 0 * b < a * b := Int.mul_lt_mul_of_neg_right ha hb + rwa [Int.zero_mul] at this + +protected theorem mul_self_le_mul_self {a b : Int} (h1 : 0 ≤ a) (h2 : a ≤ b) : a * a ≤ b * b := + Int.mul_le_mul h2 h2 h1 (Int.le_trans h1 h2) + +protected theorem mul_self_lt_mul_self {a b : Int} (h1 : 0 ≤ a) (h2 : a < b) : a * a < b * b := + Int.mul_lt_mul' (Int.le_of_lt h2) h2 h1 (Int.lt_of_le_of_lt h1 h2) + +/- ## sign -/ + +@[simp] theorem sign_zero : sign 0 = 0 := rfl +@[simp] theorem sign_one : sign 1 = 1 := rfl +theorem sign_neg_one : sign (-1) = -1 := rfl + +@[simp] theorem sign_of_add_one (x : Nat) : Int.sign (x + 1) = 1 := rfl +@[simp] theorem sign_negSucc (x : Nat) : Int.sign (Int.negSucc x) = -1 := rfl + +theorem natAbs_sign (z : Int) : z.sign.natAbs = if z = 0 then 0 else 1 := + match z with | 0 | succ _ | -[_+1] => rfl + +theorem natAbs_sign_of_nonzero {z : Int} (hz : z ≠ 0) : z.sign.natAbs = 1 := by + rw [Int.natAbs_sign, if_neg hz] + +theorem sign_ofNat_of_nonzero {n : Nat} (hn : n ≠ 0) : Int.sign n = 1 := + match n, Nat.exists_eq_succ_of_ne_zero hn with + | _, ⟨n, rfl⟩ => Int.sign_of_add_one n + +@[simp] theorem sign_neg (z : Int) : Int.sign (-z) = -Int.sign z := by + match z with | 0 | succ _ | -[_+1] => rfl + +theorem sign_mul_natAbs : ∀ a : Int, sign a * natAbs a = a + | 0 => rfl + | succ _ => Int.one_mul _ + | -[_+1] => (Int.neg_eq_neg_one_mul _).symm + +@[simp] theorem sign_mul : ∀ a b, sign (a * b) = sign a * sign b + | a, 0 | 0, b => by simp [Int.mul_zero, Int.zero_mul] + | succ _, succ _ | succ _, -[_+1] | -[_+1], succ _ | -[_+1], -[_+1] => rfl + +theorem sign_eq_one_of_pos {a : Int} (h : 0 < a) : sign a = 1 := + match a, eq_succ_of_zero_lt h with + | _, ⟨_, rfl⟩ => rfl + +theorem sign_eq_neg_one_of_neg {a : Int} (h : a < 0) : sign a = -1 := + match a, eq_negSucc_of_lt_zero h with + | _, ⟨_, rfl⟩ => rfl + +theorem eq_zero_of_sign_eq_zero : ∀ {a : Int}, sign a = 0 → a = 0 + | 0, _ => rfl + +theorem pos_of_sign_eq_one : ∀ {a : Int}, sign a = 1 → 0 < a + | (_ + 1 : Nat), _ => ofNat_lt.2 (Nat.succ_pos _) + +theorem neg_of_sign_eq_neg_one : ∀ {a : Int}, sign a = -1 → a < 0 + | (_ + 1 : Nat), h => nomatch h + | 0, h => nomatch h + | -[_+1], _ => negSucc_lt_zero _ + +theorem sign_eq_one_iff_pos (a : Int) : sign a = 1 ↔ 0 < a := + ⟨pos_of_sign_eq_one, sign_eq_one_of_pos⟩ + +theorem sign_eq_neg_one_iff_neg (a : Int) : sign a = -1 ↔ a < 0 := + ⟨neg_of_sign_eq_neg_one, sign_eq_neg_one_of_neg⟩ + +@[simp] theorem sign_eq_zero_iff_zero (a : Int) : sign a = 0 ↔ a = 0 := + ⟨eq_zero_of_sign_eq_zero, fun h => by rw [h, sign_zero]⟩ + +@[simp] theorem sign_sign : sign (sign x) = sign x := by + match x with + | 0 => rfl + | .ofNat (_ + 1) => rfl + | .negSucc _ => rfl + +@[simp] theorem sign_nonneg : 0 ≤ sign x ↔ 0 ≤ x := by + match x with + | 0 => rfl + | .ofNat (_ + 1) => + simp (config := { decide := true }) only [sign, true_iff] + exact Int.le_add_one (ofNat_nonneg _) + | .negSucc _ => simp (config := { decide := true }) [sign] + +/- ## natAbs -/ + +theorem natAbs_ne_zero {a : Int} : a.natAbs ≠ 0 ↔ a ≠ 0 := not_congr Int.natAbs_eq_zero + +theorem natAbs_mul_self : ∀ {a : Int}, ↑(natAbs a * natAbs a) = a * a + | ofNat _ => rfl + | -[_+1] => rfl + +theorem eq_nat_or_neg (a : Int) : ∃ n : Nat, a = n ∨ a = -↑n := ⟨_, natAbs_eq a⟩ + +theorem natAbs_mul_natAbs_eq {a b : Int} {c : Nat} + (h : a * b = (c : Int)) : a.natAbs * b.natAbs = c := by rw [← natAbs_mul, h, natAbs] + +@[simp] theorem natAbs_mul_self' (a : Int) : (natAbs a * natAbs a : Int) = a * a := by + rw [← Int.ofNat_mul, natAbs_mul_self] + +theorem natAbs_eq_iff {a : Int} {n : Nat} : a.natAbs = n ↔ a = n ∨ a = -↑n := by + rw [← Int.natAbs_eq_natAbs_iff, Int.natAbs_ofNat] + +@[deprecated] alias ofNat_natAbs_eq_of_nonneg := natAbs_of_nonneg + +theorem natAbs_add_le (a b : Int) : natAbs (a + b) ≤ natAbs a + natAbs b := by + suffices ∀ a b : Nat, natAbs (subNatNat a b.succ) ≤ (a + b).succ by + match a, b with + | (a:Nat), (b:Nat) => rw [ofNat_add_ofNat, natAbs_ofNat]; apply Nat.le_refl + | (a:Nat), -[b+1] => rw [natAbs_ofNat, natAbs_negSucc]; apply this + | -[a+1], (b:Nat) => + rw [natAbs_negSucc, natAbs_ofNat, Nat.succ_add, Nat.add_comm a b]; apply this + | -[a+1], -[b+1] => rw [natAbs_negSucc, succ_add]; apply Nat.le_refl + refine fun a b => subNatNat_elim a b.succ + (fun m n i => n = b.succ → natAbs i ≤ (m + b).succ) ?_ + (fun i n (e : (n + i).succ = _) => ?_) rfl + · rintro i n rfl + rw [Nat.add_comm _ i, Nat.add_assoc] + exact Nat.le_add_right i (b.succ + b).succ + · apply succ_le_succ + rw [← succ.inj e, ← Nat.add_assoc, Nat.add_comm] + apply Nat.le_add_right + +theorem natAbs_sub_le (a b : Int) : natAbs (a - b) ≤ natAbs a + natAbs b := by + rw [← Int.natAbs_neg b]; apply natAbs_add_le + +theorem negSucc_eq' (m : Nat) : -[m+1] = -m - 1 := by simp only [negSucc_eq, Int.neg_add]; rfl + +theorem natAbs_lt_natAbs_of_nonneg_of_lt {a b : Int} + (w₁ : 0 ≤ a) (w₂ : a < b) : a.natAbs < b.natAbs := + match a, b, eq_ofNat_of_zero_le w₁, eq_ofNat_of_zero_le (Int.le_trans w₁ (Int.le_of_lt w₂)) with + | _, _, ⟨_, rfl⟩, ⟨_, rfl⟩ => ofNat_lt.1 w₂ + +theorem eq_natAbs_iff_mul_eq_zero : natAbs a = n ↔ (a - n) * (a + n) = 0 := by + rw [natAbs_eq_iff, Int.mul_eq_zero, ← Int.sub_neg, Int.sub_eq_zero, Int.sub_eq_zero] + + + +/-! ### toNat -/ + +theorem toNat_eq_max : ∀ a : Int, (toNat a : Int) = max a 0 + | (n : Nat) => (Int.max_eq_left (ofNat_zero_le n)).symm + | -[n+1] => (Int.max_eq_right (Int.le_of_lt (negSucc_lt_zero n))).symm + +@[simp] theorem toNat_zero : (0 : Int).toNat = 0 := rfl + +@[simp] theorem toNat_one : (1 : Int).toNat = 1 := rfl + +@[simp] theorem toNat_of_nonneg {a : Int} (h : 0 ≤ a) : (toNat a : Int) = a := by + rw [toNat_eq_max, Int.max_eq_left h] + +@[simp] theorem toNat_ofNat (n : Nat) : toNat ↑n = n := rfl + +@[simp] theorem toNat_ofNat_add_one {n : Nat} : ((n : Int) + 1).toNat = n + 1 := rfl + +theorem self_le_toNat (a : Int) : a ≤ toNat a := by rw [toNat_eq_max]; apply Int.le_max_left + +@[simp] theorem le_toNat {n : Nat} {z : Int} (h : 0 ≤ z) : n ≤ z.toNat ↔ (n : Int) ≤ z := by + rw [← Int.ofNat_le, Int.toNat_of_nonneg h] + +@[simp] theorem toNat_lt {n : Nat} {z : Int} (h : 0 ≤ z) : z.toNat < n ↔ z < (n : Int) := by + rw [← Int.not_le, ← Nat.not_le, Int.le_toNat h] + +theorem toNat_add {a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : (a + b).toNat = a.toNat + b.toNat := + match a, b, eq_ofNat_of_zero_le ha, eq_ofNat_of_zero_le hb with + | _, _, ⟨_, rfl⟩, ⟨_, rfl⟩ => rfl + +theorem toNat_add_nat {a : Int} (ha : 0 ≤ a) (n : Nat) : (a + n).toNat = a.toNat + n := + match a, eq_ofNat_of_zero_le ha with | _, ⟨_, rfl⟩ => rfl + +@[simp] theorem pred_toNat : ∀ i : Int, (i - 1).toNat = i.toNat - 1 + | 0 => rfl + | (n+1:Nat) => by simp [ofNat_add] + | -[n+1] => rfl + +@[simp] theorem toNat_sub_toNat_neg : ∀ n : Int, ↑n.toNat - ↑(-n).toNat = n + | 0 => rfl + | (_+1:Nat) => Int.sub_zero _ + | -[_+1] => Int.zero_sub _ + +@[simp] theorem toNat_add_toNat_neg_eq_natAbs : ∀ n : Int, n.toNat + (-n).toNat = n.natAbs + | 0 => rfl + | (_+1:Nat) => Nat.add_zero _ + | -[_+1] => Nat.zero_add _ + +theorem mem_toNat' : ∀ (a : Int) (n : Nat), toNat' a = some n ↔ a = n + | (m : Nat), n => Option.some_inj.trans ofNat_inj.symm + | -[m+1], n => by constructor <;> intro. + +@[simp] theorem toNat_neg_nat : ∀ n : Nat, (-(n : Int)).toNat = 0 + | 0 => rfl + | _+1 => rfl diff --git a/Std/Data/List.lean b/Std/Data/List.lean index 137c762db9..2c472a904d 100644 --- a/Std/Data/List.lean +++ b/Std/Data/List.lean @@ -1,6 +1,7 @@ import Std.Data.List.Basic import Std.Data.List.Count import Std.Data.List.Init.Attach +import Std.Data.List.Init.Basic import Std.Data.List.Init.Lemmas import Std.Data.List.Lemmas import Std.Data.List.Pairwise diff --git a/Std/Data/List/Basic.lean b/Std/Data/List/Basic.lean index 71540888f8..42c32c7955 100644 --- a/Std/Data/List/Basic.lean +++ b/Std/Data/List/Basic.lean @@ -1438,27 +1438,6 @@ zipRight = zipWithRight prod.mk -/ @[inline] def zipRight : List α → List β → List (Option α × β) := zipWithRight Prod.mk -/-- -Version of `List.zipWith` that continues to the end of both lists, passing `none` to one argument -once the shorter list has run out. --/ --- TODO We should add a tail-recursive version as we do for other `zip` functions above. -def zipWithAll (f : Option α → Option β → γ) : List α → List β → List γ - | [], bs => bs.map fun b => f none (some b) - | a :: as, [] => (a :: as).map fun a => f (some a) none - | a :: as, b :: bs => f a b :: zipWithAll f as bs - -@[simp] theorem zipWithAll_nil_right : - zipWithAll f as [] = as.map fun a => f (some a) none := by - cases as <;> rfl - -@[simp] theorem zipWithAll_nil_left : - zipWithAll f [] bs = bs.map fun b => f none (some b) := by - rw [zipWithAll] - -@[simp] theorem zipWithAll_cons_cons : - zipWithAll f (a :: as) (b :: bs) = f (some a) (some b) :: zipWithAll f as bs := rfl - /-- If all elements of `xs` are `some xᵢ`, `allSome xs` returns the `xᵢ`. Otherwise it returns `none`. diff --git a/Std/Data/List/Init/Basic.lean b/Std/Data/List/Init/Basic.lean new file mode 100644 index 0000000000..7aa74a0a0c --- /dev/null +++ b/Std/Data/List/Init/Basic.lean @@ -0,0 +1,28 @@ +/- +Copyright (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura +-/ + +namespace List + +/-- +Version of `List.zipWith` that continues to the end of both lists, passing `none` to one argument +once the shorter list has run out. +-/ +-- TODO We should add a tail-recursive version as we do for other `zip` functions. +def zipWithAll (f : Option α → Option β → γ) : List α → List β → List γ + | [], bs => bs.map fun b => f none (some b) + | a :: as, [] => (a :: as).map fun a => f (some a) none + | a :: as, b :: bs => f a b :: zipWithAll f as bs + +@[simp] theorem zipWithAll_nil_right : + zipWithAll f as [] = as.map fun a => f (some a) none := by + cases as <;> rfl + +@[simp] theorem zipWithAll_nil_left : + zipWithAll f [] bs = bs.map fun b => f none (some b) := by + rw [zipWithAll] + +@[simp] theorem zipWithAll_cons_cons : + zipWithAll f (a :: as) (b :: bs) = f (some a) (some b) :: zipWithAll f as bs := rfl diff --git a/Std/Data/List/Init/Lemmas.lean b/Std/Data/List/Init/Lemmas.lean index 6dca88f867..c78b28ee90 100644 --- a/Std/Data/List/Init/Lemmas.lean +++ b/Std/Data/List/Init/Lemmas.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Parikshit Khanna, Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Mario Carneiro -/ import Std.Classes.SetNotation +import Std.Data.Nat.Init.Lemmas +import Std.Data.List.Init.Basic import Std.Logic namespace List @@ -13,7 +15,7 @@ open Nat /-! # Bootstrapping theorems for lists -These are theorems used in the definitions of `Std.Data.List.Basic`. +These are theorems used in the definitions of `Std.Data.List.Basic` and tactics. New theorems should be added to `Std.Data.List.Lemmas` if they are not needed by the bootstrap. -/ @@ -59,6 +61,13 @@ theorem length_eq_zero : length l = 0 ↔ l = [] := ⟨fun h => by cases h <;> simp [Membership.mem, *], fun | Or.inl rfl => by constructor | Or.inr h => by constructor; assumption⟩ +theorem mem_cons_self (a : α) (l : List α) : a ∈ a :: l := .head .. + +theorem mem_cons_of_mem (y : α) {a : α} {l : List α} : a ∈ l → a ∈ y :: l := .tail _ + +theorem eq_nil_iff_forall_not_mem {l : List α} : l = [] ↔ ∀ a, a ∉ l := by + cases l <;> simp + /-! ### append -/ @[simp 1100] theorem singleton_append : [x] ++ l = x :: l := rfl @@ -91,6 +100,9 @@ theorem append_right_inj {t₁ t₂ : List α} (s) : s ++ t₁ = s ++ t₂ ↔ t theorem append_left_inj {s₁ s₂ : List α} (t) : s₁ ++ t = s₂ ++ t ↔ s₁ = s₂ := ⟨fun h => append_inj_left' h rfl, congrArg (· ++ _)⟩ +@[simp] theorem append_eq_nil : p ++ q = [] ↔ p = [] ∧ q = [] := by + cases p <;> simp + /-! ### map -/ @[simp] theorem map_nil {f : α → β} : map f [] = [] := rfl @@ -104,6 +116,12 @@ theorem append_left_inj {s₁ s₂ : List α} (t) : s₁ ++ t = s₂ ++ t ↔ s @[simp] theorem map_id' (l : List α) : map (fun a => a) l = l := by induction l <;> simp_all +@[simp] theorem mem_map {f : α → β} : ∀ {l : List α}, b ∈ l.map f ↔ ∃ a, a ∈ l ∧ f a = b + | [] => by simp + | _ :: l => by simp [mem_map (l := l), eq_comm (a := b)] + +theorem mem_map_of_mem (f : α → β) (h : a ∈ l) : f a ∈ map f l := mem_map.2 ⟨_, h, rfl⟩ + @[simp] theorem map_map (g : β → γ) (f : α → β) (l : List α) : map g (map f l) = map (g ∘ f) l := by induction l <;> simp_all @@ -144,6 +162,11 @@ theorem reverseAux_eq (as bs : List α) : reverseAux as bs = reverse as ++ bs := theorem reverse_map (f : α → β) (l : List α) : (l.map f).reverse = l.reverse.map f := by induction l <;> simp [*] +@[simp] theorem reverse_eq_nil_iff {xs : List α} : xs.reverse = [] ↔ xs = [] := by + match xs with + | [] => simp + | x :: xs => simp + /-! ### nth element -/ theorem get_of_mem : ∀ {a} {l : List α}, a ∈ l → ∃ n, get l n = a @@ -157,6 +180,54 @@ theorem get_mem : ∀ (l : List α) n h, get l ⟨n, h⟩ ∈ l theorem mem_iff_get {a} {l : List α} : a ∈ l ↔ ∃ n, get l n = a := ⟨get_of_mem, fun ⟨_, e⟩ => e ▸ get_mem ..⟩ +theorem get?_len_le : ∀ {l : List α} {n}, length l ≤ n → l.get? n = none + | [], _, _ => rfl + | _ :: l, _+1, h => get?_len_le (l := l) <| Nat.le_of_succ_le_succ h + +theorem get?_eq_get : ∀ {l : List α} {n} (h : n < l.length), l.get? n = some (get l ⟨n, h⟩) + | _ :: _, 0, _ => rfl + | _ :: l, _+1, _ => get?_eq_get (l := l) _ + +theorem get?_eq_some : l.get? n = some a ↔ ∃ h, get l ⟨n, h⟩ = a := + ⟨fun e => + have : n < length l := Nat.gt_of_not_le fun hn => by cases get?_len_le hn ▸ e + ⟨this, by rwa [get?_eq_get this, Option.some.injEq] at e⟩, + fun ⟨h, e⟩ => e ▸ get?_eq_get _⟩ + +@[simp] theorem get?_eq_none : l.get? n = none ↔ length l ≤ n := + ⟨fun e => Nat.ge_of_not_lt (fun h' => by cases e ▸ get?_eq_some.2 ⟨h', rfl⟩), get?_len_le⟩ + +@[simp] theorem get?_map (f : α → β) : ∀ l n, (map f l).get? n = (l.get? n).map f + | [], _ => rfl + | _ :: _, 0 => rfl + | _ :: l, n+1 => get?_map f l n + +@[simp] theorem get?_concat_length : ∀ (l : List α) (a : α), (l ++ [a]).get? l.length = some a + | [], a => rfl + | b :: l, a => by rw [cons_append, length_cons]; simp only [get?, get?_concat_length] + +theorem getLast_eq_get : ∀ (l : List α) (h : l ≠ []), + getLast l h = l.get ⟨l.length - 1, by + match l with + | [] => contradiction + | a :: l => exact Nat.le_refl _⟩ + | [a], h => rfl + | a :: b :: l, h => by + simp [getLast, get, Nat.succ_sub_succ, getLast_eq_get] + +@[simp] theorem getLast?_nil : @getLast? α [] = none := rfl + +theorem getLast?_eq_getLast : ∀ l h, @getLast? α l = some (getLast l h) + | [], h => nomatch h rfl + | _::_, _ => rfl + +theorem getLast?_eq_get? : ∀ (l : List α), getLast? l = l.get? (l.length - 1) + | [] => rfl + | a::l => by rw [getLast?_eq_getLast (a::l) fun., getLast_eq_get, get?_eq_get] + +@[simp] theorem getLast?_concat (l : List α) : getLast? (l ++ [a]) = some a := by + simp [getLast?_eq_get?, Nat.succ_sub_succ] + /-! ### take and drop -/ @[simp] theorem take_append_drop : ∀ (n : Nat) (l : List α), take n l ++ drop n l = l @@ -200,6 +271,14 @@ theorem take_concat_get (l : List α) (i : Nat) (h : i < l.length) : theorem reverse_concat (l : List α) (a : α) : (l.concat a).reverse = a :: l.reverse := by rw [concat_eq_append, reverse_append]; rfl +/-! ### takeWhile and dropWhile -/ + +@[simp] theorem dropWhile_nil : ([] : List α).dropWhile p = [] := rfl + +theorem dropWhile_cons : + (x :: xs : List α).dropWhile p = if p x then xs.dropWhile p else x :: xs := by + split <;> simp_all [dropWhile] + /-! ### foldlM and foldrM -/ @[simp] theorem foldlM_reverse [Monad m] (l : List α) (f : β → α → m β) (b) : @@ -315,6 +394,31 @@ theorem mapM'_eq_mapM [Monad m] [LawfulMonad m] (f : α → m β) (l : List α) theorem find?_cons : (a::as).find? p = match p a with | true => some a | false => as.find? p := rfl +/-! ### filter -/ + +@[simp] theorem filter_nil (p : α → Bool) : filter p [] = [] := rfl + +@[simp] theorem filter_cons_of_pos {p : α → Bool} {a : α} (l) (pa : p a) : + filter p (a :: l) = a :: filter p l := by rw [filter, pa] + +@[simp] theorem filter_cons_of_neg {p : α → Bool} {a : α} (l) (pa : ¬ p a) : + filter p (a :: l) = filter p l := by rw [filter, eq_false_of_ne_true pa] + +theorem filter_cons : + (x :: xs : List α).filter p = if p x then x :: (xs.filter p) else xs.filter p := by + split <;> simp [*] + +theorem mem_filter : x ∈ filter p as ↔ x ∈ as ∧ p x := by + induction as with + | nil => simp [filter] + | cons a as ih => + by_cases h : p a <;> simp [*, or_and_right] + · exact or_congr_left (and_iff_left_of_imp fun | rfl => h).symm + · exact (or_iff_right fun ⟨rfl, h'⟩ => h h').symm + +theorem filter_eq_nil {l} : filter p l = [] ↔ ∀ a ∈ l, ¬p a := by + simp only [eq_nil_iff_forall_not_mem, mem_filter, not_and] + /-! ### findSome? -/ @[simp] theorem findSome?_nil : ([] : List α).findSome? f = none := rfl @@ -361,6 +465,32 @@ theorem lookup_cons [BEq α] {k : α} : zipWith f (a :: as) (b :: bs) = f a b :: zipWith f as bs := by rfl +theorem zipWith_get? {f : α → β → γ} : + (List.zipWith f as bs).get? i = match as.get? i, bs.get? i with + | some a, some b => some (f a b) | _, _ => none := by + induction as generalizing bs i with + | nil => cases bs with + | nil => simp + | cons b bs => simp + | cons a as aih => cases bs with + | nil => simp + | cons b bs => cases i <;> simp_all + +/-! ### zipWithAll -/ + +theorem zipWithAll_get? {f : Option α → Option β → γ} : + (zipWithAll f as bs).get? i = match as.get? i, bs.get? i with + | none, none => .none | a?, b? => some (f a? b?) := by + induction as generalizing bs i with + | nil => induction bs generalizing i with + | nil => simp + | cons b bs bih => cases i <;> simp_all + | cons a as aih => cases bs with + | nil => + specialize @aih [] + cases i <;> simp_all + | cons b bs => cases i <;> simp_all + /-! ### zip -/ @[simp] theorem zip_nil_left : zip ([] : List α) (l : List β) = [] := by @@ -433,3 +563,74 @@ attribute [simp] mapA forA filterAuxM firstM anyM allM findM? findSomeM? -- had attribute `@[simp]`. -- We don't currently provide simp lemmas, -- as this is an internal implementation and they don't seem to be needed. + +/-! ### minimum? -/ + +@[simp] theorem minimum?_nil [Min α] : ([] : List α).minimum? = none := rfl + +-- We don't put `@[simp]` on `minimum?_cons`, +-- because the definition in terms of `foldl` is not useful for proofs. +theorem minimum?_cons [Min α] {xs : List α} : (x :: xs).minimum? = foldl min x xs := rfl + +@[simp] theorem minimum?_eq_none_iff {xs : List α} [Min α] : xs.minimum? = none ↔ xs = [] := by + cases xs <;> simp [minimum?] + +theorem minimum?_mem [Min α] (min_eq_or : ∀ a b : α, min a b = a ∨ min a b = b) : + {xs : List α} → xs.minimum? = some a → a ∈ xs := by + intro xs + match xs with + | nil => simp + | x :: xs => + simp only [minimum?_cons, Option.some.injEq, List.mem_cons] + intro eq + induction xs generalizing x with + | nil => + simp at eq + simp [eq] + | cons y xs ind => + simp at eq + have p := ind _ eq + cases p with + | inl p => + cases min_eq_or x y with | _ q => simp [p, q] + | inr p => simp [p, mem_cons] + +theorem le_minimum?_iff [Min α] [LE α] + (le_min_iff : ∀ a b c : α, a ≤ min b c ↔ a ≤ b ∧ a ≤ c) : + {xs : List α} → xs.minimum? = some a → ∀ x, x ≤ a ↔ ∀ b ∈ xs, x ≤ b + | nil => by simp + | cons x xs => by + rw [minimum?] + intro eq y + simp only [Option.some.injEq] at eq + induction xs generalizing x with + | nil => + simp at eq + simp [eq] + | cons z xs ih => + simp at eq + simp [ih _ eq, le_min_iff, and_assoc] + +-- This could be refactored by designing appropriate typeclasses to replace `le_refl`, `min_eq_or`, +-- and `le_min_iff`. +theorem minimum?_eq_some_iff [Min α] [LE α] [anti : Antisymm ((· : α) ≤ ·)] + (le_refl : ∀ a : α, a ≤ a) + (min_eq_or : ∀ a b : α, min a b = a ∨ min a b = b) + (le_min_iff : ∀ a b c : α, a ≤ min b c ↔ a ≤ b ∧ a ≤ c) {xs : List α} : + xs.minimum? = some a ↔ a ∈ xs ∧ ∀ b ∈ xs, a ≤ b := by + refine ⟨fun h => ⟨minimum?_mem min_eq_or h, (le_minimum?_iff le_min_iff h _).1 (le_refl _)⟩, ?_⟩ + intro ⟨h₁, h₂⟩ + cases xs with + | nil => simp at h₁ + | cons x xs => + exact congrArg some <| anti.1 + ((le_minimum?_iff le_min_iff (xs := x::xs) rfl _).1 (le_refl _) _ h₁) + (h₂ _ (minimum?_mem min_eq_or (xs := x::xs) rfl)) + +-- A specialization of `minimum?_eq_some_iff` to Nat. +theorem minimum?_eq_some_iff' {xs : List Nat} : + xs.minimum? = some a ↔ (a ∈ xs ∧ ∀ b ∈ xs, a ≤ b) := + minimum?_eq_some_iff + (le_refl := Nat.le_refl) + (min_eq_or := fun _ _ => Nat.min_def .. ▸ by split <;> simp) + (le_min_iff := fun _ _ _ => Nat.le_min) diff --git a/Std/Data/List/Lemmas.lean b/Std/Data/List/Lemmas.lean index 3ce9de736c..7a7b2d8314 100644 --- a/Std/Data/List/Lemmas.lean +++ b/Std/Data/List/Lemmas.lean @@ -29,6 +29,9 @@ theorem tail_eq_of_cons_eq (H : h₁ :: t₁ = h₂ :: t₂) : t₁ = t₂ := (c theorem cons_inj (a : α) {l l' : List α} : a :: l = a :: l' ↔ l = l' := ⟨tail_eq_of_cons_eq, congrArg _⟩ +theorem cons_eq_cons {a b : α} {l l' : List α} : a :: l = b :: l' ↔ a = b ∧ l = l' := + List.cons.injEq .. ▸ .rfl + theorem exists_cons_of_ne_nil : ∀ {l : List α}, l ≠ [] → ∃ b L, l = b :: L | c :: l', _ => ⟨c, l', rfl⟩ @@ -45,6 +48,17 @@ theorem exists_mem_of_length_pos : ∀ {l : List α}, 0 < length l → ∃ a, a theorem length_pos_iff_exists_mem {l : List α} : 0 < length l ↔ ∃ a, a ∈ l := ⟨exists_mem_of_length_pos, fun ⟨_, h⟩ => length_pos_of_mem h⟩ +theorem exists_cons_of_length_pos : ∀ {l : List α}, 0 < l.length → ∃ h t, l = h :: t + | _::_, _ => ⟨_, _, rfl⟩ + +theorem length_pos_iff_exists_cons : + ∀ {l : List α}, 0 < l.length ↔ ∃ h t, l = h :: t := + ⟨exists_cons_of_length_pos, fun ⟨_, _, eq⟩ => eq ▸ Nat.succ_pos _⟩ + +theorem exists_cons_of_length_succ : + ∀ {l : List α}, l.length = n + 1 → ∃ h t, l = h :: t + | _::_, _ => ⟨_, _, rfl⟩ + theorem length_pos {l : List α} : 0 < length l ↔ l ≠ [] := Nat.pos_iff_ne_zero.trans (not_congr length_eq_zero) @@ -58,10 +72,6 @@ theorem length_eq_one {l : List α} : length l = 1 ↔ ∃ a, l = [a] := theorem mem_nil_iff (a : α) : a ∈ ([] : List α) ↔ False := by simp -theorem mem_cons_self (a : α) (l : List α) : a ∈ a :: l := .head .. - -theorem mem_cons_of_mem (y : α) {a : α} {l : List α} : a ∈ l → a ∈ y :: l := .tail _ - @[simp] theorem mem_toArray {a : α} {l : List α} : a ∈ l.toArray ↔ a ∈ l := by simp [Array.mem_def] @@ -133,9 +143,6 @@ theorem drop_eq_get_cons : ∀ {n} {l : List α} (h), drop n l = get l ⟨n, h theorem append_eq_append : List.append l₁ l₂ = l₁ ++ l₂ := rfl -@[simp] theorem append_eq_nil : p ++ q = [] ↔ p = [] ∧ q = [] := by - cases p <;> simp - theorem append_ne_nil_of_ne_nil_left (s t : List α) : s ≠ [] → s ++ t ≠ [] := by simp_all theorem append_ne_nil_of_ne_nil_right (s t : List α) : t ≠ [] → s ++ t ≠ [] := by simp_all @@ -178,15 +185,29 @@ theorem mem_append_right {a : α} (l₁ : List α) {l₂ : List α} (h : a ∈ l theorem mem_iff_append {a : α} {l : List α} : a ∈ l ↔ ∃ s t : List α, l = s ++ a :: t := ⟨append_of_mem, fun ⟨s, t, e⟩ => e ▸ by simp⟩ -/-! ### map -/ +/-! ### concat -/ -theorem map_singleton (f : α → β) (a : α) : map f [a] = [f a] := rfl +theorem concat_nil (a : α) : concat [] a = [a] := + rfl -@[simp] theorem mem_map {f : α → β} : ∀ {l : List α}, b ∈ l.map f ↔ ∃ a, a ∈ l ∧ f a = b - | [] => by simp - | _ :: l => by simp [mem_map (l := l), eq_comm (a := b)] +theorem concat_cons (a b : α) (l : List α) : concat (a :: l) b = a :: concat l b := + rfl -theorem mem_map_of_mem (f : α → β) (h : a ∈ l) : f a ∈ map f l := mem_map.2 ⟨_, h, rfl⟩ +theorem init_eq_of_concat_eq {a : α} {l₁ l₂ : List α} : concat l₁ a = concat l₂ a → l₁ = l₂ := by + simp + +theorem last_eq_of_concat_eq {a b : α} {l : List α} : concat l a = concat l b → a = b := by + simp + +theorem concat_ne_nil (a : α) (l : List α) : concat l a ≠ [] := by simp + +theorem concat_append (a : α) (l₁ l₂ : List α) : concat l₁ a ++ l₂ = l₁ ++ a :: l₂ := by simp + +theorem append_concat (a : α) (l₁ l₂ : List α) : l₁ ++ concat l₂ a = concat (l₁ ++ l₂) a := by simp + +/-! ### map -/ + +theorem map_singleton (f : α → β) (a : α) : map f [a] = [f a] := rfl theorem exists_of_mem_map (h : b ∈ map f l) : ∃ a, a ∈ l ∧ f a = b := mem_map.1 h @@ -225,17 +246,6 @@ theorem zipWith_foldl_eq_zip_foldl {f : α → β → γ} (i : δ): (zipWith f l₁ l₂).foldl g i = (zip l₁ l₂).foldl (fun r p => g r (f p.1 p.2)) i := by induction l₁ generalizing i l₂ <;> cases l₂ <;> simp_all -theorem zipWith_get? {f : α → β → γ} : - (List.zipWith f as bs).get? i = match as.get? i, bs.get? i with - | some a, some b => some (f a b) | _, _ => none := by - induction as generalizing bs i with - | nil => cases bs with - | nil => simp - | cons b bs => simp - | cons a as aih => cases bs with - | nil => simp - | cons b bs => cases i <;> simp_all - @[simp] theorem zipWith_eq_nil_iff {f : α → β → γ} {l l'} : zipWith f l l' = [] ↔ l = [] ∨ l' = [] := by cases l <;> cases l' <;> simp @@ -286,21 +296,6 @@ theorem zipWith_append (f : α → β → γ) (l la : List α) (l' lb : List β) simp only [length_cons, Nat.succ.injEq] at h simp [ih _ h] -/-! ### zipWithAll -/ - -theorem zipWithAll_get? {f : Option α → Option β → γ} : - (zipWithAll f as bs).get? i = match as.get? i, bs.get? i with - | none, none => .none | a?, b? => some (f a? b?) := by - induction as generalizing bs i with - | nil => induction bs generalizing i with - | nil => simp - | cons b bs bih => cases i <;> simp_all - | cons a as aih => cases bs with - | nil => - specialize @aih [] - cases i <;> simp_all - | cons b bs => cases i <;> simp_all - /-! ### zip -/ @[simp] theorem length_zip (l₁ : List α) (l₂ : List β) : @@ -455,9 +450,6 @@ fun s => Subset.trans s <| subset_append_right _ _ theorem subset_nil {l : List α} : l ⊆ [] ↔ l = [] := ⟨fun h => match l with | [] => rfl | _::_ => nomatch h (.head ..), fun | rfl => Subset.refl _⟩ -theorem eq_nil_iff_forall_not_mem {l : List α} : l = [] ↔ ∀ a, a ∉ l := - subset_nil.symm.trans <| by simp [subset_def] - 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 _) @@ -692,16 +684,11 @@ theorem getLast_singleton (a h) : @getLast α [a] h = a := rfl theorem getLast!_cons [Inhabited α] : @getLast! α _ (a::l) = getLastD l a := by simp [getLast!, getLast_eq_getLastD] -@[simp] theorem getLast?_nil : @getLast? α [] = none := rfl theorem getLast?_cons : @getLast? α (a::l) = getLastD l a := by simp [getLast?, getLast_eq_getLastD] @[simp] theorem getLast?_singleton (a : α) : getLast? [a] = a := rfl -theorem getLast?_eq_getLast : ∀ l h, @getLast? α l = some (getLast l h) - | [], h => nomatch h rfl - | _::_, _ => rfl - theorem getLast_mem : ∀ {l : List α} (h : l ≠ []), getLast l h ∈ l | [], h => absurd rfl h | [_], _ => .head .. @@ -740,10 +727,6 @@ are often used for theorems about `Array.pop`. -/ @[simp] theorem get_cons_cons_one : (a₁ :: a₂ :: as).get (1 : Fin (as.length + 2)) = a₂ := rfl -theorem get?_eq_get : ∀ {l : List α} {n} (h : n < l.length), l.get? n = some (get l ⟨n, h⟩) - | _ :: _, 0, _ => rfl - | _ :: l, _+1, _ => get?_eq_get (l := l) _ - theorem get!_cons_succ [Inhabited α] (l : List α) (a : α) (n : Nat) : (a::l).get! (n+1) = get! l n := rfl @@ -751,23 +734,10 @@ theorem get!_cons_zero [Inhabited α] (l : List α) (a : α) : (a::l).get! 0 = a theorem get!_nil [Inhabited α] (n : Nat) : [].get! n = (default : α) := rfl -theorem get?_len_le : ∀ {l : List α} {n}, length l ≤ n → l.get? n = none - | [], _, _ => rfl - | _ :: l, _+1, h => get?_len_le (l := l) <| Nat.le_of_succ_le_succ h - theorem get!_len_le [Inhabited α] : ∀ {l : List α} {n}, length l ≤ n → l.get! n = (default : α) | [], _, _ => rfl | _ :: l, _+1, h => get!_len_le (l := l) <| Nat.le_of_succ_le_succ h -theorem get?_eq_some : l.get? n = some a ↔ ∃ h, get l ⟨n, h⟩ = a := - ⟨fun e => - have : n < length l := Nat.lt_of_not_le fun hn => by cases get?_len_le hn ▸ e - ⟨this, by rwa [get?_eq_get this, Option.some.injEq] at e⟩, - fun ⟨h, e⟩ => e ▸ get?_eq_get _⟩ - -@[simp] theorem get?_eq_none : l.get? n = none ↔ length l ≤ n := - ⟨fun e => Nat.le_of_not_lt (fun h' => by cases e ▸ get?_eq_some.2 ⟨h', rfl⟩), get?_len_le⟩ - theorem get?_of_mem {a} {l : List α} (h : a ∈ l) : ∃ n, l.get? n = some a := let ⟨⟨n, _⟩, e⟩ := get_of_mem h; ⟨n, e ▸ get?_eq_get _⟩ @@ -807,11 +777,6 @@ theorem get?_inj rw [mem_iff_get?] exact ⟨_, h₂⟩; exact ⟨_ , h₂.symm⟩ -@[simp] theorem get?_map (f : α → β) : ∀ l n, (map f l).get? n = (l.get? n).map f - | [], _ => rfl - | _ :: _, 0 => rfl - | _ :: l, n+1 => get?_map f l n - @[simp] theorem get_map (f : α → β) {l n} : get (map f l) n = f (get l ⟨n, length_map l f ▸ n.2⟩) := Option.some.inj <| by rw [← get?_eq_get, get?_map, get?_eq_get]; rfl @@ -864,23 +829,6 @@ theorem get?_append {l₁ l₂ : List α} {n : Nat} (hn : n < l₁.length) : length_append .. ▸ Nat.le_add_right .. rw [get?_eq_get hn, get?_eq_get hn', get_append] -theorem getLast_eq_get : ∀ (l : List α) (h : l ≠ []), - getLast l h = l.get ⟨l.length - 1, Nat.sub_lt (length_pos.2 h) Nat.one_pos⟩ - | [a], h => by - rw [getLast_singleton, get_singleton] - | a :: b :: l, h => by rw [getLast_cons', getLast_eq_get (b :: l)]; {rfl}; exact cons_ne_nil b l - -theorem getLast?_eq_get? : ∀ (l : List α), getLast? l = l.get? (l.length - 1) - | [] => rfl - | a::l => by rw [getLast?_eq_getLast (a::l) fun., getLast_eq_get, get?_eq_get] - -@[simp] theorem get?_concat_length : ∀ (l : List α) (a : α), (l ++ [a]).get? l.length = some a - | [], a => rfl - | b :: l, a => by rw [cons_append, length_cons]; simp only [get?, get?_concat_length] - -@[simp] theorem getLast?_concat (l : List α) : getLast? (l ++ [a]) = some a := by - simp [getLast?_eq_get?] - @[simp] theorem getLastD_concat (a b l) : @getLastD α (l ++ [b]) a = b := by rw [getLastD_eq_getLast?, getLast?_concat]; rfl @@ -1184,9 +1132,6 @@ theorem all_eq_not_any_not (l : List α) (p : α → Bool) : l.all p = !l.any (! @[simp] theorem mem_reverse {x : α} {as : List α} : x ∈ reverse as ↔ x ∈ as := by simp [reverse] -@[simp] theorem reverse_eq_nil_iff {xs : List α} : xs.reverse = [] ↔ xs = [] := by - induction xs <;> simp - /-! ### insert -/ section insert @@ -1406,18 +1351,6 @@ end erase /-! ### filter and partition -/ -@[simp] theorem filter_nil (p : α → Bool) : filter p [] = [] := rfl - -@[simp] theorem filter_cons_of_pos {p : α → Bool} {a : α} (l) (pa : p a) : - filter p (a :: l) = a :: filter p l := by rw [filter, pa] - -@[simp] theorem filter_cons_of_neg {p : α → Bool} {a : α} (l) (pa : ¬ p a) : - filter p (a :: l) = filter p l := by rw [filter, eq_false_of_ne_true pa] - -theorem filter_cons : - (x :: xs : List α).filter p = if p x then x :: (xs.filter p) else xs.filter p := by - split <;> simp [*] - @[simp] theorem filter_append {p : α → Bool} : ∀ (l₁ l₂ : List α), filter p (l₁ ++ l₂) = filter p l₁ ++ filter p l₂ | [], l₂ => rfl @@ -1427,14 +1360,6 @@ theorem filter_cons : | [] => .slnil | a :: l => by rw [filter]; split <;> simp [Sublist.cons, Sublist.cons₂, filter_sublist l] -theorem mem_filter : x ∈ filter p as ↔ x ∈ as ∧ p x := by - induction as with - | nil => simp [filter] - | cons a as ih => - by_cases h : p a <;> simp [*, or_and_right] - · exact or_congr_left (and_iff_left_of_imp fun | rfl => h).symm - · exact (or_iff_right fun ⟨rfl, h'⟩ => h h').symm - @[simp] theorem partition_eq_filter_filter (p : α → Bool) (l : List α) : partition p l = (filter p l, filter (not ∘ p) l) := by simp [partition, aux] where aux : ∀ l {as bs}, partition.loop p l (as, bs) = @@ -1543,9 +1468,6 @@ theorem map_filter (f : β → α) (l : List β) : filter p (map f l) = map f (f | [] => rfl | a :: l => by by_cases hp : p a <;> by_cases hq : q a <;> simp [hp, hq, filter_filter _ l] -theorem filter_eq_nil {l} : filter p l = [] ↔ ∀ a ∈ l, ¬p a := by - simp only [eq_nil_iff_forall_not_mem, mem_filter, not_and] - theorem filter_eq_self {l} : filter p l = l ↔ ∀ a ∈ l, p a := by induction l with simp | cons a l ih => @@ -2191,12 +2113,6 @@ theorem disjoint_take_drop : ∀ {l : List α}, l.Nodup → m ≤ n → Disjoint | [] => rfl | x :: xs => by simp [takeWhile, dropWhile]; cases p x <;> simp [takeWhile_append_dropWhile p xs] -@[simp] theorem dropWhile_nil : ([] : List α).dropWhile p = [] := rfl - -theorem dropWhile_cons : - (x :: xs : List α).dropWhile p = if p x then xs.dropWhile p else x :: xs := by - split <;> simp_all [dropWhile] - theorem dropWhile_append {xs ys : List α} : (xs ++ ys).dropWhile p = if (xs.dropWhile p).isEmpty then ys.dropWhile p else xs.dropWhile p ++ ys := by @@ -2404,61 +2320,6 @@ theorem reverse_range' : ∀ s n : Nat, reverse (range' s n) = map (s + n - 1 - @[simp] theorem enum_length : (enum l).length = l.length := enumFrom_length -/-! ### minimum? -/ - -@[simp] theorem minimum?_nil [Min α] : ([] : List α).minimum? = none := rfl - --- We don't put `@[simp]` on `minimum?_cons`, --- because the definition in terms of `foldl` is not useful for proofs. -theorem minimum?_cons [Min α] {xs : List α} : (x :: xs).minimum? = foldl min x xs := rfl - -@[simp] theorem minimum?_eq_none_iff {xs : List α} [Min α] : xs.minimum? = none ↔ xs = [] := by - cases xs <;> simp [minimum?] - -theorem minimum?_mem [Min α] (min_eq_or : ∀ a b : α, min a b = a ∨ min a b = b) : - {xs : List α} → xs.minimum? = some a → a ∈ xs - | nil => by simp - | cons x xs => by - rw [minimum?]; rintro ⟨⟩ - induction xs generalizing x with simp at * - | cons y xs ih => - rcases ih (min x y) with h | h <;> simp [h] - simp [← or_assoc, min_eq_or x y] - -theorem le_minimum?_iff [Min α] [LE α] - (le_min_iff : ∀ a b c : α, a ≤ min b c ↔ a ≤ b ∧ a ≤ c) : - {xs : List α} → xs.minimum? = some a → ∀ x, x ≤ a ↔ ∀ b ∈ xs, x ≤ b - | nil => by simp - | cons x xs => by - rw [minimum?]; rintro ⟨⟩ y - induction xs generalizing x with - | nil => simp - | cons y xs ih => simp [ih, le_min_iff, and_assoc] - --- This could be refactored by designing appropriate typeclasses to replace `le_refl`, `min_eq_or`, --- and `le_min_iff`. -theorem minimum?_eq_some_iff [Min α] [LE α] [anti : Antisymm ((· : α) ≤ ·)] - (le_refl : ∀ a : α, a ≤ a) - (min_eq_or : ∀ a b : α, min a b = a ∨ min a b = b) - (le_min_iff : ∀ a b c : α, a ≤ min b c ↔ a ≤ b ∧ a ≤ c) {xs : List α} : - xs.minimum? = some a ↔ a ∈ xs ∧ ∀ b ∈ xs, a ≤ b := by - refine ⟨fun h => ⟨minimum?_mem min_eq_or h, (le_minimum?_iff le_min_iff h _).1 (le_refl _)⟩, ?_⟩ - intro ⟨h₁, h₂⟩ - cases xs with - | nil => simp at h₁ - | cons x xs => - exact congrArg some <| anti.1 - ((le_minimum?_iff le_min_iff (xs := x::xs) rfl _).1 (le_refl _) _ h₁) - (h₂ _ (minimum?_mem min_eq_or (xs := x::xs) rfl)) - --- A specialization of `minimum?_eq_some_iff` to Nat. -theorem minimum?_eq_some_iff' {xs : List Nat} : - xs.minimum? = some a ↔ (a ∈ xs ∧ ∀ b ∈ xs, a ≤ b) := - minimum?_eq_some_iff - (le_refl := Nat.le_refl) - (min_eq_or := fun _ _ => Nat.min_def .. ▸ by split <;> simp) - (le_min_iff := fun _ _ _ => Nat.le_min) - /-! ### maximum? -/ @[simp] theorem maximum?_nil [Max α] : ([] : List α).maximum? = none := rfl diff --git a/Std/Data/Nat/Basic.lean b/Std/Data/Nat/Basic.lean index 84900593c9..9a861fdf18 100644 --- a/Std/Data/Nat/Basic.lean +++ b/Std/Data/Nat/Basic.lean @@ -128,7 +128,7 @@ where iter n next else guess -termination_by iter guess => guess + termination_by guess /-! ### testBit diff --git a/Std/Data/Nat/Bitwise.lean b/Std/Data/Nat/Bitwise.lean index 9a9fb8263b..9183341c10 100644 --- a/Std/Data/Nat/Bitwise.lean +++ b/Std/Data/Nat/Bitwise.lean @@ -13,19 +13,30 @@ It is primarily intended to support the bitvector library. import Std.Data.Bool import Std.Data.Nat.Lemmas import Std.Tactic.RCases +import Std.Tactic.Omega +import Std.Tactic.Basic namespace Nat @[local simp] private theorem one_div_two : 1/2 = 0 := by trivial -private theorem two_pow_succ_sub_one_div : (2 ^ (n+1) - 1) / 2 = 2^n - 1 := by - apply fun x => (Nat.sub_eq_of_eq_add x).symm - apply Eq.trans _ - apply Nat.add_mul_div_left _ _ Nat.zero_lt_two - rw [Nat.mul_one, ←Nat.sub_add_comm (Nat.pow_two_pos _)] - rw [ Nat.add_sub_assoc Nat.zero_lt_two ] - simp [Nat.pow_succ, Nat.mul_comm _ 2, Nat.mul_add_div] +private theorem two_pow_succ_sub_succ_div_two : (2 ^ (n+1) - (x + 1)) / 2 = 2^n - (x/2 + 1) := by + if h : x + 1 ≤ 2 ^ (n + 1) then + apply fun x => (Nat.sub_eq_of_eq_add x).symm + apply Eq.trans _ + apply Nat.add_mul_div_left _ _ Nat.zero_lt_two + rw [← Nat.sub_add_comm h] + rw [Nat.add_sub_assoc (by omega)] + rw [Nat.pow_succ'] + rw [Nat.mul_add_div Nat.zero_lt_two] + simp [show (2 * (x / 2 + 1) - (x + 1)) / 2 = 0 by omega] + else + rw [Nat.pow_succ'] at * + omega + +private theorem two_pow_succ_sub_one_div_two : (2 ^ (n+1) - 1) / 2 = 2^n - 1 := + two_pow_succ_sub_succ_div_two private theorem two_mul_sub_one {n : Nat} (n_pos : n > 0) : (2*n - 1) % 2 = 1 := by match n with @@ -72,10 +83,10 @@ noncomputable def div2Induction {motive : Nat → Sort u} @[simp] theorem zero_testBit (i : Nat) : testBit 0 i = false := by simp only [testBit, zero_shiftRight, zero_and, bne_self_eq_false] -theorem testBit_zero_is_mod2 (x : Nat) : testBit x 0 = decide (x % 2 = 1) := by +@[simp] theorem testBit_zero (x : Nat) : testBit x 0 = decide (x % 2 = 1) := by cases mod_two_eq_zero_or_one x with | _ p => simp [testBit, p] -theorem testBit_succ_div2 (x i : Nat) : testBit x (succ i) = testBit (x/2) i := by +@[simp] theorem testBit_succ (x i : Nat) : testBit x (succ i) = testBit (x/2) i := by unfold testBit simp [shiftRight_succ_inside] @@ -85,8 +96,7 @@ theorem testBit_to_div_mod {x : Nat} : testBit x i = decide (x / 2^i % 2 = 1) := unfold testBit cases mod_two_eq_zero_or_one x with | _ xz => simp [xz] | succ i hyp => - simp [testBit_succ_div2, hyp, - Nat.div_div_eq_div_mul, Nat.pow_succ, Nat.mul_comm 2] + simp [hyp, Nat.div_div_eq_div_mul, Nat.pow_succ'] theorem ne_zero_implies_bit_true {x : Nat} (xnz : x ≠ 0) : ∃ i, testBit x i := by induction x using div2Induction with @@ -98,10 +108,10 @@ theorem ne_zero_implies_bit_true {x : Nat} (xnz : x ≠ 0) : ∃ i, testBit x i simp only [mod2_eq, ne_eq, Nat.mul_eq_zero, Nat.add_zero, false_or] at xnz have ⟨d, dif⟩ := hyp x_pos xnz apply Exists.intro (d+1) - simp [testBit_succ_div2, dif] + simp_all | Or.inr mod2_eq => apply Exists.intro 0 - simp [testBit_zero_is_mod2, mod2_eq] + simp_all theorem ne_implies_bit_diff {x y : Nat} (p : x ≠ y) : ∃ i, testBit x i ≠ testBit y i := by induction y using Nat.div2Induction generalizing x with @@ -120,11 +130,10 @@ theorem ne_implies_bit_diff {x y : Nat} (p : x ≠ y) : ∃ i, testBit x i ≠ t Nat.zero_lt_succ, Nat.mul_left_cancel_iff] at p have ⟨i, ieq⟩ := hyp ypos p apply Exists.intro (i+1) - simp [testBit_succ_div2] - exact ieq + simpa else apply Exists.intro 0 - simp only [testBit_zero_is_mod2] + simp only [testBit_zero] revert lsb_diff cases mod_two_eq_zero_or_one x with | _ p => cases mod_two_eq_zero_or_one y with | _ q => @@ -145,7 +154,7 @@ theorem eq_of_testBit_eq {x y : Nat} (pred : ∀i, testBit x i = testBit y i) : theorem ge_two_pow_implies_high_bit_true {x : Nat} (p : x ≥ 2^n) : ∃ i, i ≥ n ∧ testBit x i := by induction x using div2Induction generalizing n with | ind x hyp => - have x_pos : x > 0 := Nat.lt_of_lt_of_le (Nat.pow_two_pos n) p + have x_pos : x > 0 := Nat.lt_of_lt_of_le (Nat.two_pow_pos n) p have x_ne_zero : x ≠ 0 := Nat.ne_of_gt x_pos match n with | zero => @@ -153,16 +162,14 @@ theorem ge_two_pow_implies_high_bit_true {x : Nat} (p : x ≥ 2^n) : ∃ i, i exact Exists.intro j (And.intro (Nat.zero_le _) jp) | succ n => have x_ge_n : x / 2 ≥ 2 ^ n := by - simp [le_div_iff_mul_le, ← Nat.pow_succ'] - exact p + simpa [le_div_iff_mul_le, ← Nat.pow_succ'] using p have ⟨j, jp⟩ := @hyp x_pos n x_ge_n apply Exists.intro (j+1) apply And.intro case left => exact (Nat.succ_le_succ jp.left) case right => - simp [testBit_succ_div2 _ j] - exact jp.right + simpa using jp.right theorem testBit_implies_ge {x : Nat} (p : testBit x i = true) : x ≥ 2^i := by simp only [testBit_to_div_mod] at p @@ -170,7 +177,7 @@ theorem testBit_implies_ge {x : Nat} (p : testBit x i = true) : x ≥ 2^i := by have x_lt : x < 2^i := Nat.lt_of_not_le not_ge simp [div_eq_of_lt x_lt] at p -theorem testBit_lt_two {x i : Nat} (lt : x < 2^i) : x.testBit i = false := by +theorem testBit_lt_two_pow {x i : Nat} (lt : x < 2^i) : x.testBit i = false := by match p : x.testBit i with | false => trivial | true => @@ -201,7 +208,7 @@ private theorem testBit_succ_zero : testBit (x + 1) 0 = not (testBit x 0) := by simp [p] theorem testBit_two_pow_add_eq (x i : Nat) : testBit (2^i + x) i = not (testBit x i) := by - simp [testBit_to_div_mod, add_div_left, Nat.pow_two_pos, succ_mod_two] + simp [testBit_to_div_mod, add_div_left, Nat.two_pow_pos, succ_mod_two] cases mod_two_eq_zero_or_one (x / 2 ^ i) with | _ p => simp [p] @@ -222,7 +229,7 @@ theorem testBit_two_pow_add_gt {i j : Nat} (j_lt_i : j < i) (x : Nat) : have i_def : i = j + (i-j) := (Nat.add_sub_cancel' (Nat.le_of_lt j_lt_i)).symm rw [i_def] simp only [testBit_to_div_mod, Nat.pow_add, - Nat.add_comm x, Nat.mul_add_div (Nat.pow_two_pos _)] + Nat.add_comm x, Nat.mul_add_div (Nat.two_pow_pos _)] match i_sub_j_eq : i - j with | 0 => exfalso @@ -231,7 +238,7 @@ theorem testBit_two_pow_add_gt {i j : Nat} (j_lt_i : j < i) (x : Nat) : | d+1 => simp [pow_succ, Nat.mul_comm _ 2, Nat.mul_add_mod] -theorem testBit_mod_two_pow (x j i : Nat) : +@[simp] theorem testBit_mod_two_pow (x j i : Nat) : testBit (x % 2^j) i = (decide (i < j) && testBit x i) := by induction x using Nat.strongInductionOn generalizing j i with | ind x hyp => @@ -244,37 +251,51 @@ theorem testBit_mod_two_pow (x j i : Nat) : · have x_lt : x < 2^i := calc x < 2^j := x_lt_j _ ≤ 2^i := Nat.pow_le_pow_of_le_right Nat.zero_lt_two i_ge_j - simp [Nat.testBit_lt_two x_lt] + simp [Nat.testBit_lt_two_pow x_lt] · generalize y_eq : x - 2^j = y have x_eq : x = y + 2^j := Nat.eq_add_of_sub_eq x_ge_j y_eq - simp only [Nat.pow_two_pos, x_eq, Nat.le_add_left, true_and, ite_true] + simp only [Nat.two_pow_pos, x_eq, Nat.le_add_left, true_and, ite_true] have y_lt_x : y < x := by simp [x_eq] - exact Nat.lt_add_of_pos_right (Nat.pow_two_pos j) + exact Nat.lt_add_of_pos_right (Nat.two_pow_pos j) simp only [hyp y y_lt_x] if i_lt_j : i < j then rw [ Nat.add_comm _ (2^_), testBit_two_pow_add_gt i_lt_j] else simp [i_lt_j] -private theorem testBit_one_zero : testBit 1 0 = true := by trivial +theorem testBit_one_zero : testBit 1 0 = true := by trivial -@[simp] theorem testBit_two_pow_sub_one (n i : Nat) : testBit (2^n-1) i = decide (i < n) := by - induction i generalizing n with +theorem testBit_two_pow_sub_succ (h₂ : x < 2 ^ n) (i : Nat) : + testBit (2^n - (x + 1)) i = (decide (i < n) && ! testBit x i) := by + induction i generalizing n x with | zero => - simp [testBit_zero_is_mod2] + simp only [testBit_zero, zero_eq, Bool.and_eq_true, decide_eq_true_eq, + Bool.not_eq_true'] match n with - | 0 => - simp + | 0 => simp | n+1 => - simp [Nat.pow_succ, Nat.mul_comm _ 2, two_mul_sub_one, Nat.pow_two_pos] - | succ i hyp => - simp [testBit_succ_div2] + -- just logic + omega: + simp only [zero_lt_succ, decide_True, Bool.true_and] + rw [Nat.pow_succ', ← decide_not, decide_eq_decide] + rw [Nat.pow_succ'] at h₂ + omega + | succ i ih => + simp only [testBit_succ] match n with | 0 => - simp [pow_zero] + simp only [pow_zero, succ_sub_succ_eq_sub, Nat.zero_sub, Nat.zero_div, zero_testBit] + rw [decide_eq_false] <;> simp | n+1 => - simp [two_pow_succ_sub_one_div, hyp, Nat.succ_lt_succ_iff] + rw [Nat.two_pow_succ_sub_succ_div_two, ih] + · simp [Nat.succ_lt_succ_iff] + · rw [Nat.pow_succ'] at h₂ + omega + +@[simp] theorem testBit_two_pow_sub_one (n i : Nat) : testBit (2^n-1) i = decide (i < n) := by + rw [testBit_two_pow_sub_succ] + · simp + · exact Nat.two_pow_pos _ theorem testBit_bool_to_nat (b : Bool) (i : Nat) : testBit (Bool.toNat b) i = (decide (i = 0) && b) := by @@ -305,11 +326,11 @@ theorem testBit_bitwise cases i with | zero => cases p : f (decide (x % 2 = 1)) (decide (y % 2 = 1)) <;> - simp [p, testBit_zero_is_mod2, Nat.mul_add_mod, mod_eq_of_lt] + simp [p, Nat.mul_add_mod, mod_eq_of_lt] | succ i => have hyp_i := hyp i (Nat.le_refl (i+1)) cases p : f (decide (x % 2 = 1)) (decide (y % 2 = 1)) <;> - simp [p, testBit_succ_div2, one_div_two, hyp_i, Nat.mul_add_div] + simp [p, one_div_two, hyp_i, Nat.mul_add_div] /-! ### bitwise -/ @@ -328,9 +349,7 @@ private theorem eq_0_of_lt (x : Nat) : x < 2^ 0 ↔ x = 0 := eq_0_of_lt_one x private theorem zero_lt_pow (n : Nat) : 0 < 2^n := by induction n case zero => simp [eq_0_of_lt] - case succ n hyp => - simp [pow_succ] - exact (Nat.mul_lt_mul_of_pos_right hyp Nat.zero_lt_two : 0 < 2 ^ n * 2) + case succ n hyp => simpa [pow_succ] private theorem div_two_le_of_lt_two {m n : Nat} (p : m < 2 ^ succ n) : m / 2 < 2^n := by simp [div_lt_iff_lt_mul Nat.zero_lt_two] @@ -365,14 +384,14 @@ theorem bitwise_lt_two_pow (left : x < 2^n) (right : y < 2^n) : (Nat.bitwise f x /-! ### and -/ -theorem testBit_and (x y i : Nat) : (x &&& y).testBit i = (x.testBit i && y.testBit i) := by +@[simp] theorem testBit_and (x y i : Nat) : (x &&& y).testBit i = (x.testBit i && y.testBit i) := by simp [HAnd.hAnd, AndOp.and, land, testBit_bitwise ] theorem and_lt_two_pow (x : Nat) {y n : Nat} (right : y < 2^n) : (x &&& y) < 2^n := by apply lt_pow_two_of_testBit intro i i_ge_n have yf : testBit y i = false := by - apply Nat.testBit_lt_two + apply Nat.testBit_lt_two_pow apply Nat.lt_of_lt_of_le right exact pow_le_pow_of_le_right Nat.zero_lt_two i_ge_n simp [testBit_and, yf] @@ -399,7 +418,7 @@ theorem and_pow_two_identity {x : Nat} (lt : x < 2^n) : x &&& 2^n-1 = x := by unfold bitwise simp [@eq_comm _ 0] -theorem testBit_or (x y i : Nat) : (x ||| y).testBit i = (x.testBit i || y.testBit i) := by +@[simp] theorem testBit_or (x y i : Nat) : (x ||| y).testBit i = (x.testBit i || y.testBit i) := by simp [HOr.hOr, OrOp.or, lor, testBit_bitwise ] theorem or_lt_two_pow {x y n : Nat} (left : x < 2^n) (right : y < 2^n) : x ||| y < 2^n := @@ -407,7 +426,8 @@ theorem or_lt_two_pow {x y n : Nat} (left : x < 2^n) (right : y < 2^n) : x ||| y /-! ### xor -/ -theorem testBit_xor (x y i : Nat) : (x ^^^ y).testBit i = Bool.xor (x.testBit i) (y.testBit i) := by +@[simp] theorem testBit_xor (x y i : Nat) : + (x ^^^ y).testBit i = Bool.xor (x.testBit i) (y.testBit i) := by simp [HXor.hXor, Xor.xor, xor, testBit_bitwise ] theorem xor_lt_two_pow {x y n : Nat} (left : x < 2^n) (right : y < 2^n) : x ^^^ y < 2^n := @@ -432,7 +452,7 @@ theorem testBit_mul_pow_two_add (a : Nat) {b i : Nat} (b_lt : b < 2^i) (j : Nat) rw [← congrArg (j+·) (Nat.succ_pred i_sub_j_nez)] rw [i_def] simp only [testBit_to_div_mod, Nat.pow_add, Nat.mul_assoc] - simp only [Nat.mul_add_div (Nat.pow_two_pos _), Nat.mul_add_mod] + simp only [Nat.mul_add_div (Nat.two_pow_pos _), Nat.mul_add_mod] simp [Nat.pow_succ, Nat.mul_comm _ 2, Nat.mul_assoc, Nat.mul_add_mod] | inr j_ge => have j_def : j = i + (j-i) := (Nat.add_sub_cancel' j_ge).symm @@ -445,11 +465,11 @@ theorem testBit_mul_pow_two_add (a : Nat) {b i : Nat} (b_lt : b < 2^i) (j : Nat) ←Nat.div_div_eq_div_mul, Nat.mul_add_div, Nat.div_eq_of_lt b_lt, - Nat.pow_two_pos i] + Nat.two_pow_pos i] theorem testBit_mul_pow_two : testBit (2 ^ i * a) j = (decide (j ≥ i) && testBit a (j-i)) := by - have gen := testBit_mul_pow_two_add a (Nat.pow_two_pos i) j + have gen := testBit_mul_pow_two_add a (Nat.two_pow_pos i) j simp at gen rw [gen] cases Nat.lt_or_ge j i with @@ -467,13 +487,13 @@ theorem mul_add_lt_is_or {b : Nat} (b_lt : b < 2^i) (a : Nat) : 2^i * a + b = 2^ have b_lt_j := calc b < 2 ^ i := b_lt _ ≤ 2 ^ j := Nat.pow_le_pow_of_le_right Nat.zero_lt_two i_le - simp [i_le, j_lt, testBit_lt_two, b_lt_j] + simp [i_le, j_lt, testBit_lt_two_pow, b_lt_j] /-! ### shiftLeft and shiftRight -/ -theorem testBit_shiftLeft (x : Nat) : testBit (x <<< i) j = +@[simp] theorem testBit_shiftLeft (x : Nat) : testBit (x <<< i) j = (decide (j ≥ i) && testBit x (j-i)) := by simp [shiftLeft_eq, Nat.mul_comm _ (2^_), testBit_mul_pow_two] -theorem testBit_shiftRight (x : Nat) : testBit (x >>> i) j = testBit x (i+j) := by +@[simp] theorem testBit_shiftRight (x : Nat) : testBit (x >>> i) j = testBit x (i+j) := by simp [testBit, ←shiftRight_add] diff --git a/Std/Data/Nat/Init/Lemmas.lean b/Std/Data/Nat/Init/Lemmas.lean index 20876de423..5c8bbfc379 100644 --- a/Std/Data/Nat/Init/Lemmas.lean +++ b/Std/Data/Nat/Init/Lemmas.lean @@ -52,7 +52,8 @@ protected theorem le_max_left (a b : Nat) : a ≤ max a b := by rw [Nat.max_def] protected theorem le_max_right (a b : Nat) : b ≤ max a b := Nat.max_comm .. ▸ Nat.le_max_left .. -protected theorem pow_two_pos (w : Nat) : 0 < 2^w := Nat.pos_pow_of_pos _ (by decide) +protected theorem two_pow_pos (w : Nat) : 0 < 2^w := Nat.pos_pow_of_pos _ (by decide) +@[deprecated] alias pow_two_pos := Nat.two_pow_pos -- deprecated 2024-02-09 @[simp] protected theorem not_le {a b : Nat} : ¬ a ≤ b ↔ b < a := ⟨Nat.gt_of_not_le, Nat.not_le_of_gt⟩ diff --git a/Std/Data/Nat/Lemmas.lean b/Std/Data/Nat/Lemmas.lean index bc3fdf344b..4a0db49aaa 100644 --- a/Std/Data/Nat/Lemmas.lean +++ b/Std/Data/Nat/Lemmas.lean @@ -4,12 +4,19 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Std.Logic -import Std.Tactic.Basic import Std.Tactic.Alias +import Std.Tactic.RCases import Std.Data.Nat.Init.Lemmas import Std.Data.Nat.Basic import Std.Data.Ord +/-! # Basic lemmas about natural numbers + +The primary purpose of the lemmas in this file is to assist with reasoning +about sizes of objects, array indices and such. For a more thorough development +of the theory of natural numbers, we recommend using Mathlib. +-/ + namespace Nat /-! ### rec/cases -/ @@ -165,7 +172,7 @@ protected theorem le_antisymm_iff {a b : Nat} : a = b ↔ a ≤ b ∧ b ≤ a := protected alias eq_iff_le_and_ge := Nat.le_antisymm_iff protected theorem lt_or_gt_of_ne {a b : Nat} : a ≠ b → a < b ∨ b < a := by - rw [← Nat.not_le, ← Nat.not_le, ← Decidable.not_and, and_comm] + rw [← Nat.not_le, ← Nat.not_le, ← Decidable.not_and_iff_or_not_not, and_comm] exact mt Nat.le_antisymm_iff.2 protected alias lt_or_lt_of_ne := Nat.lt_or_gt_of_ne @[deprecated] protected alias lt_connex := Nat.lt_or_gt_of_ne @@ -192,6 +199,17 @@ protected theorem lt_or_eq_of_le {n m : Nat} (h : n ≤ m) : n < m ∨ n = m := protected theorem le_iff_lt_or_eq {n m : Nat} : n ≤ m ↔ n < m ∨ n = m := ⟨Nat.lt_or_eq_of_le, fun | .inl h => Nat.le_of_lt h | .inr rfl => Nat.le_refl _⟩ +protected theorem lt_succ_iff : m < succ n ↔ m ≤ n := ⟨le_of_lt_succ, lt_succ_of_le⟩ + +protected theorem lt_succ_iff_lt_or_eq : m < succ n ↔ m < n ∨ m = n := + Nat.lt_succ_iff.trans Nat.le_iff_lt_or_eq + +protected theorem eq_of_lt_succ_of_not_lt (hmn : m < n + 1) (h : ¬ m < n) : m = n := + (Nat.lt_succ_iff_lt_or_eq.1 hmn).resolve_left h + +protected theorem eq_of_le_of_lt_succ (h₁ : n ≤ m) (h₂ : m < n + 1) : m = n := + Nat.le_antisymm (le_of_succ_le_succ h₂) h₁ + /-! ## compare -/ theorem compare_def_lt (a b : Nat) : @@ -444,7 +462,7 @@ protected theorem sub_right_comm (m n k : Nat) : m - n - k = m - k - n := by protected theorem add_sub_cancel_right (n m : Nat) : (n + m) - m = n := Nat.add_sub_cancel .. -protected theorem add_sub_cancel' {n m : Nat} (h : m ≤ n) : m + (n - m) = n := by +@[simp] protected theorem add_sub_cancel' {n m : Nat} (h : m ≤ n) : m + (n - m) = n := by rw [Nat.add_comm, Nat.sub_add_cancel h] theorem succ_sub_one (n) : succ n - 1 = n := rfl @@ -564,7 +582,16 @@ theorem sub_lt_succ (a b) : a - b < succ a := lt_succ_of_le (sub_le a b) theorem sub_one_sub_lt (h : i < n) : n - 1 - i < n := by rw [Nat.sub_right_comm]; exact Nat.sub_one_lt_of_le (Nat.sub_pos_of_lt h) (Nat.sub_le ..) -/-! ## min/max -/ +protected theorem exists_eq_add_of_le (h : m ≤ n) : ∃ k : Nat, n = m + k := + ⟨n - m, (add_sub_of_le h).symm⟩ + +protected theorem exists_eq_add_of_le' (h : m ≤ n) : ∃ k : Nat, n = k + m := + ⟨n - m, (Nat.sub_add_cancel h).symm⟩ + +protected theorem exists_eq_add_of_lt (h : m < n) : ∃ k : Nat, n = m + k + 1 := + ⟨n - (m + 1), by rw [Nat.add_right_comm, add_sub_of_le h]⟩ + +/-! ### min/max -/ theorem succ_min_succ (x y) : min (succ x) (succ y) = succ (min x y) := by cases Nat.le_total x y with @@ -738,7 +765,7 @@ protected theorem mul_min_mul_left (a b c : Nat) : min (a * b) (a * c) = a * min repeat rw [Nat.mul_comm a] exact Nat.mul_min_mul_right .. -/-! ## mul -/ +/-! ### mul -/ @[deprecated Nat.mul_le_mul_left] protected theorem mul_le_mul_of_nonneg_left {a b c : Nat} : a ≤ b → c * a ≤ c * b := @@ -839,7 +866,21 @@ protected theorem mul_self_sub_mul_self_eq (a b : Nat) : a * a - b * b = (a + b) rw [Nat.mul_sub_left_distrib, Nat.right_distrib, Nat.right_distrib, Nat.mul_comm b a, Nat.sub_add_eq, Nat.add_sub_cancel] -/-! ## div/mod -/ +protected theorem pos_of_mul_pos_left {a b : Nat} (h : 0 < a * b) : 0 < b := by + by_contra w; simp_all + +protected theorem pos_of_mul_pos_right {a b : Nat} (h : 0 < a * b) : 0 < a := by + by_contra w; simp_all + +@[simp] protected theorem mul_pos_iff_of_pos_left {a b : Nat} (h : 0 < a) : + 0 < a * b ↔ 0 < b := + ⟨Nat.pos_of_mul_pos_left, Nat.mul_pos h⟩ + +@[simp] protected theorem mul_pos_iff_of_pos_right {a b : Nat} (h : 0 < b) : + 0 < a * b ↔ 0 < a := + ⟨Nat.pos_of_mul_pos_right, fun w => Nat.mul_pos w h⟩ + +/-! ### div/mod -/ -- TODO mod_core_congr, mod_def @@ -1045,6 +1086,13 @@ theorem mul_mod_mul_left (z x y : Nat) : (z * x) % (z * y) = z * (x % y) := theorem mul_mod_mul_right (z x y : Nat) : (x * z) % (y * z) = (x % y) * z := by rw [Nat.mul_comm x z, Nat.mul_comm y z, Nat.mul_comm (x % y) z]; apply mul_mod_mul_left +@[simp] theorem mod_mod_of_dvd (a : Nat) (h : c ∣ b) : a % b % c = a % c := by + conv => + rhs + rw [← mod_add_div a b] + obtain ⟨x, rfl⟩ := h + rw [Nat.mul_assoc, add_mul_mod_self_left] + -- TODO cont_to_bool_mod_two theorem sub_mul_mod {x k n : Nat} (h₁ : n*k ≤ x) : (x - n*k) % n = x % n := by @@ -1142,6 +1190,75 @@ protected theorem mul_pow (a b n : Nat) : (a * b) ^ n = a ^ n * b ^ n := by | zero => rw [Nat.pow_zero, Nat.pow_zero, Nat.pow_zero, Nat.mul_one] | succ _ ih => rw [Nat.pow_succ, Nat.pow_succ, Nat.pow_succ, Nat.mul_mul_mul_comm, ih] +protected alias pow_le_pow_left := pow_le_pow_of_le_left +protected alias pow_le_pow_right := pow_le_pow_of_le_right + +protected theorem one_lt_two_pow (h : n ≠ 0) : 1 < 2 ^ n := + match n, h with + | n+1, _ => by + rw [Nat.pow_succ', ← Nat.one_mul 1] + exact Nat.mul_lt_mul_of_lt_of_le' (by decide) (Nat.two_pow_pos n) (by decide) + +@[simp] protected theorem one_lt_two_pow_iff : 1 < 2 ^ n ↔ n ≠ 0 := + ⟨(by intro h p; subst p; simp at h), Nat.one_lt_two_pow⟩ + +protected theorem one_le_two_pow : 1 ≤ 2 ^ n := by + if h : n = 0 then + subst h; simp + else + exact Nat.le_of_lt (Nat.one_lt_two_pow h) + +protected theorem pow_pos (h : 0 < a) : 0 < a^n := + match n with + | 0 => Nat.zero_lt_one + | _ + 1 => Nat.mul_pos (Nat.pow_pos h) h + +protected theorem pow_lt_pow_succ (h : 1 < a) : a ^ n < a ^ (n + 1) := by + rw [Nat.pow_succ] + conv => lhs; rw [← Nat.mul_one (a^n)] + exact Nat.mul_lt_mul_of_le_of_lt (Nat.le_refl _) h (Nat.pow_pos (Nat.lt_trans Nat.zero_lt_one h)) + +protected theorem pow_lt_pow_of_lt {a n m : Nat} (h : 1 < a) (w : n < m) : a ^ n < a ^ m := by + have := Nat.exists_eq_add_of_lt w + cases this + case intro k p => + rw [Nat.add_right_comm] at p + subst p + rw [Nat.pow_add] + conv => lhs; rw [← Nat.mul_one (a^n)] + have t : 0 < a ^ k := Nat.pow_pos (Nat.lt_trans Nat.zero_lt_one h) + exact Nat.mul_lt_mul_of_lt_of_le (Nat.pow_lt_pow_succ h) t t + +protected theorem pow_le_pow_of_le {a n m : Nat} (h : 1 < a) (w : n ≤ m) : a ^ n ≤ a ^ m := by + cases Nat.lt_or_eq_of_le w + case inl lt => + exact Nat.le_of_lt (Nat.pow_lt_pow_of_lt h lt) + case inr eq => + subst eq + exact Nat.le_refl _ + +protected theorem pow_le_pow_iff_right {a n m : Nat} (h : 1 < a) : + a ^ n ≤ a ^ m ↔ n ≤ m := by + constructor + · by_contra w + simp at w + apply Nat.lt_irrefl (a ^ n) + exact Nat.lt_of_le_of_lt w.1 (Nat.pow_lt_pow_of_lt h w.2) + · intro w + cases Nat.eq_or_lt_of_le w + case inl eq => subst eq; apply Nat.le_refl + case inr lt => exact Nat.le_of_lt (Nat.pow_lt_pow_of_lt h lt) + +protected theorem pow_lt_pow_iff_right {a n m : Nat} (h : 1 < a) : + a ^ n < a ^ m ↔ n < m := by + constructor + · by_contra w + simp at w + apply Nat.lt_irrefl (a ^ n) + exact Nat.lt_of_lt_of_le w.1 (Nat.pow_le_pow_of_le h w.2) + · intro w + exact Nat.pow_lt_pow_of_lt h w + /-! ### log2 -/ theorem le_log2 (h : n ≠ 0) : k ≤ n.log2 ↔ 2 ^ k ≤ n := by @@ -1285,6 +1402,60 @@ protected theorem dvd_of_mul_dvd_mul_left protected theorem dvd_of_mul_dvd_mul_right (kpos : 0 < k) (H : m * k ∣ n * k) : m ∣ n := by rw [Nat.mul_comm m k, Nat.mul_comm n k] at H; exact Nat.dvd_of_mul_dvd_mul_left kpos H +theorem pow_dvd_pow_iff_pow_le_pow {k l : Nat} : + ∀ {x : Nat}, 0 < x → (x ^ k ∣ x ^ l ↔ x ^ k ≤ x ^ l) + | x + 1, w => by + constructor + · intro a + exact le_of_dvd (Nat.pow_pos (succ_pos x)) a + · intro a + cases x + case zero => simp + case succ x => + have le := + (Nat.pow_le_pow_iff_right (Nat.succ_le_succ (Nat.succ_le_succ (Nat.zero_le _)))).mp a + refine ⟨(x + 2) ^ (l - k), ?_⟩ + rw [← Nat.pow_add, Nat.add_comm k, Nat.sub_add_cancel le] + +/-- If `1 < x`, then `x^k` divides `x^l` if and only if `k` is at most `l`. -/ +theorem pow_dvd_pow_iff_le_right {x k l : Nat} (w : 1 < x) : x ^ k ∣ x ^ l ↔ k ≤ l := by + rw [pow_dvd_pow_iff_pow_le_pow (lt_of_succ_lt w), Nat.pow_le_pow_iff_right w] + +theorem pow_dvd_pow_iff_le_right' {b k l : Nat} : (b + 2) ^ k ∣ (b + 2) ^ l ↔ k ≤ l := + pow_dvd_pow_iff_le_right (Nat.lt_of_sub_eq_succ rfl) + +protected theorem eq_mul_of_div_eq_right {a b c : Nat} (H1 : b ∣ a) (H2 : a / b = c) : + a = b * c := by + rw [← H2, Nat.mul_div_cancel' H1] + +protected theorem div_eq_iff_eq_mul_right {a b c : Nat} (H : 0 < b) (H' : b ∣ a) : + a / b = c ↔ a = b * c := + ⟨Nat.eq_mul_of_div_eq_right H', Nat.div_eq_of_eq_mul_right H⟩ + +protected theorem div_eq_iff_eq_mul_left {a b c : Nat} (H : 0 < b) (H' : b ∣ a) : + a / b = c ↔ a = c * b := by + rw [Nat.mul_comm]; exact Nat.div_eq_iff_eq_mul_right H H' + +protected theorem pow_dvd_pow {m n : Nat} (a : Nat) (h : m ≤ n) : a ^ m ∣ a ^ n := by + cases Nat.exists_eq_add_of_le h + case intro k p => + subst p + rw [Nat.pow_add] + apply Nat.dvd_mul_right + +protected theorem pow_sub_mul_pow (a : Nat) {m n : Nat} (h : m ≤ n) : + a ^ (n - m) * a ^ m = a ^ n := by + rw [← Nat.pow_add, Nat.sub_add_cancel h] + +theorem pow_dvd_of_le_of_pow_dvd {p m n k : Nat} (hmn : m ≤ n) (hdiv : p ^ n ∣ k) : p ^ m ∣ k := + Nat.dvd_trans (Nat.pow_dvd_pow _ hmn) hdiv + +theorem dvd_of_pow_dvd {p k m : Nat} (hk : 1 ≤ k) (hpk : p ^ k ∣ m) : p ∣ m := by + rw [← Nat.pow_one p]; exact pow_dvd_of_le_of_pow_dvd hk hpk + +protected theorem pow_div {x m n : Nat} (h : n ≤ m) (hx : 0 < x) : x ^ m / x ^ n = x ^ (m - n) := by + rw [Nat.div_eq_iff_eq_mul_left (Nat.pow_pos hx) (Nat.pow_dvd_pow _ h), Nat.pow_sub_mul_pow _ h] + /-! ### sum -/ @[simp] theorem sum_nil : Nat.sum [] = 0 := rfl @@ -1361,3 +1532,34 @@ theorem mul_add_mod (m x y : Nat) : (m * x + y) % m = y % m := by cases n · exact (m % 0).div_zero · case succ n => exact Nat.div_eq_of_lt (m.mod_lt n.succ_pos) + +/-! ### Decidability of predicates -/ + +instance decidableBallLT : + ∀ (n : Nat) (P : ∀ k, k < n → Prop) [∀ n h, Decidable (P n h)], Decidable (∀ n h, P n h) +| 0, _, _ => isTrue fun _ => (by cases ·) +| n + 1, P, H => + match decidableBallLT n (P · <| lt_succ_of_lt ·) with + | isFalse h => isFalse (h fun _ _ => · _ _) + | isTrue h => + match H n Nat.le.refl with + | isFalse p => isFalse (p <| · _ _) + | isTrue p => isTrue fun _ h' => (Nat.lt_succ_iff_lt_or_eq.1 h').elim (h _) fun hn => hn ▸ p + +instance decidableForallFin (P : Fin n → Prop) [DecidablePred P] : Decidable (∀ i, P i) := + decidable_of_iff (∀ k h, P ⟨k, h⟩) ⟨fun m ⟨k, h⟩ => m k h, fun m k h => m ⟨k, h⟩⟩ + +instance decidableBallLE (n : Nat) (P : ∀ k, k ≤ n → Prop) [∀ n h, Decidable (P n h)] : + Decidable (∀ n h, P n h) := + decidable_of_iff (∀ (k) (h : k < succ n), P k (le_of_lt_succ h)) + ⟨fun m k h => m k (lt_succ_of_le h), fun m k _ => m k _⟩ + +instance decidableExistsLT [h : DecidablePred p] : DecidablePred fun n => ∃ m : Nat, m < n ∧ p m + | 0 => isFalse (by simp only [not_lt_zero, false_and, exists_const, not_false_eq_true]) + | n + 1 => + @decidable_of_decidable_of_iff _ _ (@instDecidableOr _ _ (decidableExistsLT (p := p) n) (h n)) + (by simp only [Nat.lt_succ_iff_lt_or_eq, or_and_right, exists_or, exists_eq_left]) + +instance decidableExistsLE [DecidablePred p] : DecidablePred fun n => ∃ m : Nat, m ≤ n ∧ p m := + fun n => decidable_of_iff (∃ m, m < n + 1 ∧ p m) + (exists_congr fun _ => and_congr_left' Nat.lt_succ_iff) diff --git a/Std/Data/PairingHeap.lean b/Std/Data/PairingHeap.lean index a4d192b44f..fbcdcdcb7f 100644 --- a/Std/Data/PairingHeap.lean +++ b/Std/Data/PairingHeap.lean @@ -173,7 +173,7 @@ by repeatedly pulling the minimum element out of the heap. | some (hd, tl) => have : tl.size < s.size := by simp_arith [Heap.size_deleteMin_lt eq] do foldM le tl (← f init hd) f -termination_by _ => s.size +termination_by s.size /-- `O(n log n)`. Fold over the elements of a heap in increasing order, diff --git a/Std/Data/RBMap/Basic.lean b/Std/Data/RBMap/Basic.lean index 2080cde801..bf69dc8b5a 100644 --- a/Std/Data/RBMap/Basic.lean +++ b/Std/Data/RBMap/Basic.lean @@ -360,7 +360,7 @@ def append : RBNode α → RBNode α → RBNode α | bc => balLeft a x (node black bc y d) | a@(node black ..), node red b x c => node red (append a b) x c | node red a x b, c@(node black ..) => node red a x (append b c) -termination_by _ x y => x.size + y.size +termination_by x y => x.size + y.size /-! ## erase -/ diff --git a/Std/Data/RBMap/WF.lean b/Std/Data/RBMap/WF.lean index 1d6a76c350..c786122429 100644 --- a/Std/Data/RBMap/WF.lean +++ b/Std/Data/RBMap/WF.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro -/ import Std.Logic import Std.Data.RBMap.Basic +import Std.Tactic.SeqFocus /-! # Lemmas for Red-black trees @@ -308,7 +309,7 @@ protected theorem All.append (hl : l.All p) (hr : r.All p) : (append l r).All p have := hb.append hc; split <;> simp_all [All.balLeft] · simp_all [hl.append hr.2.1] · simp_all [hl.2.2.append hr] -termination_by _ => l.size + r.size +termination_by l.size + r.size /-- The `append` function preserves the ordering invariants. -/ protected theorem Ordered.append {l : RBNode α} {v : α} {r : RBNode α} @@ -341,7 +342,7 @@ protected theorem Ordered.append {l : RBNode α} {v : α} {r : RBNode α} exact ⟨(vx.trans_r lv).append bx, yc, hl.append lv vb hb, hc⟩ · have ⟨xv, _, bv⟩ := lv; have ⟨ax, xb, ha, hb⟩ := hl exact ⟨ax, xb.append (xv.trans_l vr), ha, hb.append bv vr hr⟩ -termination_by _ => l.size + r.size +termination_by l.size + r.size /-- The balance properties of the `append` function. -/ protected theorem Balanced.append {l r : RBNode α} @@ -380,7 +381,7 @@ protected theorem Balanced.append {l r : RBNode α} · have .red ha hb := hl; have IH := hb.append hr have .black hc hd := hr; have ⟨c, IH⟩ := IH.of_false (· rfl rfl) exact .redred (fun.) ha IH -termination_by _ => l.size + r.size +termination_by l.size + r.size /-! ## erase -/ diff --git a/Std/Data/Range/Lemmas.lean b/Std/Data/Range/Lemmas.lean index 1ff0fc8b6f..958d074a6d 100644 --- a/Std/Data/Range/Lemmas.lean +++ b/Std/Data/Range/Lemmas.lean @@ -1,4 +1,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 Std.Tactic.ByCases +import Std.Tactic.SeqFocus import Std.Data.List.Lemmas import Std.Data.List.Init.Attach diff --git a/Std/Data/Rat/Lemmas.lean b/Std/Data/Rat/Lemmas.lean index 3bf7c41b0d..0401dd6084 100644 --- a/Std/Data/Rat/Lemmas.lean +++ b/Std/Data/Rat/Lemmas.lean @@ -1,5 +1,11 @@ -import Std.Data.Int.Lemmas +/- +Copyright (c) 2022 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ +import Std.Data.Int.Init.Lemmas import Std.Data.Rat.Basic +import Std.Tactic.SeqFocus /-! # Additional lemmas about the Rational Numbers -/ diff --git a/Std/Data/String/Basic.lean b/Std/Data/String/Basic.lean index 014d2fc9b6..072fbe0222 100644 --- a/Std/Data/String/Basic.lean +++ b/Std/Data/String/Basic.lean @@ -91,7 +91,7 @@ where spos else spos -termination_by loop => s.stopPos.byteIdx - spos.byteIdx + termination_by s.stopPos.byteIdx - spos.byteIdx /-- Returns the longest common suffix of two substrings. @@ -112,7 +112,7 @@ where spos else spos -termination_by loop => spos.byteIdx + termination_by spos.byteIdx /-- If `pre` is a prefix of `s`, i.e. `s = pre ++ t`, returns the remainder `t`. diff --git a/Std/Data/String/Lemmas.lean b/Std/Data/String/Lemmas.lean index bc2864d07a..958daba086 100644 --- a/Std/Data/String/Lemmas.lean +++ b/Std/Data/String/Lemmas.lean @@ -7,6 +7,7 @@ import Std.Data.Char import Std.Data.Nat.Lemmas import Std.Data.List.Lemmas import Std.Data.String.Basic +import Std.Tactic.SeqFocus import Std.Tactic.Ext.Attr import Std.Tactic.Simpa diff --git a/Std/Data/UInt.lean b/Std/Data/UInt.lean index 6b8d79bf09..fe87ee037a 100644 --- a/Std/Data/UInt.lean +++ b/Std/Data/UInt.lean @@ -1,3 +1,8 @@ +/- +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 Std.Data.Nat.Init.Lemmas import Std.Tactic.Ext.Attr @@ -75,7 +80,7 @@ theorem UInt64.toNat_lt (x : UInt64) : x.toNat < 2 ^ 64 := x.val.isLt @[simp] theorem USize.val_val_eq_toNat (x : USize) : x.val.val = x.toNat := rfl theorem USize.size_eq : USize.size = 2 ^ System.Platform.numBits := by - have : 1 ≤ 2 ^ System.Platform.numBits := Nat.succ_le_of_lt (Nat.pow_two_pos _) + have : 1 ≤ 2 ^ System.Platform.numBits := Nat.succ_le_of_lt (Nat.two_pow_pos _) rw [USize.size, Nat.succ_eq_add_one, Nat.sub_eq, Nat.sub_add_cancel this] theorem USize.le_size : 2 ^ 32 ≤ USize.size := by diff --git a/Std/Lean/Delaborator.lean b/Std/Lean/Delaborator.lean index 4c70f3ac31..98aa890e01 100644 --- a/Std/Lean/Delaborator.lean +++ b/Std/Lean/Delaborator.lean @@ -7,36 +7,6 @@ import Lean.PrettyPrinter open Lean PrettyPrinter Delaborator SubExpr -/-- -This is similar to `withAppFnArgs` but it handles construction of an "over-application". -For example, suppose we want to implement a delaborator for applications of `f : Foo A → A` -like `f x` as `F[x]`, but because `A` is a type variable we might encounter a term of the form -`@f (A → B) x y` which has an additional argument `y`. - -Most of the built in delaborators will deliberately fail on such an example, resulting in -delaborated syntax `f x y`, but this combinator can be used if we want to display `F[x] y` -instead. - -* `arity`: the expected number of arguments to `f`. - The combinator will fail if fewer than this number of arguments are passed -* `x`: constructs data corresponding to the main application (`f x` in the example) --/ -def Lean.PrettyPrinter.Delaborator.withOverApp (arity : Nat) (x : Delab) : Delab := do - let n := (← getExpr).getAppNumArgs - guard (n ≥ arity) - let kinds ← getParamKinds - let rec - /-- Inner loop of `withOverApp`. -/ - loop : Nat → DelabM (Term × Array Term) - | 0 => return (← x, #[]) - | n+1 => do - let mut (fnStx, args) ← withAppFn (loop n) - if kinds.get? (n + arity) |>.all (·.bInfo.isExplicit) then - args := args.push (← withAppArg delab) - pure (fnStx, args) - let (fnStx, argStxs) ← loop (n - arity) - return Syntax.mkApp fnStx argStxs - /-- Pretty print a const expression using `delabConst` and generate terminfo. This function avoids inserting `@` if the constant is for a function whose first argument is implicit, which is what the default `toMessageData` for `Expr` does. diff --git a/Std/Lean/InfoTree.lean b/Std/Lean/InfoTree.lean index f0977a19e9..d69450d575 100644 --- a/Std/Lean/InfoTree.lean +++ b/Std/Lean/InfoTree.lean @@ -14,7 +14,7 @@ partial def InfoTree.foldInfo' (init : α) (f : ContextInfo → InfoTree → α where /-- `foldInfo'.go` is like `foldInfo'` but with an additional outer context parameter `ctx?`. -/ go ctx? a - | context ctx t => go ctx a t + | context ctx t => go (ctx.mergeIntoOuter? ctx?) a t | t@(node i ts) => let a := match ctx? with | none => a diff --git a/Std/Lean/Meta/Basic.lean b/Std/Lean/Meta/Basic.lean index b75dacbf57..a0e1b94b68 100644 --- a/Std/Lean/Meta/Basic.lean +++ b/Std/Lean/Meta/Basic.lean @@ -60,14 +60,6 @@ Check whether a metavariable is declared in the given `MetavarContext`. def isExprMVarDeclared (mctx : MetavarContext) (mvarId : MVarId) : Bool := mctx.decls.contains mvarId -/-- -Add a delayed assignment for the given metavariable. You must make sure that -the metavariable is not already assigned or delayed-assigned. --/ -def delayedAssignExprMVar (mctx : MetavarContext) (mvarId : MVarId) - (ass : DelayedMetavarAssignment) : MetavarContext := - { mctx with dAssignment := mctx.dAssignment.insert mvarId ass } - /-- Erase any assignment or delayed assignment of the given metavariable. -/ @@ -156,14 +148,6 @@ Check whether a metavariable is declared. def isDeclared [Monad m] [MonadMCtx m] (mvarId : MVarId) : m Bool := return (← getMCtx).isExprMVarDeclared mvarId -/-- -Add a delayed assignment for the given metavariable. You must make sure that -the metavariable is not already assigned or delayed-assigned. --/ -def delayedAssign [MonadMCtx m] (mvarId : MVarId) - (ass : DelayedMetavarAssignment) : m Unit := - modifyMCtx (·.delayedAssignExprMVar mvarId ass) - /-- Erase any assignment or delayed assignment of the given metavariable. -/ @@ -407,7 +391,7 @@ where match ← observing? (f g) with | some gs' => go n true gs' (gs::stk) acc | none => go n p gs stk (acc.push g) -termination_by _ n p gs stk _ => (n, stk, gs) +termination_by n p gs stk _ => (n, stk, gs) /-- `repeat' f` runs `f` on all of the goals to produce a new list of goals, diff --git a/Std/Lean/Meta/DiscrTree.lean b/Std/Lean/Meta/DiscrTree.lean index 9250e07855..d6940b31bb 100644 --- a/Std/Lean/Meta/DiscrTree.lean +++ b/Std/Lean/Meta/DiscrTree.lean @@ -165,11 +165,11 @@ Inserts a new key into a discrimination tree, but only if it is not of the form `#[*]` or `#[=, *, *, *]`. -/ def insertIfSpecific [BEq α] (d : DiscrTree α) - (keys : Array DiscrTree.Key) (v : α) (config : WhnfCoreConfig) : DiscrTree α := + (keys : Array DiscrTree.Key) (v : α) : DiscrTree α := if keys == #[Key.star] || keys == #[Key.const `Eq 3, Key.star, Key.star, Key.star] then d else - d.insertCore keys v config + d.insertCore keys v variable {m : Type → Type} [Monad m] diff --git a/Std/Lean/Meta/LazyDiscrTree.lean b/Std/Lean/Meta/LazyDiscrTree.lean index ab45798e64..f8b0adadd8 100644 --- a/Std/Lean/Meta/LazyDiscrTree.lean +++ b/Std/Lean/Meta/LazyDiscrTree.lean @@ -5,6 +5,7 @@ Authors: Joe Hendrix, Scott Morrison -/ import Lean.Meta.DiscrTree +import Std.Lean.Name import Std.Data.Nat.Init.Lemmas /-! @@ -558,8 +559,8 @@ private def MatchResult.push (r : MatchResult α) (score : Nat) (e : Array α) : loop (a.push #[]) else { elts := a.push #[e] } + termination_by score - a.size loop r.elts - termination_by loop _ => score - a.size private partial def MatchResult.toArray (mr : MatchResult α) : Array α := loop (Array.mkEmpty n) mr.elts @@ -845,7 +846,7 @@ private def createImportedEnvironmentSeq (env : Environment) go d tree (start+1) stop else toFlat d tree - termination_by go _ idx stop => stop - idx + termination_by stop - start /-- Get the results of each task and merge using combining function -/ private def combineGet [Append α] (z : α) (tasks : Array (Task α)) : α := @@ -873,9 +874,9 @@ def createImportedEnvironment (env : Environment) tasks.push <$> (createImportedEnvironmentSeq env act start n).asTask else pure tasks + termination_by env.header.moduleData.size - idx let tasks ← go #[] 0 0 0 let r := combineGet default tasks if p : r.errors.size > 0 then throw r.errors[0].exception pure <| r.tree.toLazy - termination_by go _ idx => env.header.moduleData.size - idx diff --git a/Std/Lean/Meta/Simp.lean b/Std/Lean/Meta/Simp.lean index 7c9793b5f4..a00a5727d0 100644 --- a/Std/Lean/Meta/Simp.lean +++ b/Std/Lean/Meta/Simp.lean @@ -53,15 +53,18 @@ def mkSimpContext' (simpTheorems : SimpTheorems) (stx : Syntax) (eraseLocal : Bo simpOnlyBuiltins.foldlM (·.addConst ·) {} else pure simpTheorems + let simprocs ← if simpOnly then pure {} else Simp.getSimprocs let congrTheorems ← Meta.getSimpCongrTheorems - let r ← elabSimpArgs stx[4] (eraseLocal := eraseLocal) (kind := kind) { + let r ← elabSimpArgs stx[4] (eraseLocal := eraseLocal) (kind := kind) (simprocs := #[simprocs]) { config := (← elabSimpConfig stx[1] (kind := kind)) simpTheorems := #[simpTheorems], congrTheorems } if !r.starArg || ignoreStarArg then return { r with dischargeWrapper } else - let mut simpTheorems := r.ctx.simpTheorems + let ctx := r.ctx + let mut simpTheorems := ctx.simpTheorems + let simprocs := r.simprocs /- When using `zeta := false`, we do not expand let-declarations when using `[*]`. Users must explicitly include it in the list. @@ -70,7 +73,9 @@ def mkSimpContext' (simpTheorems : SimpTheorems) (stx : Syntax) (eraseLocal : Bo for h in hs do unless simpTheorems.isErased (.fvar h) do simpTheorems ← simpTheorems.addTheorem (.fvar h) (← h.getDecl).toExpr - return { ctx := { r.ctx with simpTheorems }, dischargeWrapper } + let ctx := { ctx with simpTheorems } + return { ctx, simprocs, dischargeWrapper } + end Simp diff --git a/Std/Lean/Syntax.lean b/Std/Lean/Syntax.lean new file mode 100644 index 0000000000..5fa6e341b0 --- /dev/null +++ b/Std/Lean/Syntax.lean @@ -0,0 +1,30 @@ +/- +Copyright (c) 2022 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Gabriel Ebner +-/ +import Lean.Syntax + +/-! +# Helper functions for working with typed syntaxes. +-/ + +namespace Lean + +/-- +Applies the given function to every subsyntax. + +Like `Syntax.replaceM` but for typed syntax. +(Note there are no guarantees of type correctness here.) +-/ +def TSyntax.replaceM [Monad M] (f : Syntax → M (Option Syntax)) (stx : TSyntax k) : M (TSyntax k) := + .mk <$> stx.1.replaceM f + +/-- +Constructs a typed separated array from elements. +The given array does not include the separators. + +Like `Syntax.SepArray.ofElems` but for typed syntax. +-/ +def Syntax.TSepArray.ofElems {sep} (elems : Array (TSyntax k)) : TSepArray k sep := + .mk (SepArray.ofElems (sep := sep) elems).1 diff --git a/Std/Logic.lean b/Std/Logic.lean index 2afce52521..4759fdec94 100644 --- a/Std/Logic.lean +++ b/Std/Logic.lean @@ -3,9 +3,11 @@ 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 Std.Tactic.Basic +import Std.Tactic.Init +import Std.Tactic.NoMatch import Std.Tactic.Alias import Std.Tactic.Lint.Misc +import Std.Tactic.ByCases instance {f : α → β} [DecidablePred p] : DecidablePred (p ∘ f) := inferInstanceAs <| DecidablePred fun x => p (f x) @@ -592,7 +594,7 @@ theorem Decidable.imp_or' [Decidable b] : (a → b ∨ c) ↔ (a → b) ∨ (a if h : b then by simp [h] else by rw [eq_false h, false_or]; exact (or_iff_right_of_imp fun hx x => (hx x).elim).symm -theorem Decidable.not_imp [Decidable a] : ¬(a → b) ↔ a ∧ ¬b := +theorem Decidable.not_imp_iff_and_not [Decidable a] : ¬(a → b) ↔ a ∧ ¬b := ⟨fun h => ⟨of_not_imp h, not_of_not_imp h⟩, not_imp_of_and_not⟩ theorem Decidable.peirce (a b : Prop) [Decidable a] : ((a → b) → a) → a := @@ -624,17 +626,17 @@ theorem Decidable.iff_iff_not_or_and_or_not [Decidable a] [Decidable b] : theorem Decidable.not_and_not_right [Decidable b] : ¬(a ∧ ¬b) ↔ (a → b) := ⟨fun h ha => not_imp_symm (And.intro ha) h, fun h ⟨ha, hb⟩ => hb <| h ha⟩ -theorem Decidable.not_and [Decidable a] : ¬(a ∧ b) ↔ ¬a ∨ ¬b := +theorem Decidable.not_and_iff_or_not_not [Decidable a] : ¬(a ∧ b) ↔ ¬a ∨ ¬b := ⟨fun h => if ha : a then .inr (h ⟨ha, ·⟩) else .inl ha, not_and_of_not_or_not⟩ -theorem Decidable.not_and' [Decidable b] : ¬(a ∧ b) ↔ ¬a ∨ ¬b := +theorem Decidable.not_and_iff_or_not_not' [Decidable b] : ¬(a ∧ b) ↔ ¬a ∨ ¬b := ⟨fun h => if hb : b then .inl (h ⟨·, hb⟩) else .inr hb, not_and_of_not_or_not⟩ theorem Decidable.or_iff_not_and_not [Decidable a] [Decidable b] : a ∨ b ↔ ¬(¬a ∧ ¬b) := by rw [← not_or, not_not] theorem Decidable.and_iff_not_or_not [Decidable a] [Decidable b] : a ∧ b ↔ ¬(¬a ∨ ¬b) := by - rw [← not_and, not_not] + rw [← not_and_iff_or_not_not, not_not] theorem Decidable.imp_iff_right_iff [Decidable a] : (a → b ↔ b) ↔ a ∨ b := ⟨fun H => (Decidable.em a).imp_right fun ha' => H.1 fun ha => (ha' ha).elim, @@ -665,13 +667,11 @@ This is the same as `decidable_of_iff` but the iff is flipped. -/ instance Decidable.predToBool (p : α → Prop) [DecidablePred p] : CoeDep (α → Prop) p (α → Bool) := ⟨fun b => decide <| p b⟩ -theorem Bool.false_ne_true : false ≠ true := fun. - /-- Prove that `a` is decidable by constructing a boolean `b` and a proof that `b ↔ a`. (This is sometimes taken as an alternate definition of decidability.) -/ def decidable_of_bool : ∀ (b : Bool), (b ↔ a) → Decidable a | true, h => isTrue (h.1 rfl) - | false, h => isFalse (mt h.2 Bool.false_ne_true) + | false, h => isFalse (mt h.2 Bool.noConfusion) protected theorem Decidable.not_forall {p : α → Prop} [Decidable (∃ x, ¬p x)] [∀ x, Decidable (p x)] : (¬∀ x, p x) ↔ ∃ x, ¬p x := @@ -718,6 +718,15 @@ theorem or_iff_not_imp_left : a ∨ b ↔ (¬a → b) := theorem or_iff_not_imp_right : a ∨ b ↔ (¬b → a) := Decidable.or_iff_not_imp_right +theorem not_imp_iff_and_not : ¬(a → b) ↔ a ∧ ¬b := + Decidable.not_imp_iff_and_not + +theorem not_and_iff_or_not_not : ¬(a ∧ b) ↔ ¬a ∨ ¬b := + Decidable.not_and_iff_or_not_not + +theorem not_iff : ¬(a ↔ b) ↔ (¬a ↔ b) := + Decidable.not_iff + end Classical /-! ## equality -/ @@ -923,13 +932,6 @@ example [Subsingleton α] (p : α → Prop) : Subsingleton (Subtype p) := theorem false_ne_true : False ≠ True := fun h => h.symm ▸ trivial -theorem Bool.eq_false_or_eq_true : (b : Bool) → b = true ∨ b = false - | true => .inl rfl - | false => .inr rfl - -theorem Bool.eq_false_iff {b : Bool} : b = false ↔ b ≠ true := - ⟨ne_true_of_eq_false, eq_false_of_ne_true⟩ - theorem Bool.eq_iff_iff {a b : Bool} : a = b ↔ (a ↔ b) := by cases b <;> simp theorem ne_comm {α} {a b : α} : a ≠ b ↔ b ≠ a := ⟨Ne.symm, Ne.symm⟩ diff --git a/Std/Tactic/Basic.lean b/Std/Tactic/Basic.lean index f03b59f39f..9872a7361d 100644 --- a/Std/Tactic/Basic.lean +++ b/Std/Tactic/Basic.lean @@ -1,165 +1,11 @@ -/- -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.ElabTerm import Std.Linter -import Std.Tactic.NoMatch -import Std.Tactic.GuardExpr import Std.Tactic.ByCases +import Std.Tactic.Init +import Std.Tactic.NoMatch import Std.Tactic.SeqFocus import Std.Tactic.ShowTerm import Std.Tactic.SimpTrace -import Lean.Elab.Tactic.ElabTerm -import Std.Lean.Meta.Basic -import Std.Lean.Tactic import Std.Util.ProofWanted -namespace Std.Tactic -open Lean Parser.Tactic Elab Command Elab.Tactic Meta - -/-- `exfalso` converts a goal `⊢ tgt` into `⊢ False` by applying `False.elim`. -/ -macro "exfalso" : tactic => `(tactic| refine False.elim ?_) - -/-- -`_` in tactic position acts like the `done` tactic: it fails and gives the list -of goals if there are any. It is useful as a placeholder after starting a tactic block -such as `by _` to make it syntactically correct and show the current goal. --/ -macro "_" : tactic => `(tactic| {}) - -@[inherit_doc failIfSuccess] -syntax (name := failIfSuccessConv) "fail_if_success " Conv.convSeq : conv - -attribute [tactic failIfSuccessConv] evalFailIfSuccess - -/-- We allow the `rfl` tactic to also use `Iff.rfl`. -/ --- `rfl` was defined earlier in Lean4, at src/Lean/Init/Tactics.lean --- Later we want to allow `rfl` to use all relations marked with an attribute. -macro_rules | `(tactic| rfl) => `(tactic| exact Iff.rfl) - -macro_rules | `(tactic| rfl) => `(tactic| exact HEq.rfl) - -/-- `rwa` calls `rw`, then closes any remaining goals using `assumption`. -/ -macro "rwa " rws:rwRuleSeq loc:(location)? : tactic => - `(tactic| (rw $rws:rwRuleSeq $[$loc:location]?; assumption)) - -/-- -Like `exact`, but takes a list of terms and checks that all goals are discharged after the tactic. --/ -elab (name := exacts) "exacts " "[" hs:term,* "]" : tactic => do - for stx in hs.getElems do - evalTactic (← `(tactic| exact $stx)) - evalTactic (← `(tactic| done)) - -/-- -`by_contra h` proves `⊢ p` by contradiction, -introducing a hypothesis `h : ¬p` and proving `False`. -* If `p` is a negation `¬q`, `h : q` will be introduced instead of `¬¬q`. -* If `p` is decidable, it uses `Decidable.byContradiction` instead of `Classical.byContradiction`. -* If `h` is omitted, the introduced variable `_: ¬p` will be anonymous. --/ -macro (name := byContra) tk:"by_contra" e?:(ppSpace colGt binderIdent)? : tactic => do - let e := match e? with - | some e => match e with - | `(binderIdent| $e:ident) => e - | e => Unhygienic.run `(_%$e) -- HACK: hover fails without Unhygienic here - | none => Unhygienic.run `(_%$tk) - `(tactic| first - | guard_target = Not _; intro $e:term - | refine Decidable.byContradiction fun $e => ?_ - | refine Classical.byContradiction fun $e => ?_) - -/-- -Given a proof `h` of `p`, `absurd h` changes the goal to `⊢ ¬ p`. -If `p` is a negation `¬q` then the goal is changed to `⊢ q` instead. --/ -macro "absurd " h:term : tactic => - `(tactic| first | refine absurd ?_ $h | refine absurd $h ?_) - -/-- -`iterate n tac` runs `tac` exactly `n` times. -`iterate tac` runs `tac` repeatedly until failure. - -To run multiple tactics, one can do `iterate (tac₁; tac₂; ⋯)` or -```lean -iterate - tac₁ - tac₂ - ⋯ -``` --/ -syntax "iterate" (ppSpace num)? ppSpace tacticSeq : tactic -macro_rules - | `(tactic| iterate $seq:tacticSeq) => - `(tactic| try ($seq:tacticSeq); iterate $seq:tacticSeq) - | `(tactic| iterate $n $seq:tacticSeq) => - match n.1.toNat with - | 0 => `(tactic| skip) - | n+1 => `(tactic| ($seq:tacticSeq); iterate $(quote n) $seq:tacticSeq) - -/-- -`repeat' tac` runs `tac` on all of the goals to produce a new list of goals, -then runs `tac` again on all of those goals, and repeats until `tac` fails on all remaining goals. --/ -elab "repeat' " tac:tacticSeq : tactic => do - setGoals (← repeat' (evalTacticAtRaw tac) (← getGoals)) - -/-- -`repeat1' tac` applies `tac` to main goal at least once. If the application succeeds, -the tactic is applied recursively to the generated subgoals until it eventually fails. --/ -elab "repeat1' " tac:tacticSeq : tactic => do - setGoals (← repeat1' (evalTacticAtRaw tac) (← getGoals)) - -/-- `subst_eqs` applies `subst` to all equalities in the context as long as it makes progress. -/ -elab "subst_eqs" : tactic => Elab.Tactic.liftMetaTactic1 (·.substEqs) - -/-- `split_ands` applies `And.intro` until it does not make progress. -/ -syntax "split_ands" : tactic -macro_rules | `(tactic| split_ands) => `(tactic| repeat' refine And.intro ?_ ?_) - -/-- -`fapply e` is like `apply e` but it adds goals in the order they appear, -rather than putting the dependent goals first. --/ -elab "fapply " e:term : tactic => - evalApplyLikeTactic (·.apply (cfg := {newGoals := .all})) e - -/-- -`eapply e` is like `apply e` but it does not add subgoals for variables that appear -in the types of other goals. Note that this can lead to a failure where there are -no goals remaining but there are still metavariables in the term: -``` -example (h : ∀ x : Nat, x = x → True) : True := by - eapply h - rfl - -- no goals --- (kernel) declaration has metavariables '_example' -``` --/ -elab "eapply " e:term : tactic => - evalApplyLikeTactic (·.apply (cfg := {newGoals := .nonDependentOnly})) e - -/-- -Tries to solve the goal using a canonical proof of `True`, or the `rfl` tactic. -Unlike `trivial` or `trivial'`, does not use the `contradiction` tactic. --/ -macro (name := triv) "triv" : tactic => - `(tactic| first | exact trivial | rfl | fail "triv tactic failed") - -/-- `conv` tactic to close a goal using an equality theorem. -/ -macro (name := Conv.exact) "exact " t:term : conv => `(conv| tactic => exact $t) - -/-- The `conv` tactic `equals` claims that the currently focused subexpression is equal - to the given expression, and proves this claim using the given tactic. -``` -example (P : (Nat → Nat) → Prop) : P (fun n => n - n) := by - conv in (_ - _) => equals 0 => - -- current goal: ⊢ n - n = 0 - apply Nat.sub_self - -- current goal: P (fun n => 0) -``` --/ -macro (name := Conv.equals) "equals " t:term " => " tac:tacticSeq : conv => - `(conv| tactic => show (_ = $t); next => $tac) +-- This is an import only file for common tactics used throughout Std diff --git a/Std/Tactic/Classical.lean b/Std/Tactic/Classical.lean new file mode 100644 index 0000000000..3bd386c858 --- /dev/null +++ b/Std/Tactic/Classical.lean @@ -0,0 +1,51 @@ +/- +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 + +/-! # `classical` and `classical!` tactics -/ + +namespace Std.Tactic +open Lean Meta + +/-- +`classical!` adds a proof of `Classical.propDecidable` as a local variable, which makes it +available for instance search and effectively makes all propositions decidable. +``` +noncomputable def foo : Bool := by + classical! + have := ∀ p, decide p -- uses the classical instance + exact decide (0 < 1) -- uses the classical instance even though `0 < 1` is decidable +``` +Consider using `classical` instead if you want to use the decidable instance when available. +-/ +macro (name := classical!) "classical!" : tactic => + `(tactic| have em := Classical.propDecidable) + +/-- +`classical tacs` runs `tacs` in a scope where `Classical.propDecidable` is a low priority +local instance. It differs from `classical!` in that `classical!` uses a local variable, +which has high priority: +``` +noncomputable def foo : Bool := by + classical! + have := ∀ p, decide p -- uses the classical instance + exact decide (0 < 1) -- uses the classical instance even though `0 < 1` is decidable + +def bar : Bool := by + classical + have := ∀ p, decide p -- uses the classical instance + exact decide (0 < 1) -- uses the decidable 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 + modifyEnv Meta.instanceExtension.pushScope + Meta.addInstance ``Classical.propDecidable .local 10 + try Elab.Tactic.evalTactic tacs + finally modifyEnv Meta.instanceExtension.popScope diff --git a/Std/Tactic/CoeExt.lean b/Std/Tactic/CoeExt.lean index 1e0e464f72..8f5508f617 100644 --- a/Std/Tactic/CoeExt.lean +++ b/Std/Tactic/CoeExt.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner, Mario Carneiro -/ import Lean.PrettyPrinter.Delaborator.Builtins -import Std.Lean.Delaborator + open Lean Elab.Term Meta Std /-! diff --git a/Std/Tactic/Congr.lean b/Std/Tactic/Congr.lean index 1096d9efb0..7a784feb95 100644 --- a/Std/Tactic/Congr.lean +++ b/Std/Tactic/Congr.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Miyahara Kō -/ import Lean.Meta.Tactic.Congr +import Lean.Elab.Tactic.Config import Std.Tactic.RCases import Std.Tactic.Ext diff --git a/Std/Tactic/Ext.lean b/Std/Tactic/Ext.lean index fcce1fc297..c4629e65af 100644 --- a/Std/Tactic/Ext.lean +++ b/Std/Tactic/Ext.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Gabriel Ebner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner, Mario Carneiro -/ -import Std.Tactic.Basic import Std.Tactic.RCases import Std.Tactic.Ext.Attr diff --git a/Std/Tactic/Ext/Attr.lean b/Std/Tactic/Ext/Attr.lean index 88877e16fc..e7b19480ef 100644 --- a/Std/Tactic/Ext/Attr.lean +++ b/Std/Tactic/Ext/Attr.lean @@ -39,7 +39,7 @@ initialize extExtension : SimpleScopedEnvExtension ExtTheorem ExtTheorems ← registerSimpleScopedEnvExtension { addEntry := fun { tree, erased } thm => - { tree := tree.insertCore thm.keys thm extExt.config, erased := erased.erase thm.declName } + { tree := tree.insertCore thm.keys thm, erased := erased.erase thm.declName } initial := {} } diff --git a/Std/Tactic/Init.lean b/Std/Tactic/Init.lean new file mode 100644 index 0000000000..075e179ed8 --- /dev/null +++ b/Std/Tactic/Init.lean @@ -0,0 +1,161 @@ +/- +Copyright (c) 2021 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ +import Std.Tactic.GuardExpr +import Std.Lean.Meta.Basic +import Std.Lean.Tactic + +/-! +# Simple tactics that are used throughout Std. +-/ + +namespace Std.Tactic +open Lean Parser.Tactic Elab Command Elab.Tactic Meta + +/-- `exfalso` converts a goal `⊢ tgt` into `⊢ False` by applying `False.elim`. -/ +macro "exfalso" : tactic => `(tactic| refine False.elim ?_) + +/-- +`_` in tactic position acts like the `done` tactic: it fails and gives the list +of goals if there are any. It is useful as a placeholder after starting a tactic block +such as `by _` to make it syntactically correct and show the current goal. +-/ +macro "_" : tactic => `(tactic| {}) + +@[inherit_doc failIfSuccess] +syntax (name := failIfSuccessConv) "fail_if_success " Conv.convSeq : conv + +attribute [tactic failIfSuccessConv] evalFailIfSuccess + +/-- We allow the `rfl` tactic to also use `Iff.rfl`. -/ +-- `rfl` was defined earlier in Lean4, at src/Lean/Init/Tactics.lean +-- Later we want to allow `rfl` to use all relations marked with an attribute. +macro_rules | `(tactic| rfl) => `(tactic| exact Iff.rfl) + +macro_rules | `(tactic| rfl) => `(tactic| exact HEq.rfl) + +/-- `rwa` calls `rw`, then closes any remaining goals using `assumption`. -/ +macro "rwa " rws:rwRuleSeq loc:(location)? : tactic => + `(tactic| (rw $rws:rwRuleSeq $[$loc:location]?; assumption)) + +/-- +Like `exact`, but takes a list of terms and checks that all goals are discharged after the tactic. +-/ +elab (name := exacts) "exacts " "[" hs:term,* "]" : tactic => do + for stx in hs.getElems do + evalTactic (← `(tactic| exact $stx)) + evalTactic (← `(tactic| done)) + +/-- +`by_contra h` proves `⊢ p` by contradiction, +introducing a hypothesis `h : ¬p` and proving `False`. +* If `p` is a negation `¬q`, `h : q` will be introduced instead of `¬¬q`. +* If `p` is decidable, it uses `Decidable.byContradiction` instead of `Classical.byContradiction`. +* If `h` is omitted, the introduced variable `_: ¬p` will be anonymous. +-/ +macro (name := byContra) tk:"by_contra" e?:(ppSpace colGt binderIdent)? : tactic => do + let e := match e? with + | some e => match e with + | `(binderIdent| $e:ident) => e + | e => Unhygienic.run `(_%$e) -- HACK: hover fails without Unhygienic here + | none => Unhygienic.run `(_%$tk) + `(tactic| first + | guard_target = Not _; intro $e:term + | refine Decidable.byContradiction fun $e => ?_ + | refine Classical.byContradiction fun $e => ?_) + +/-- +Given a proof `h` of `p`, `absurd h` changes the goal to `⊢ ¬ p`. +If `p` is a negation `¬q` then the goal is changed to `⊢ q` instead. +-/ +macro "absurd " h:term : tactic => + `(tactic| first | refine absurd ?_ $h | refine absurd $h ?_) + +/-- +`iterate n tac` runs `tac` exactly `n` times. +`iterate tac` runs `tac` repeatedly until failure. + +To run multiple tactics, one can do `iterate (tac₁; tac₂; ⋯)` or +```lean +iterate + tac₁ + tac₂ + ⋯ +``` +-/ +syntax "iterate" (ppSpace num)? ppSpace tacticSeq : tactic +macro_rules + | `(tactic| iterate $seq:tacticSeq) => + `(tactic| try ($seq:tacticSeq); iterate $seq:tacticSeq) + | `(tactic| iterate $n $seq:tacticSeq) => + match n.1.toNat with + | 0 => `(tactic| skip) + | n+1 => `(tactic| ($seq:tacticSeq); iterate $(quote n) $seq:tacticSeq) + +/-- +`repeat' tac` runs `tac` on all of the goals to produce a new list of goals, +then runs `tac` again on all of those goals, and repeats until `tac` fails on all remaining goals. +-/ +elab "repeat' " tac:tacticSeq : tactic => do + setGoals (← repeat' (evalTacticAtRaw tac) (← getGoals)) + +/-- +`repeat1' tac` applies `tac` to main goal at least once. If the application succeeds, +the tactic is applied recursively to the generated subgoals until it eventually fails. +-/ +elab "repeat1' " tac:tacticSeq : tactic => do + setGoals (← repeat1' (evalTacticAtRaw tac) (← getGoals)) + +/-- `subst_eqs` applies `subst` to all equalities in the context as long as it makes progress. -/ +elab "subst_eqs" : tactic => Elab.Tactic.liftMetaTactic1 (·.substEqs) + +/-- `split_ands` applies `And.intro` until it does not make progress. -/ +syntax "split_ands" : tactic +macro_rules | `(tactic| split_ands) => `(tactic| repeat' refine And.intro ?_ ?_) + +/-- +`fapply e` is like `apply e` but it adds goals in the order they appear, +rather than putting the dependent goals first. +-/ +elab "fapply " e:term : tactic => + evalApplyLikeTactic (·.apply (cfg := {newGoals := .all})) e + +/-- +`eapply e` is like `apply e` but it does not add subgoals for variables that appear +in the types of other goals. Note that this can lead to a failure where there are +no goals remaining but there are still metavariables in the term: +``` +example (h : ∀ x : Nat, x = x → True) : True := by + eapply h + rfl + -- no goals +-- (kernel) declaration has metavariables '_example' +``` +-/ +elab "eapply " e:term : tactic => + evalApplyLikeTactic (·.apply (cfg := {newGoals := .nonDependentOnly})) e + +/-- +Tries to solve the goal using a canonical proof of `True`, or the `rfl` tactic. +Unlike `trivial` or `trivial'`, does not use the `contradiction` tactic. +-/ +macro (name := triv) "triv" : tactic => + `(tactic| first | exact trivial | rfl | fail "triv tactic failed") + +/-- `conv` tactic to close a goal using an equality theorem. -/ +macro (name := Conv.exact) "exact " t:term : conv => `(conv| tactic => exact $t) + +/-- The `conv` tactic `equals` claims that the currently focused subexpression is equal + to the given expression, and proves this claim using the given tactic. +``` +example (P : (Nat → Nat) → Prop) : P (fun n => n - n) := by + conv in (_ - _) => equals 0 => + -- current goal: ⊢ n - n = 0 + apply Nat.sub_self + -- current goal: P (fun n => 0) +``` +-/ +macro (name := Conv.equals) "equals " t:term " => " tac:tacticSeq : conv => + `(conv| tactic => show (_ = $t); next => $tac) diff --git a/Std/Tactic/LibrarySearch.lean b/Std/Tactic/LibrarySearch.lean index 59e62005f0..f5829ef49d 100644 --- a/Std/Tactic/LibrarySearch.lean +++ b/Std/Tactic/LibrarySearch.lean @@ -7,8 +7,10 @@ import Std.Lean.CoreM import Std.Lean.Expr import Std.Lean.Meta.DiscrTree import Std.Lean.Meta.LazyDiscrTree +import Std.Lean.Parser import Std.Data.Option.Basic import Std.Tactic.SolveByElim +import Std.Tactic.TryThis import Std.Util.Pickle /-! @@ -88,7 +90,7 @@ def cachePath : IO FilePath := do private def addPath [BEq α] (config : WhnfCoreConfig) (tree : DiscrTree α) (tp : Expr) (v : α) : MetaM (DiscrTree α) := do let k ← DiscrTree.mkPath tp config - pure <| tree.insertCore k v config + pure <| tree.insertCore k v /-- Adds a constant with given name to tree. -/ private def updateTree (config : WhnfCoreConfig) (tree : DiscrTree (Name × DeclMod)) @@ -221,7 +223,7 @@ If `solveByElim` succeeds, we return `[]` as the list of new subgoals, otherwise the full list of subgoals. -/ private def librarySearchLemma (cfg : ApplyConfig) (act : List MVarId → MetaM (List MVarId)) - (allowFailure : Bool) (cand : Candidate) : MetaM (List MVarId) := do + (allowFailure : MVarId → MetaM Bool) (cand : Candidate) : MetaM (List MVarId) := do let ((goal, mctx), (name, mod)) := cand withTraceNode `Tactic.stdLibrarySearch (return m!"{emoji ·} trying {name} with {mod} ") do setMCtx mctx @@ -230,7 +232,7 @@ private def librarySearchLemma (cfg : ApplyConfig) (act : List MVarId → MetaM try act newGoals catch _ => - if allowFailure then + if ← allowFailure goal then pure newGoals else failure @@ -366,7 +368,8 @@ def solveByElim (required : List Expr) (exfalso : Bool) (goals : List MVarId) (m -- There is only a marginal decrease in performance for using the `symm` option for `solveByElim`. -- (measured via `lake build && time lake env lean test/librarySearch.lean`). let cfg : SolveByElim.Config := - { maxDepth, exfalso := exfalso, symm := true, commitIndependentGoals := true } + { maxDepth, exfalso := exfalso, symm := true, commitIndependentGoals := true, + transparency := ← getTransparency } let ⟨lemmas, ctx⟩ ← SolveByElim.mkAssumptionSet false false [] [] #[] let cfg := if !required.isEmpty then cfg.requireUsingAll required else cfg SolveByElim.solveByElim cfg lemmas ctx goals @@ -390,7 +393,7 @@ private def librarySearchEmoji : Except ε (Option α) → String private def librarySearch' (goal : MVarId) (tactic : List MVarId → MetaM (List MVarId)) - (allowFailure : Bool) + (allowFailure : MVarId → MetaM Bool) (leavePercentHeartbeats : Nat) : MetaM (Option (Array (List MVarId × MetavarContext))) := do withTraceNode `Tactic.stdLibrarySearch (return m!"{librarySearchEmoji ·} {← goal.getType}") do @@ -438,7 +441,7 @@ this is not currently tracked.) -/ def librarySearch (goal : MVarId) (tactic : Bool → List MVarId → MetaM (List MVarId)) - (allowFailure : Bool := true) + (allowFailure : MVarId → MetaM Bool := fun _ => pure true) (leavePercentHeartbeats : Nat := 10) : MetaM (Option (Array (List MVarId × MetavarContext))) := do (tactic true [goal] *> pure none) <|> @@ -478,7 +481,10 @@ def exact? (tk : Syntax) (required : Option (Array (TSyntax `term))) | some t => let _ <- mkInitialTacticInfo t throwError "Do not yet support custom std_exact?/std_apply? tactics." - match ← librarySearch goal tactic (allowFailure := required.isEmpty) with + let allowFailure := fun g => do + let g ← g.withContext (instantiateMVars (.mvar g)) + return required.all fun e => e.occurs g + match ← librarySearch goal tactic allowFailure with -- Found goal that closed problem | none => addExactSuggestion tk (← instantiateMVars (mkMVar mvar)).headBeta diff --git a/Std/Tactic/Lint/Misc.lean b/Std/Tactic/Lint/Misc.lean index 1d328b63cc..77773eb5a7 100644 --- a/Std/Tactic/Lint/Misc.lean +++ b/Std/Tactic/Lint/Misc.lean @@ -226,11 +226,13 @@ Return a list of unused `let_fun` terms in an expression. def findUnusedHaves (e : Expr) : MetaM (Array MessageData) := do let res ← IO.mkRef #[] forEachExpr e fun e => do - let some (n, t, _v, b) := e.letFun? | return - if n.isInternal then return - if b.hasLooseBVars then return - let msg ← addMessageContextFull m!"unnecessary have {n.eraseMacroScopes} : {t}" - res.modify (·.push msg) + match e.letFun? with + | some (n, t, _, b) => + if n.isInternal then return + if b.hasLooseBVars then return + let msg ← addMessageContextFull m!"unnecessary have {n.eraseMacroScopes} : {t}" + res.modify (·.push msg) + | _ => return res.get /-- A linter for checking that declarations don't have unused term mode have statements. We do not diff --git a/Std/Tactic/Lint/Simp.lean b/Std/Tactic/Lint/Simp.lean index 525107e803..b467032f0c 100644 --- a/Std/Tactic/Lint/Simp.lean +++ b/Std/Tactic/Lint/Simp.lean @@ -112,10 +112,10 @@ 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 (⟨lhs', prf1, _⟩, prf1Lems) ← + let ({ expr := lhs', proof? := prf1, .. }, prf1Lems) ← decorateError "simplify fails on left-hand side:" <| simp lhs ctx if prf1Lems.contains (.decl declName) then return none - let (⟨rhs', _, _⟩, used_lemmas) ← + let ({ expr := rhs', .. }, used_lemmas) ← decorateError "simplify fails on right-hand side:" <| simp rhs ctx (usedSimps := prf1Lems) let lhs'EqRhs' ← isSimpEq lhs' rhs' (whnfFirst := false) let lhsInNF ← isSimpEq lhs' lhs diff --git a/Std/Tactic/NormCast.lean b/Std/Tactic/NormCast.lean index e894460106..0dafd4770a 100644 --- a/Std/Tactic/NormCast.lean +++ b/Std/Tactic/NormCast.lean @@ -24,13 +24,13 @@ initialize registerTraceClass `Tactic.norm_cast /-- Prove `a = b` using the given simp set. -/ def proveEqUsing (s : SimpTheorems) (a b : Expr) : MetaM (Option Simp.Result) := do let go : SimpM (Option Simp.Result) := do - let methods := Simp.DefaultMethods.methods - let a' ← Simp.simp a methods - let b' ← Simp.simp b methods + let a' ← Simp.simp a + let b' ← Simp.simp b unless ← isDefEq a'.expr b'.expr do return none - mkEqTrans a' (← mkEqSymm b b') + a'.mkEqTrans (← mkEqSymm b b') withReducible do - (go { simpTheorems := #[s], congrTheorems := ← Meta.getSimpCongrTheorems }).run' {} + (go (← Simp.mkDefaultMethods).toMethodsRef + { simpTheorems := #[s], congrTheorems := ← Meta.getSimpCongrTheorems }).run' {} /-- Prove `a = b` by simplifying using move and squash lemmas. -/ def proveEqUsingDown (a b : Expr) : MetaM (Option Simp.Result) := do @@ -136,9 +136,10 @@ It tries to rewrite an expression using the elim and move lemmas. On failure, it calls the splitting procedure heuristic. -/ partial def upwardAndElim (up : SimpTheorems) (e : Expr) : SimpM Simp.Step := do - let r ← Simp.rewrite? e up.post up.erased prove (tag := "squash") (rflOnly := false) + let r ← withDischarger prove do + Simp.rewrite? e up.post up.erased (tag := "squash") (rflOnly := false) let r := r.getD { expr := e } - let r ← mkEqTrans r <|← splittingProcedure r.expr + let r ← r.mkEqTrans (← splittingProcedure r.expr) if r.expr == e then return Simp.Step.done {expr := e} return Simp.Step.visit r @@ -169,7 +170,7 @@ def derive (e : Expr) : MetaM Simp.Result := do } let congrTheorems ← Meta.getSimpCongrTheorems - let r := {expr := e} + let r : Simp.Result := { expr := e } let withTrace phase := withTraceNode `Tactic.norm_cast fun | .ok r => return m!"{r.expr} (after {phase})" @@ -178,17 +179,17 @@ def derive (e : Expr) : MetaM Simp.Result := do -- step 1: pre-processing of numerals let r ← withTrace "pre-processing numerals" do let post e := return Simp.Step.done (← try numeralToCoe e catch _ => pure {expr := e}) - Simp.mkEqTrans r (← Simp.main r.expr { config, congrTheorems } (methods := { post })).1 + r.mkEqTrans (← Simp.main r.expr { config, congrTheorems } (methods := { post })).1 -- step 2: casts are moved upwards and eliminated let r ← withTrace "moving upward, splitting and eliminating" do let post := upwardAndElim (← normCastExt.up.getTheorems) - Simp.mkEqTrans r (← Simp.main r.expr { config, congrTheorems } (methods := { post })).1 + r.mkEqTrans (← Simp.main r.expr { config, congrTheorems } (methods := { post })).1 -- step 3: casts are squashed let r ← withTrace "squashing" do let simpTheorems := #[← normCastExt.squash.getTheorems] - mkEqTrans r (← simp r.expr { simpTheorems, config, congrTheorems }).1 + r.mkEqTrans (← simp r.expr { simpTheorems, config, congrTheorems }).1 return r @@ -204,7 +205,7 @@ elab "mod_cast " e:term : term <= expectedType => do let eTy' ← derive eTy unless ← isDefEq eTy'.expr expectedType'.expr do throwTypeMismatchError "mod_cast" expectedType'.expr eTy'.expr e - let eTy_eq_expectedType ← mkEqTrans eTy' (← mkEqSymm expectedType expectedType') + let eTy_eq_expectedType ← eTy'.mkEqTrans (← mkEqSymm expectedType expectedType') mkCast eTy_eq_expectedType e open Tactic Parser.Tactic Elab.Tactic @@ -325,8 +326,8 @@ syntax (name := pushCast) "push_cast" (config)? (discharger)? (&" only")? (" [" (simpStar <|> simpErase <|> simpLemma),* "]")? (location)? : tactic @[inherit_doc pushCast, tactic pushCast] def evalPushCast : Tactic := fun stx => do - let { ctx, dischargeWrapper, .. } ← withMainContext do + let { ctx, simprocs, dischargeWrapper } ← withMainContext do mkSimpContext' (← pushCastExt.getTheorems) stx (eraseLocal := false) let ctx := { ctx with config := { ctx.config with failIfUnchanged := false } } dischargeWrapper.with fun discharge? => - discard <| simpLocation ctx discharge? (expandOptLocation stx[5]) + discard <| simpLocation ctx simprocs discharge? (expandOptLocation stx[5]) diff --git a/Std/Tactic/Omega.lean b/Std/Tactic/Omega.lean index 766b49f000..2e9cff9076 100644 --- a/Std/Tactic/Omega.lean +++ b/Std/Tactic/Omega.lean @@ -1,3 +1,8 @@ +/- +Copyright (c) 2023 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison +-/ import Std.Tactic.Omega.Frontend diff --git a/Std/Tactic/Omega/Coeffs/IntList.lean b/Std/Tactic/Omega/Coeffs/IntList.lean index bcae0c9927..94be9e134c 100644 --- a/Std/Tactic/Omega/Coeffs/IntList.lean +++ b/Std/Tactic/Omega/Coeffs/IntList.lean @@ -1,4 +1,10 @@ +/- +Copyright (c) 2023 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison +-/ import Std.Tactic.Omega.IntList +import Std.Data.List.Basic /-! # `Coeffs` as a wrapper for `IntList` @@ -59,6 +65,11 @@ abbrev length (xs : Coeffs) := List.length xs abbrev leading (xs : Coeffs) : Int := IntList.leading xs /-- Shim for `List.map`. -/ abbrev map (f : Int → Int) (xs : Coeffs) : Coeffs := List.map f xs +/-- Shim for `.enum.find?`. -/ +abbrev findIdx? (f : Int → Bool) (xs : Coeffs) : Option Nat := + List.findIdx? f xs + -- We could avoid `Std.Data.List.Basic` by using the less efficient: + -- xs.enum.find? (f ·.2) |>.map (·.1) /-- Shim for `IntList.bmod`. -/ abbrev bmod (x : Coeffs) (m : Nat) : Coeffs := IntList.bmod x m /-- Shim for `IntList.bmod_dot_sub_dot_bmod`. -/ diff --git a/Std/Tactic/Omega/Config.lean b/Std/Tactic/Omega/Config.lean index ea61f87280..eda2b57b7c 100644 --- a/Std/Tactic/Omega/Config.lean +++ b/Std/Tactic/Omega/Config.lean @@ -32,6 +32,11 @@ structure OmegaConfig where `0 ≤ a ∧ Int.natAbs a = a ∨ a < 0 ∧ Int.natAbs a = - a` for later splitting. -/ splitNatAbs : Bool := true + /-- + Whenever `min a b` or `max a b` is found, rewrite in terms of the definition + `if a ≤ b ...`, for later case splitting. + -/ + splitMinMax : Bool := true /-- Allow elaboration of `OmegaConfig` arguments to tactics. diff --git a/Std/Tactic/Omega/Constraint.lean b/Std/Tactic/Omega/Constraint.lean index a4834d0936..6d81f585ac 100644 --- a/Std/Tactic/Omega/Constraint.lean +++ b/Std/Tactic/Omega/Constraint.lean @@ -5,7 +5,6 @@ Authors: Scott Morrison -/ import Std.Classes.Order import Std.Tactic.RCases -import Std.Tactic.NormCast import Std.Tactic.Omega.Coeffs.IntList /-! @@ -212,26 +211,26 @@ def div (c : Constraint) (k : Nat) : Constraint where theorem div_sat (c : Constraint) (t : Int) (k : Nat) (n : k ≠ 0) (h : (k : Int) ∣ t) (w : c.sat t) : (c.div k).sat (t / k) := by - replace n : k > 0 := Nat.pos_of_ne_zero n + replace n : (k : Int) > 0 := Int.ofNat_lt.mpr (Nat.pos_of_ne_zero n) rcases c with ⟨_ | l, _ | u⟩ · simp_all [sat, div] - · simp_all [sat, div] + · simp [sat, div] at w ⊢ apply Int.le_of_sub_nonneg - rw [← Int.sub_ediv_of_dvd _ h, ← ge_iff_le, Int.div_nonneg_iff_of_pos (mod_cast n)] + rw [← Int.sub_ediv_of_dvd _ h, ← ge_iff_le, Int.div_nonneg_iff_of_pos n] exact Int.sub_nonneg_of_le w - · simp_all [sat, div] + · simp [sat, div] at w ⊢ apply Int.le_of_sub_nonneg rw [Int.sub_neg, ← Int.add_ediv_of_dvd_left h, ← ge_iff_le, - Int.div_nonneg_iff_of_pos (mod_cast n)] + Int.div_nonneg_iff_of_pos n] exact Int.sub_nonneg_of_le w - · simp_all [sat, div] + · simp [sat, div] at w ⊢ constructor · apply Int.le_of_sub_nonneg rw [Int.sub_neg, ← Int.add_ediv_of_dvd_left h, ← ge_iff_le, - Int.div_nonneg_iff_of_pos (mod_cast n)] + Int.div_nonneg_iff_of_pos n] exact Int.sub_nonneg_of_le w.1 · apply Int.le_of_sub_nonneg - rw [← Int.sub_ediv_of_dvd _ h, ← ge_iff_le, Int.div_nonneg_iff_of_pos (mod_cast n)] + rw [← Int.sub_ediv_of_dvd _ h, ← ge_iff_le, Int.div_nonneg_iff_of_pos n] exact Int.sub_nonneg_of_le w.2 /-- diff --git a/Std/Tactic/Omega/Core.lean b/Std/Tactic/Omega/Core.lean index 654319f512..fd5738f601 100644 --- a/Std/Tactic/Omega/Core.lean +++ b/Std/Tactic/Omega/Core.lean @@ -6,9 +6,8 @@ Authors: Scott Morrison import Std.Tactic.Omega.OmegaM import Std.Tactic.Omega.Constraint import Std.Tactic.Omega.MinNatAbs -import Std.Data.HashMap.Basic -open Lean (HashSet) +open Lean (HashMap HashSet) namespace Std.Tactic.Omega @@ -216,10 +215,6 @@ def comboProof (s t : Constraint) (a : Int) (x : Coeffs) (b : Int) (y : Coeffs) mkApp9 (.const ``combo_sat' []) (toExpr s) (toExpr t) (toExpr a) (toExpr x) (toExpr b) (toExpr y) v px py -/-- Construct the term with type hint `(Eq.refl a : a = b)`-/ -def mkEqReflWithExpectedType (a b : Expr) : MetaM Expr := do - mkExpectedTypeHint (← mkEqRefl a) (← mkEq a b) - /-- Construct the proof term associated to a `bmod` step. -/ def bmodProof (m : Nat) (r : Int) (i : Nat) (x : Coeffs) (v : Expr) (w : Expr) : MetaM Expr := do let m := toExpr m diff --git a/Std/Tactic/Omega/Frontend.lean b/Std/Tactic/Omega/Frontend.lean index 16b64475f4..7ed447c428 100644 --- a/Std/Tactic/Omega/Frontend.lean +++ b/Std/Tactic/Omega/Frontend.lean @@ -49,10 +49,6 @@ structure MetaProblem where /-- Facts which have already been processed; we keep these to avoid duplicates. -/ processedFacts : HashSet Expr := ∅ -/-- Construct the term with type hint `(Eq.refl a : a = b)`-/ -def mkEqReflWithExpectedType (a b : Expr) : MetaM Expr := do - mkExpectedTypeHint (← mkEqRefl a) (← mkEq a b) - /-- Construct the `rfl` proof that `lc.eval atoms = e`. -/ def mkEvalRflProof (e : Expr) (lc : LinearCombo) : OmegaM Expr := do mkEqReflWithExpectedType e (mkApp2 (.const ``LinearCombo.eval []) (toExpr lc) (← atomsCoeffs)) @@ -126,7 +122,13 @@ partial def asLinearComboImpl (e : Expr) : OmegaM (LinearCombo × OmegaM Expr × | some i => let lc := {const := i} return ⟨lc, mkEvalRflProof e lc, ∅⟩ - | none => match e.getAppFnArgs with + | none => + if e.isFVar then + if let some v ← e.fvarId!.getValue? then + rewrite e (← mkEqReflWithExpectedType e v) + else + mkAtomLinearCombo e + else match e.getAppFnArgs with | (``HAdd.hAdd, #[_, _, _, _, e₁, e₂]) => do let (l₁, prf₁, facts₁) ← asLinearCombo e₁ let (l₂, prf₂, facts₂) ← asLinearCombo e₂ @@ -156,22 +158,32 @@ partial def asLinearComboImpl (e : Expr) : OmegaM (LinearCombo × OmegaM Expr × (← mkEqSymm neg_eval) pure (-l, prf', facts) | (``HMul.hMul, #[_, _, _, _, x, y]) => - let (xl, xprf, xfacts) ← asLinearCombo x - let (yl, yprf, yfacts) ← asLinearCombo y - if xl.coeffs.isZero ∨ yl.coeffs.isZero then - let prf : OmegaM Expr := do - let h ← mkDecideProof (mkApp2 (.const ``Or []) - (.app (.const ``Coeffs.isZero []) (toExpr xl.coeffs)) - (.app (.const ``Coeffs.isZero []) (toExpr yl.coeffs))) - let mul_eval := - mkApp4 (.const ``LinearCombo.mul_eval []) (toExpr xl) (toExpr yl) (← atomsCoeffs) h - mkEqTrans - (← mkAppM ``Int.mul_congr #[← xprf, ← yprf]) - (← mkEqSymm mul_eval) - pure (LinearCombo.mul xl yl, prf, xfacts.merge yfacts) - else - mkAtomLinearCombo e - | (``HMod.hMod, #[_, _, _, _, n, k]) => rewrite e (mkApp2 (.const ``Int.emod_def []) n k) + -- If we decide not to expand out the multiplication, + -- we have to revert the `OmegaM` state so that any new facts about the factors + -- can still be reported when they are visited elsewhere. + let r? ← commitWhen do + let (xl, xprf, xfacts) ← asLinearCombo x + let (yl, yprf, yfacts) ← asLinearCombo y + if xl.coeffs.isZero ∨ yl.coeffs.isZero then + let prf : OmegaM Expr := do + let h ← mkDecideProof (mkApp2 (.const ``Or []) + (.app (.const ``Coeffs.isZero []) (toExpr xl.coeffs)) + (.app (.const ``Coeffs.isZero []) (toExpr yl.coeffs))) + let mul_eval := + mkApp4 (.const ``LinearCombo.mul_eval []) (toExpr xl) (toExpr yl) (← atomsCoeffs) h + mkEqTrans + (← mkAppM ``Int.mul_congr #[← xprf, ← yprf]) + (← mkEqSymm mul_eval) + pure (some (LinearCombo.mul xl yl, prf, xfacts.merge yfacts), true) + else + pure (none, false) + match r? with + | some r => pure r + | none => mkAtomLinearCombo e + | (``HMod.hMod, #[_, _, _, _, n, k]) => + match natCast? k with + | some _ => rewrite e (mkApp2 (.const ``Int.emod_def []) n k) + | none => mkAtomLinearCombo e | (``HDiv.hDiv, #[_, _, _, _, x, z]) => match intCast? z with | some 0 => rewrite e (mkApp (.const ``Int.ediv_zero []) x) @@ -181,7 +193,25 @@ partial def asLinearComboImpl (e : Expr) : OmegaM (LinearCombo × OmegaM Expr × else mkAtomLinearCombo e | _ => mkAtomLinearCombo e - | (``Nat.cast, #[_, _, n]) => match n.getAppFnArgs with + | (``Min.min, #[_, _, a, b]) => + if (← cfg).splitMinMax then + rewrite e (mkApp2 (.const ``Int.min_def []) a b) + else + mkAtomLinearCombo e + | (``Max.max, #[_, _, a, b]) => + if (← cfg).splitMinMax then + rewrite e (mkApp2 (.const ``Int.max_def []) a b) + else + mkAtomLinearCombo e + | (``Nat.cast, #[.const ``Int [], i, n]) => + match n with + | .fvar h => + if let some v ← h.getValue? then + rewrite e (← mkEqReflWithExpectedType e + (mkApp3 (.const ``Nat.cast [0]) (.const ``Int []) i v)) + else + mkAtomLinearCombo e + | _ => match n.getAppFnArgs with | (``Nat.succ, #[n]) => rewrite e (.app (.const ``Int.ofNat_succ []) n) | (``HAdd.hAdd, #[_, _, _, _, a, b]) => rewrite e (mkApp2 (.const ``Int.ofNat_add []) a b) | (``HMul.hMul, #[_, _, _, _, a, b]) => rewrite e (mkApp2 (.const ``Int.ofNat_mul []) a b) @@ -198,6 +228,13 @@ partial def asLinearComboImpl (e : Expr) : OmegaM (LinearCombo × OmegaM Expr × | .app (.app (.app (.app (.const ``Prod.mk [u, 0]) _) _) x) y => rewrite e (mkApp3 (.const ``Int.ofNat_snd_mk [u]) α x y) | _ => mkAtomLinearCombo e + | (``Min.min, #[_, _, a, b]) => rewrite e (mkApp2 (.const ``Int.ofNat_min []) a b) + | (``Max.max, #[_, _, a, b]) => rewrite e (mkApp2 (.const ``Int.ofNat_max []) a b) + | (``Int.natAbs, #[n]) => + if (← cfg).splitNatAbs then + rewrite e (mkApp (.const ``Int.ofNat_natAbs []) n) + else + mkAtomLinearCombo e | _ => mkAtomLinearCombo e | (``Prod.fst, #[α, β, p]) => match p with | .app (.app (.app (.app (.const ``Prod.mk [u, v]) _) _) x) y => @@ -361,7 +398,7 @@ partial def addFact (p : MetaProblem) (h : Expr) : OmegaM (MetaProblem × Nat) : | (``Dvd.dvd, #[.const ``Nat [], _, k, x]) => p.addFact (mkApp3 (.const ``Nat.mod_eq_zero_of_dvd []) k x h) | (``Dvd.dvd, #[.const ``Int [], _, k, x]) => - p.addFact (mkApp3 (.const ``Int.mod_eq_zero_of_dvd []) k x h) + p.addFact (mkApp3 (.const ``Int.emod_eq_zero_of_dvd []) k x h) | (``And, #[t₁, t₂]) => do let (p₁, n₁) ← p.addFact (mkApp3 (.const ``And.left []) t₁ t₂ h) let (p₂, n₂) ← p₁.addFact (mkApp3 (.const ``And.right []) t₁ t₂ h) @@ -434,7 +471,7 @@ partial def splitDisjunction (m : MetaProblem) (g : MVarId) : OmegaM Unit := g.w let (⟨g₁, h₁⟩, ⟨g₂, h₂⟩) ← cases₂ g h trace[omega] "Adding facts:\n{← g₁.withContext <| inferType (.fvar h₁)}" let m₁ := { m with facts := [.fvar h₁], disjunctions := t } - let r ← savingState do + let r ← withoutModifyingState do let (m₁, n) ← g₁.withContext m₁.processFacts if 0 < n then omegaImpl m₁ g₁ @@ -514,13 +551,15 @@ a natural subtraction appearing in a hypothesis, and try again. The options ``` -omega (config := { splitDisjunctions := true, splitNatSub := true, splitNatAbs := true}) +omega (config := + { splitDisjunctions := true, splitNatSub := true, splitNatAbs := true, splitMinMax := true }) ``` can be used to: * `splitDisjunctions`: split any disjunctions found in the context, if the problem is not otherwise solvable. * `splitNatSub`: for each appearance of `((a - b : Nat) : Int)`, split on `a ≤ b` if necessary. * `splitNatAbs`: for each appearance of `Int.natAbs a`, split on `0 ≤ a` if necessary. +* `splitMinMax`: for each occurrence of `min a b`, split on `min a b = a ∨ min a b = b` Currently, all of these are on by default. -/ syntax (name := omegaSyntax) "omega" (config)? : tactic diff --git a/Std/Tactic/Omega/Int.lean b/Std/Tactic/Omega/Int.lean index 84bd182075..147de0ec31 100644 --- a/Std/Tactic/Omega/Int.lean +++ b/Std/Tactic/Omega/Int.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import Std.Classes.Order -import Std.Data.Int.Lemmas +import Std.Data.Int.Init.Order import Std.Data.Prod.Lex import Std.Tactic.LeftRight @@ -19,6 +19,20 @@ If you do find a use for them, please move them into the appropriate file and na namespace Std.Tactic.Omega.Int +theorem ofNat_pow (a b : Nat) : ((a ^ b : Nat) : Int) = (a : Int) ^ b := by + induction b with + | zero => rfl + | succ b ih => rw [Nat.pow_succ, Int.ofNat_mul, ih]; rfl + +theorem pos_pow_of_pos (a : Int) (b : Nat) (h : 0 < a) : 0 < a ^ b := by + rw [Int.eq_natAbs_of_zero_le (Int.le_of_lt h), ← Int.ofNat_zero, ← Int.ofNat_pow, Int.ofNat_lt] + exact Nat.pos_pow_of_pos _ (Int.natAbs_pos.mpr (Int.ne_of_gt h)) + +theorem ofNat_pos {a : Nat} : 0 < (a : Int) ↔ 0 < a := by + rw [← Int.ofNat_zero, Int.ofNat_lt] + +alias ⟨_, ofNat_pos_of_pos⟩ := Int.ofNat_pos + theorem natCast_ofNat {x : Nat} : @Nat.cast Int instNatCastInt (no_index (OfNat.ofNat x)) = OfNat.ofNat x := rfl @@ -66,6 +80,21 @@ theorem ofNat_congr {a b : Nat} (h : a = b) : (a : Int) = (b : Int) := congrArg theorem ofNat_sub_sub {a b c : Nat} : ((a - b - c : Nat) : Int) = ((a - (b + c) : Nat) : Int) := congrArg _ (Nat.sub_sub _ _ _) +theorem ofNat_min (a b : Nat) : ((min a b : Nat) : Int) = min (a : Int) (b : Int) := by + simp only [Nat.min_def, Int.min_def, Int.ofNat_le] + split <;> rfl + +theorem ofNat_max (a b : Nat) : ((max a b : Nat) : Int) = max (a : Int) (b : Int) := by + simp only [Nat.max_def, Int.max_def, Int.ofNat_le] + split <;> rfl + +theorem ofNat_natAbs (a : Int) : (a.natAbs : Int) = if 0 ≤ a then a else -a := by + rw [Int.natAbs] + split <;> rename_i n + · simp only [Int.ofNat_eq_coe] + rw [if_pos (Int.ofNat_nonneg n)] + · simp; rfl + theorem natAbs_dichotomy {a : Int} : 0 ≤ a ∧ a.natAbs = a ∨ a < 0 ∧ a.natAbs = -a := by by_cases h : 0 ≤ a · left @@ -75,6 +104,9 @@ theorem natAbs_dichotomy {a : Int} : 0 ≤ a ∧ a.natAbs = a ∨ a < 0 ∧ a.na rw [Int.ofNat_natAbs_of_nonpos (Int.le_of_lt h)] simp_all +theorem neg_le_natAbs {a : Int} : -a ≤ a.natAbs := by + simpa using Int.le_natAbs (a := -a) + theorem add_le_iff_le_sub (a b c : Int) : a + b ≤ c ↔ a ≤ c - b := by conv => lhs @@ -116,7 +148,7 @@ theorem of_lex (w : Prod.Lex r s p q) : r p.fst q.fst ∨ p.fst = q.fst ∧ s p. theorem of_not_lex {α} {r : α → α → Prop} [DecidableEq α] {β} {s : β → β → Prop} {p q : α × β} (w : ¬ Prod.Lex r s p q) : ¬ r p.fst q.fst ∧ (p.fst ≠ q.fst ∨ ¬ s p.snd q.snd) := by - rw [Prod.lex_def, not_or, Decidable.not_and] at w + rw [Prod.lex_def, not_or, Decidable.not_and_iff_or_not_not] at w exact w theorem fst_mk : (Prod.mk x y).fst = x := rfl diff --git a/Std/Tactic/Omega/IntList.lean b/Std/Tactic/Omega/IntList.lean index c96044c62b..8aaaa1942e 100644 --- a/Std/Tactic/Omega/IntList.lean +++ b/Std/Tactic/Omega/IntList.lean @@ -3,10 +3,12 @@ Copyright (c) 2023 Lean FRO, LLC. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import Std.Data.List.Lemmas +import Std.Data.List.Init.Lemmas import Std.Data.Nat.Gcd -import Std.Data.Int.DivMod +import Std.Data.Int.Init.DivMod +import Std.Data.Option.Lemmas import Std.Tactic.Replace +import Std.Tactic.Simpa /-- A type synonym for `List Int`, used by `omega` for dense representation of coefficients. @@ -33,9 +35,9 @@ theorem get_of_length_le {xs : IntList} (h : xs.length ≤ i) : xs.get i = 0 := rw [get, List.get?_eq_none.mpr h] rfl -theorem lt_length_of_get_nonzero {xs : IntList} (h : xs.get i ≠ 0) : i < xs.length := by - revert h - simpa using mt get_of_length_le +-- theorem lt_length_of_get_nonzero {xs : IntList} (h : xs.get i ≠ 0) : i < xs.length := by +-- revert h +-- simpa using mt get_of_length_le /-- Like `List.set`, but right-pad with zeroes as necessary first. -/ def set (xs : IntList) (i : Nat) (y : Int) : IntList := @@ -50,20 +52,6 @@ def set (xs : IntList) (i : Nat) (y : Int) : IntList := @[simp] theorem set_cons_zero : set (x :: xs) 0 y = y :: xs := rfl @[simp] theorem set_cons_succ : set (x :: xs) (i+1) y = x :: set xs i y := rfl -theorem set_get_eq : get (set xs i y) j = if i = j then y else xs.get j := by - induction xs generalizing i j with - | nil => - induction i generalizing j with - | zero => cases j <;> simp [set] - | succ i => cases j <;> simp_all [set] - | cons x xs ih => - induction i with - | zero => cases j <;> simp [set] - | succ i => cases j <;> simp_all [set] - -@[simp] theorem set_get_self : get (set xs i y) i = y := by simp [set_get_eq] -@[simp] theorem set_get_of_ne (h : i ≠ j) : get (set xs i y) j = xs.get j := by simp [set_get_eq, h] - /-- Returns the leading coefficient, i.e. the first non-zero entry. -/ def leading (xs : IntList) : Int := xs.find? (! · == 0) |>.getD 0 @@ -146,15 +134,6 @@ theorem smul_def (xs : IntList) (i : Int) : i * xs = xs.map fun x => i * x := rf def combo (a : Int) (xs : IntList) (b : Int) (ys : IntList) : IntList := List.zipWithAll (fun x y => a * x.getD 0 + b * y.getD 0) xs ys -theorem combo_def (xs ys : IntList) : - combo a xs b ys = List.zipWithAll (fun x y => a * x.getD 0 + b * y.getD 0) xs ys := - rfl - -@[simp] theorem combo_get (xs ys : IntList) (i : Nat) : - (combo a xs b ys).get i = a * xs.get i + b * ys.get i := by - simp only [combo_def, get, List.zipWithAll_get?, List.get?_eq_none] - cases xs.get? i <;> cases ys.get? i <;> simp - theorem combo_eq_smul_add_smul (a : Int) (xs : IntList) (b : Int) (ys : IntList) : combo a xs b ys = a * xs + b * ys := by dsimp [combo] @@ -165,14 +144,6 @@ theorem combo_eq_smul_add_smul (a : Int) (xs : IntList) (b : Int) (ys : IntList) | nil => simp; rfl | cons y ys => simp_all -theorem mul_comm (xs ys : IntList) : xs * ys = ys * xs := by - induction xs generalizing ys with - | nil => simp - | cons x xs ih => cases ys <;> simp_all [Int.mul_comm] - -@[simp] theorem neg_neg {xs : IntList} : - - xs = xs := by - induction xs <;> simp_all - attribute [local simp] add_def mul_def in theorem mul_distrib_left (xs ys zs : IntList) : (xs + ys) * zs = xs * zs + ys * zs := by induction xs generalizing ys zs with @@ -208,9 +179,6 @@ theorem sub_eq_add_neg (xs ys : IntList) : xs - ys = xs + (-ys) := by | nil => simp | cons y ys => simp_all [Int.sub_eq_add_neg] -@[simp] theorem sub_get (xs ys : IntList) (i : Nat) : (xs - ys).get i = xs.get i - ys.get i := by - rw [sub_eq_add_neg, add_get, neg_get, ← Int.sub_eq_add_neg] - @[simp] theorem mul_smul_left {i : Int} {xs ys : IntList} : (i * xs) * ys = i * (xs * ys) := by induction xs generalizing ys with | nil => simp @@ -256,8 +224,8 @@ example : IntList.dot [a, b, c] [x, y, z] = IntList.dot [a, b, c] [x, y, z, w] : @[simp] theorem dot_nil_right : dot xs ([] : IntList) = 0 := by simp [dot] @[simp] theorem dot_cons₂ : dot (x::xs) (y::ys) = x * y + dot xs ys := rfl -theorem dot_comm (xs ys : IntList) : dot xs ys = dot ys xs := by - rw [dot, dot, mul_comm] +-- theorem dot_comm (xs ys : IntList) : dot xs ys = dot ys xs := by +-- rw [dot, dot, mul_comm] @[simp] theorem dot_set_left (xs ys : IntList) (i : Nat) (z : Int) : dot (xs.set i z) ys = dot xs ys + (z - xs.get i) * ys.get i := by @@ -279,10 +247,6 @@ theorem dot_comm (xs ys : IntList) : dot xs ys = dot ys xs := by | nil => simp | cons y ys => simp_all [Int.add_assoc] -@[simp] theorem dot_set_right (xs ys : IntList) (i : Nat) (z : Int) : - dot xs (ys.set i z) = dot xs ys + xs.get i * (z - ys.get i) := by - rw [dot_comm, dot_set_left, dot_comm, Int.mul_comm] - theorem dot_distrib_left (xs ys zs : IntList) : (xs + ys).dot zs = xs.dot zs + ys.dot zs := by simp [dot, mul_distrib_left, sum_add] @@ -292,9 +256,6 @@ theorem dot_distrib_left (xs ys zs : IntList) : (xs + ys).dot zs = xs.dot zs + y @[simp] theorem dot_smul_left (xs ys : IntList) (i : Int) : (i * xs).dot ys = i * xs.dot ys := by simp [dot] -theorem dot_sub_left (xs ys zs : IntList) : (xs - ys).dot zs = xs.dot zs - ys.dot zs := by - rw [sub_eq_add_neg, dot_distrib_left, dot_neg_left, ← Int.sub_eq_add_neg] - theorem dot_of_left_zero (w : ∀ x, x ∈ xs → x = 0) : dot xs ys = 0 := by induction xs generalizing ys with | nil => simp @@ -308,38 +269,12 @@ theorem dot_of_left_zero (w : ∀ x, x ∈ xs → x = 0) : dot xs ys = 0 := by apply w exact List.mem_cons_of_mem _ m -theorem dvd_dot_of_dvd_left (w : ∀ x, x ∈ xs → m ∣ x) : m ∣ dot xs ys := by - induction xs generalizing ys with - | nil => exact Int.dvd_zero m - | cons x xs ih => - cases ys with - | nil => exact Int.dvd_zero m - | cons y ys => - rw [dot_cons₂] - apply Int.dvd_add - · apply Int.dvd_trans - rotate_left - · apply Int.dvd_mul_right - · apply w - apply List.mem_cons_self - · apply ih - intro x m - apply w x - apply List.mem_cons_of_mem _ m - /-- Division of an `IntList` by a integer. -/ def sdiv (xs : IntList) (g : Int) : IntList := xs.map fun x => x / g @[simp] theorem sdiv_nil : sdiv [] g = [] := rfl @[simp] theorem sdiv_cons : sdiv (x::xs) g = (x / g) :: sdiv xs g := rfl -@[simp] theorem sdiv_get {xs : IntList} {g : Int} {i} : (xs.sdiv g).get i = xs.get i / g := by - simp only [sdiv, get, List.get?_map] - cases xs.get? i <;> simp - -theorem mem_sdiv {xs : IntList} (h : x ∈ xs) : x / g ∈ xs.sdiv g := by - apply List.mem_map_of_mem _ h - /-- The gcd of the absolute values of the entries of an `IntList`. -/ def gcd (xs : IntList) : Nat := xs.foldr (fun x g => Nat.gcd x.natAbs g) 0 @@ -428,10 +363,6 @@ theorem dot_eq_zero_of_left_eq_zero {xs ys : IntList} (h : ∀ x ∈ xs, x = 0) rw [dot_cons₂, h x (List.mem_cons_self _ _), ih (fun x m => h x (List.mem_cons_of_mem _ m)), Int.zero_mul, Int.add_zero] -theorem dot_eq_zero_of_gcd_left_eq_zero {xs ys : IntList} (h : xs.gcd = 0) : dot xs ys = 0 := by - simp at h - simp_all - theorem dot_sdiv_left (xs ys : IntList) {d : Int} (h : d ∣ xs.gcd) : dot (xs.sdiv d) ys = (dot xs ys) / d := by induction xs generalizing ys with @@ -445,88 +376,6 @@ theorem dot_sdiv_left (xs ys : IntList) {d : Int} (h : d ∣ xs.gcd) : have w : d ∣ (IntList.gcd xs : Int) := Int.dvd_trans h (gcd_cons_div_right') simp_all [Int.add_ediv_of_dvd_left, Int.mul_ediv_assoc'] -@[simp] theorem dot_sdiv_gcd_left (xs ys : IntList) : - dot (xs.sdiv xs.gcd) ys = (dot xs ys) / xs.gcd := - dot_sdiv_left xs ys (by exact Int.dvd_refl _) - -/-- The leading sign in an `IntList`. -/ -def leadingSign (xs : IntList) : Int := - match xs with - | [] => 0 - | 0 :: xs => leadingSign xs - | x :: _ => x.sign - -@[simp] theorem leadingSign_nil : leadingSign [] = 0 := rfl -@[simp] theorem leadingSign_cons_zero : leadingSign (0 :: xs) = leadingSign xs := rfl -theorem leadingSign_cons : leadingSign (x :: xs) = if x = 0 then leadingSign xs else x.sign := by - split <;> rename_i h - · subst h - rfl - · rw [leadingSign] - intro w - exact h w - -theorem leadingSign_neg {xs : IntList} : (-xs).leadingSign = - xs.leadingSign := by - induction xs with - | nil => simp - | cons x xs ih => - by_cases h : x = 0 - · subst h - simp_all - · simp_all [leadingSign_cons] - -/-- -Trim trailing zeroes from a `List Int`, returning `none` if none were removed. --/ --- We could implement this without any `List.reverse`s, but at the expense of either: --- * an `Array.toList` and scanning for zeroes from the left, or --- * a `List.toArray` and an `Array.toList`. --- It's unclear either would be appreciably faster. -def trim? (xs : IntList) : Option IntList := - match xs.reverse with - | [] => none - | 0 :: xs => (xs.dropWhile (· == 0)).reverse - | _ :: _ => none - -theorem trim?_isSome {xs : IntList} : xs.trim?.isSome = (xs.getLast? = some 0) := by - dsimp [trim?] - split <;> rename_i h h' - · simp_all - · replace h' := congrArg List.reverse h' - simp at h' - simp [h'] - · replace h' := congrArg List.reverse h' - simp at h' - simpa [h'] using h - -/-- Trailing trailing zeroes from a `List Int`. -/ -def trim (xs : IntList) : IntList := xs.trim?.getD xs - -theorem trim?_eq_some {xs : IntList} (w : xs.trim? = some t) : xs.trim = t := by simp [trim, w] - -theorem trim_spec {xs : IntList} : xs.trim = (xs.reverse.dropWhile (· == 0)).reverse := by - dsimp [trim, trim?] - split <;> rename_i h h' - · simp_all - · simp_all [List.dropWhile_cons] - · simp_all [List.dropWhile_cons] - rw [if_neg h, ← h', List.reverse_reverse] - -@[simp] theorem trim_nil : trim [] = [] := rfl - -@[simp] theorem trim_append_zero {xs : IntList} : (xs ++ [0]).trim = xs.trim := by - simp [trim_spec, List.dropWhile] - -@[simp] theorem trim_neg {xs : IntList} : (-xs).trim = -xs.trim := by - simp only [trim_spec, neg_def, List.reverse_map] - generalize xs.reverse = xs' - induction xs' with - | nil => simp - | cons x xs' ih => - simp only [List.map_cons, List.dropWhile_cons, Int.neg_eq_zero, beq_iff_eq] - split <;> - simp_all [List.reverse_map] - /-- Apply "balanced mod" to each entry in an `IntList`. -/ abbrev bmod (x : IntList) (m : Nat) : IntList := x.map (Int.bmod · m) diff --git a/Std/Tactic/Omega/LinearCombo.lean b/Std/Tactic/Omega/LinearCombo.lean index 9f71deb9af..379dee9db5 100644 --- a/Std/Tactic/Omega/LinearCombo.lean +++ b/Std/Tactic/Omega/LinearCombo.lean @@ -23,7 +23,7 @@ structure LinearCombo where /-- Constant term. -/ const : Int := 0 /-- Coefficients of the atoms. -/ - coeffs : Coeffs := {} + coeffs : Coeffs := [] deriving DecidableEq, Repr namespace LinearCombo @@ -40,7 +40,7 @@ instance : ToExpr LinearCombo where instance : Inhabited LinearCombo := ⟨{const := 1}⟩ -@[ext] theorem ext {a b : LinearCombo} (w₁ : a.const = b.const) (w₂ : a.coeffs = b.coeffs) : +theorem ext {a b : LinearCombo} (w₁ : a.const = b.const) (w₂ : a.coeffs = b.coeffs) : a = b := by cases a; cases b subst w₁; subst w₂ @@ -118,7 +118,7 @@ instance : Neg LinearCombo := ⟨neg⟩ theorem sub_eq_add_neg (l₁ l₂ : LinearCombo) : l₁ - l₂ = l₁ + -l₂ := by rcases l₁ with ⟨a₁, c₁⟩; rcases l₂ with ⟨a₂, c₂⟩ - ext1 + apply ext · simp [Int.sub_eq_add_neg] · simp [Coeffs.sub_eq_add_neg] diff --git a/Std/Tactic/Omega/Logic.lean b/Std/Tactic/Omega/Logic.lean index 018fd30fcb..7f699a7731 100644 --- a/Std/Tactic/Omega/Logic.lean +++ b/Std/Tactic/Omega/Logic.lean @@ -29,4 +29,4 @@ theorem Decidable.not_iff_iff_and_not_or_not_and [Decidable a] [Decidable b] : alias ⟨Decidable.and_not_or_not_and_of_not_iff, _⟩ := Decidable.not_iff_iff_and_not_or_not_and -alias ⟨Decidable.and_not_of_not_imp, _⟩ := Decidable.not_imp +alias ⟨Decidable.and_not_of_not_imp, _⟩ := Decidable.not_imp_iff_and_not diff --git a/Std/Tactic/Omega/MinNatAbs.lean b/Std/Tactic/Omega/MinNatAbs.lean index 9cf1944f7f..c5725be026 100644 --- a/Std/Tactic/Omega/MinNatAbs.lean +++ b/Std/Tactic/Omega/MinNatAbs.lean @@ -3,8 +3,9 @@ Copyright (c) 2023 Lean FRO, LLC. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import Std.Data.List.Lemmas -import Std.Data.Int.Lemmas +import Std.Data.List.Init.Lemmas +import Std.Data.Int.Init.Order +import Std.Data.Option.Lemmas import Std.Tactic.LeftRight /-! @@ -17,7 +18,6 @@ so we keep it in the `Std/Tactic/Omega` directory -/ -open Classical namespace List @@ -29,6 +29,7 @@ We completely characterize the function via -/ def nonzeroMinimum (xs : List Nat) : Nat := xs.filter (· ≠ 0) |>.minimum? |>.getD 0 +open Classical in @[simp] theorem nonzeroMinimum_eq_zero_iff {xs : List Nat} : xs.nonzeroMinimum = 0 ↔ ∀ x ∈ xs, x = 0 := by simp [nonzeroMinimum, Option.getD_eq_iff, minimum?_eq_none_iff, minimum?_eq_some_iff', @@ -84,13 +85,14 @@ theorem nonzeroMinimum_eq_of_nonzero {xs : List Nat} (h : xs.nonzeroMinimum ≠ theorem nonzeroMinimum_le_iff {xs : List Nat} {y : Nat} : xs.nonzeroMinimum ≤ y ↔ xs.nonzeroMinimum = 0 ∨ ∃ x ∈ xs, x ≤ y ∧ x ≠ 0 := by refine ⟨fun h => ?_, fun h => ?_⟩ - · rw [Decidable.or_iff_not_imp_right] - simp only [ne_eq, not_exists, not_and, not_not, nonzeroMinimum_eq_zero_iff] + · rw [Classical.or_iff_not_imp_right] + simp only [ne_eq, not_exists, not_and, Classical.not_not, nonzeroMinimum_eq_zero_iff] intro w apply nonzeroMinimum_eq_zero_iff.mp - by_cases p : xs.nonzeroMinimum = 0 - · exact p - · exact w _ (nonzeroMinimum_mem p) h + if p : xs.nonzeroMinimum = 0 then + exact p + else + exact w _ (nonzeroMinimum_mem p) h · match h with | .inl h => simp [h] | .inr ⟨x, m, le, ne⟩ => exact Nat.le_trans (nonzeroMinimum_le m ne) le diff --git a/Std/Tactic/Omega/OmegaM.lean b/Std/Tactic/Omega/OmegaM.lean index 59c6edc72b..0be6dc0383 100644 --- a/Std/Tactic/Omega/OmegaM.lean +++ b/Std/Tactic/Omega/OmegaM.lean @@ -27,7 +27,10 @@ The main functions are: `k * a ≤ x < k * a + k` * for each new atom of the form `((a - b : Nat) : Int)`, the fact: `b ≤ a ∧ ((a - b : Nat) : Int) = a - b ∨ a < b ∧ ((a - b : Nat) : Int) = 0` - + * for each new atom of the form `min a b`, the facts `min a b ≤ a` and `min a b ≤ b` + (and similarly for `max`) + * for each new atom of the form `if P then a else b`, the disjunction: + `(P ∧ (if P then a else b) = a) ∨ (¬ P ∧ (if P then a else b) = b)` The `OmegaM` monad also keeps an internal cache of visited expressions (not necessarily atoms, but arbitrary subexpressions of one side of a linear relation) to reduce duplication. @@ -81,16 +84,21 @@ def atomsList : OmegaM Expr := do mkListLit (.const ``Int []) (← atoms) def atomsCoeffs : OmegaM Expr := do return .app (.const ``Coeffs.ofList []) (← atomsList) +/-- Run an `OmegaM` computation, restoring the state afterwards depending on the result. -/ +def commitWhen (t : OmegaM (α × Bool)) : OmegaM α := do + let state ← getThe State + let cache ← getThe Cache + let (a, r) ← t + if !r then do + modifyThe State fun _ => state + modifyThe Cache fun _ => cache + pure a + /-- Run an `OmegaM` computation, restoring the state afterwards. -/ -def savingState (t : OmegaM α) : OmegaM α := do - let state ← getThe State - let cache ← getThe Cache - let r ← t - modifyThe State fun _ => state - modifyThe Cache fun _ => cache - pure r +def withoutModifyingState (t : OmegaM α) : OmegaM α := + commitWhen (do pure (← t, false)) /-- Wrapper around `Expr.nat?` that also allows `Nat.cast`. -/ def natCast? (n : Expr) : Option Nat := @@ -104,35 +112,83 @@ def intCast? (n : Expr) : Option Int := | (``Nat.cast, #[_, _, n]) => n.nat? | _ => n.int? +theorem ite_disjunction {α : Type u} {P : Prop} [Decidable P] {a b : α} : + (P ∧ (if P then a else b) = a) ∨ (¬ P ∧ (if P then a else b) = b) := by + by_cases P <;> simp_all + +/-- Construct the term with type hint `(Eq.refl a : a = b)`-/ +def mkEqReflWithExpectedType (a b : Expr) : MetaM Expr := do + mkExpectedTypeHint (← mkEqRefl a) (← mkEq a b) + /-- Analyzes a newly recorded atom, returning a collection of interesting facts about it that should be added to the context. -/ def analyzeAtom (e : Expr) : OmegaM (HashSet Expr) := do match e.getAppFnArgs with - | (``Nat.cast, #[_, _, e']) => + | (``Nat.cast, #[.const ``Int [], _, e']) => -- Casts of natural numbers are non-negative. let mut r := {Expr.app (.const ``Int.ofNat_nonneg []) e'} - match (← cfg).splitNatSub, (← cfg).splitNatAbs, e'.getAppFnArgs with - | true, _, (``HSub.hSub, #[_, _, _, _, a, b]) => + match (← cfg).splitNatSub, e'.getAppFnArgs with + | true, (``HSub.hSub, #[_, _, _, _, a, b]) => -- `((a - b : Nat) : Int)` gives a dichotomy r := r.insert (mkApp2 (.const ``Int.ofNat_sub_dichotomy []) a b) - | _, true, (``Int.natAbs, #[x]) => - -- `(a.natAbs : Int)` gives a dichotomy - r := r.insert (mkApp (.const ``Int.natAbs_dichotomy []) x) - | _, _,_ => pure () + | _, (``Int.natAbs, #[x]) => + r := r.insert (mkApp (.const ``Int.le_natAbs []) x) + r := r.insert (mkApp (.const ``Int.neg_le_natAbs []) x) + | _, (``Fin.val, #[n, i]) => + r := r.insert (mkApp2 (.const ``Fin.isLt []) n i) + | _, _ => pure () return r | (``HDiv.hDiv, #[_, _, _, _, x, k]) => match natCast? k with | none | some 0 => pure ∅ | some _ => -- `k * x/k ≤ x < k * x/k + k` - let ne_zero := mkApp3 (.const ``Ne [.succ .zero]) (.const ``Int []) k (toExpr (0 : Int)) - let pos := mkApp4 (.const ``LT.lt [.zero]) (.const ``Int []) (.const ``Int.instLTInt []) + let ne_zero := mkApp3 (.const ``Ne [1]) (.const ``Int []) k (toExpr (0 : Int)) + let pos := mkApp4 (.const ``LT.lt [0]) (.const ``Int []) (.const ``Int.instLTInt []) (toExpr (0 : Int)) k pure <| {mkApp3 (.const ``Int.mul_ediv_self_le []) x k (← mkDecideProof ne_zero), mkApp3 (.const ``Int.lt_mul_ediv_self_add []) x k (← mkDecideProof pos)} + | (``HMod.hMod, #[_, _, _, _, x, k]) => + match k.getAppFnArgs with + | (``HPow.hPow, #[_, _, _, _, b, exp]) => match natCast? b with + | none + | some 0 => pure ∅ + | some _ => + let b_pos := mkApp4 (.const ``LT.lt [0]) (.const ``Int []) (.const ``Int.instLTInt []) + (toExpr (0 : Int)) b + let pow_pos := mkApp3 (.const ``Int.pos_pow_of_pos []) b exp (← mkDecideProof b_pos) + pure <| + {mkApp3 (.const ``Int.emod_nonneg []) x k + (mkApp3 (.const ``Int.ne_of_gt []) k (toExpr (0 : Int)) pow_pos), + mkApp3 (.const ``Int.emod_lt_of_pos []) x k pow_pos} + | (``Nat.cast, #[.const ``Int [], _, k']) => + match k'.getAppFnArgs with + | (``HPow.hPow, #[_, _, _, _, b, exp]) => match natCast? b with + | none + | some 0 => pure ∅ + | some _ => + let b_pos := mkApp4 (.const ``LT.lt [0]) (.const ``Nat []) (.const ``instLTNat []) + (toExpr (0 : Nat)) b + let pow_pos := mkApp3 (.const ``Nat.pos_pow_of_pos []) b exp (← mkDecideProof b_pos) + let cast_pos := mkApp2 (.const ``Int.ofNat_pos_of_pos []) k' pow_pos + pure <| + {mkApp3 (.const ``Int.emod_nonneg []) x k + (mkApp3 (.const ``Int.ne_of_gt []) k (toExpr (0 : Int)) cast_pos), + mkApp3 (.const ``Int.emod_lt_of_pos []) x k cast_pos} + | _ => pure ∅ + | _ => pure ∅ + | (``Min.min, #[_, _, x, y]) => + pure <| {mkApp2 (.const ``Int.min_le_left []) x y, mkApp2 (.const ``Int.min_le_right []) x y} + | (``Max.max, #[_, _, x, y]) => + pure <| {mkApp2 (.const ``Int.le_max_left []) x y, mkApp2 (.const ``Int.le_max_right []) x y} + | (``ite, #[α, i, dec, t, e]) => + if α == (.const ``Int []) then + pure <| {mkApp5 (.const ``ite_disjunction [0]) α i dec t e} + else + pure {} | _ => pure ∅ /-- @@ -148,7 +204,6 @@ Return its index, and, if it is new, a collection of interesting facts about the def lookup (e : Expr) : OmegaM (Nat × Option (HashSet Expr)) := do let c ← getThe State for h : i in [:c.atoms.size] do - have : i < c.atoms.size := h.2 if ← isDefEq e c.atoms[i] then return (i, none) trace[omega] "New atom: {e}" diff --git a/Std/Tactic/RCases.lean b/Std/Tactic/RCases.lean index 37157ad50c..68c8cf75cb 100644 --- a/Std/Tactic/RCases.lean +++ b/Std/Tactic/RCases.lean @@ -310,7 +310,7 @@ def processConstructor (ref : Syntax) (info : Array ParamInfo) | [p] => ([p.name?.getD `_], [p]) | ps => ([`_], [(bif explicit then .explicit ref else id) (.tuple ref ps)]) else ([], []) -termination_by _ => info.size - idx +termination_by info.size - idx /-- Takes a list of constructor names, and an (alternation) list of patterns, and matches each @@ -487,10 +487,15 @@ The terminating continuation used in `rcasesCore` and `rcasesContinue`. We speci `α` to `Array MVarId` to collect the list of goals, and given the list of `clears`, it attempts to clear them from the goal and adds the goal to the list. -/ -def finish (g : MVarId) (fs : FVarSubst) (clears : Array FVarId) +def finish (toTag : Array (Ident × FVarId) := #[]) + (g : MVarId) (fs : FVarSubst) (clears : Array FVarId) (gs : Array MVarId) : TermElabM (Array MVarId) := do let cs : Array Expr := (clears.map fs.get).filter Expr.isFVar - gs.push <$> tryClearMany' g (cs.map Expr.fvarId!) + let g ← tryClearMany' g (cs.map Expr.fvarId!) + g.withContext do + for (stx, fvar) in toTag do + Term.addLocalVarInfo stx (fs.get fvar) + return gs.push g open Elab @@ -511,7 +516,7 @@ partial def RCasesPatt.parse (stx : Syntax) : MetaM RCasesPatt := -- extracted from elabCasesTargets /-- Generalize all the arguments as specified in `args` to fvars if they aren't already -/ def generalizeExceptFVar (goal : MVarId) (args : Array GeneralizeArg) : - MetaM (Array Expr × MVarId) := do + MetaM (Array Expr × Array FVarId × MVarId) := do let argsToGeneralize := args.filter fun arg => !(arg.expr.isFVar && arg.hName?.isNone) let (fvarIdsNew, goal) ← goal.generalize argsToGeneralize let mut result := #[] @@ -522,14 +527,14 @@ def generalizeExceptFVar (goal : MVarId) (args : Array GeneralizeArg) : else result := result.push (mkFVar fvarIdsNew[j]!) j := j+1 - pure (result, goal) + pure (result, fvarIdsNew[j:], goal) /-- Given a list of targets of the form `e` or `h : e`, and a pattern, match all the targets against the pattern. Returns the list of produced subgoals. -/ -def rcases (tgts : Array (Option Name × Syntax)) - (pat : RCasesPatt) (g : MVarId) : TermElabM (List MVarId) := Term.withSynthesize do +def rcases (tgts : Array (Option Ident × Syntax)) + (pat : RCasesPatt) (g : MVarId) : TermElabM (List MVarId) := Term.withSynthesize do let pats ← match tgts.size with | 0 => return [g] | 1 => pure [pat] @@ -541,9 +546,11 @@ def rcases (tgts : Array (Option Name × Syntax)) pure (.typed ref pat (← Term.exprToSyntax ty), some ty) | _ => pure (pat, none) let expr ← Term.ensureHasType ty (← Term.elabTerm tgt ty) - pure (pat, { expr, xName? := pat.name?, hName? : GeneralizeArg }) - let (vs, g) ← generalizeExceptFVar g args - (·.toList) <$> rcasesContinue g {} #[] #[] (pats.zip vs).toList finish + pure (pat, { expr, xName? := pat.name?, hName? := hName?.map (·.getId) : GeneralizeArg }) + let (vs, hs, g) ← generalizeExceptFVar g args + let toTag := tgts.filterMap (·.1) |>.zip hs + let gs ← rcasesContinue g {} #[] #[] (pats.zip vs).toList (finish (toTag := toTag)) + pure gs.toList /-- The `obtain` tactic in the no-target case. Given a type `T`, create a goal `|- T` and @@ -688,7 +695,7 @@ elab (name := rcases) tk:"rcases" tgts:casesTarget,* pat:((" with " rcasesPatLo) | #[] => pure $ RCasesPatt.tuple tk [] | _ => throwUnsupportedSyntax let tgts := tgts.getElems.map fun tgt => - (if tgt.raw[0].isNone then none else some tgt.raw[0][0].getId, tgt.raw[1]) + (if tgt.raw[0].isNone then none else some ⟨tgt.raw[0][0]⟩, tgt.raw[1]) let g ← getMainGoal g.withContext do replaceMainGoal (← RCases.rcases tgts pat g) diff --git a/Std/Tactic/Relation/Rfl.lean b/Std/Tactic/Relation/Rfl.lean index 4675f3c6e0..fb14972946 100644 --- a/Std/Tactic/Relation/Rfl.lean +++ b/Std/Tactic/Relation/Rfl.lean @@ -24,7 +24,7 @@ def reflExt.config : WhnfCoreConfig := {} initialize reflExt : SimpleScopedEnvExtension (Name × Array DiscrTree.Key) (DiscrTree Name) ← registerSimpleScopedEnvExtension { - addEntry := fun dt (n, ks) => dt.insertCore ks n reflExt.config + addEntry := fun dt (n, ks) => dt.insertCore ks n initial := {} } diff --git a/Std/Tactic/Relation/Symm.lean b/Std/Tactic/Relation/Symm.lean index 6e09d318a3..2369a4efd6 100644 --- a/Std/Tactic/Relation/Symm.lean +++ b/Std/Tactic/Relation/Symm.lean @@ -27,7 +27,7 @@ def symmExt.config : WhnfCoreConfig := {} initialize symmExt : SimpleScopedEnvExtension (Name × Array DiscrTree.Key) (DiscrTree Name) ← registerSimpleScopedEnvExtension { - addEntry := fun dt (n, ks) => dt.insertCore ks n symmExt.config + addEntry := fun dt (n, ks) => dt.insertCore ks n initial := {} } diff --git a/Std/Tactic/SimpTrace.lean b/Std/Tactic/SimpTrace.lean index 1496fdea11..59f9e85425 100644 --- a/Std/Tactic/SimpTrace.lean +++ b/Std/Tactic/SimpTrace.lean @@ -36,44 +36,10 @@ syntax (name := simpTrace) "simp?" "!"? simpTraceArgsRest : tactic macro tk:"simp?!" rest:simpTraceArgsRest : tactic => `(tactic| simp?%$tk ! $rest) open TSyntax.Compat in --- adapted from traceSimpCall /-- Constructs the syntax for a simp call, for use with `simp?`. -/ def mkSimpCallStx (stx : Syntax) (usedSimps : UsedSimps) : MetaM (TSyntax `tactic) := do - let mut stx := stx.unsetTrailing - if stx[3].isNone then - stx := stx.setArg 3 (mkNullNode #[mkAtom "only"]) - let mut args := #[] - let mut localsOrStar := some #[] - let lctx ← getLCtx - let env ← getEnv - for (thm, _) in usedSimps.toArray.qsort (·.2 < ·.2) do - match thm with - | .decl declName inv => -- global definitions in the environment - if env.contains declName && !simpOnlyBuiltins.contains declName then - args := args.push (← if inv then - `(Parser.Tactic.simpLemma| ← $(mkIdent (← unresolveNameGlobal declName)):ident) - else - `(Parser.Tactic.simpLemma| $(mkIdent (← unresolveNameGlobal declName)):ident)) - | .fvar fvarId => -- local hypotheses in the context - if let some ldecl := lctx.find? fvarId then - localsOrStar := localsOrStar.bind fun locals => - if !ldecl.userName.isInaccessibleUserName && - (lctx.findFromUserName? ldecl.userName).get!.fvarId == ldecl.fvarId then - some (locals.push ldecl.userName) - else - none - -- Note: the `if let` can fail for `simp (config := {contextual := true})` when - -- rewriting with a variable that was introduced in a scope. In that case we just ignore. - | .stx _ thmStx => -- simp theorems provided in the local invocation - args := args.push thmStx - | .other _ => -- Ignore "special" simp lemmas such as constructed by `simp_all`. - pure () -- We can't display them anyway. - if let some locals := localsOrStar then - args := args ++ (← locals.mapM fun id => `(Parser.Tactic.simpLemma| $(mkIdent id):ident)) - else - args := args.push (← `(Parser.Tactic.simpStar| *)) - let argsStx := if args.isEmpty then #[] else #[mkAtom "[", (mkAtom ",").mkSep args, mkAtom "]"] - pure <| stx.setArg 4 (mkNullNode argsStx) + let stx := stx.unsetTrailing + mkSimpOnly stx usedSimps elab_rules : tactic | `(tactic| @@ -82,9 +48,11 @@ elab_rules : tactic `(tactic| simp!%$tk $(config)? $(discharger)? $[only%$o]? $[[$args,*]]? $(loc)?) else `(tactic| simp%$tk $(config)? $(discharger)? $[only%$o]? $[[$args,*]]? $(loc)?) - let { ctx, dischargeWrapper } ← withMainContext <| mkSimpContext stx (eraseLocal := false) + let { ctx, simprocs, dischargeWrapper } ← + withMainContext <| mkSimpContext stx (eraseLocal := false) let usedSimps ← dischargeWrapper.with fun discharge? => - simpLocation ctx discharge? <| (loc.map expandLocation).getD (.targets #[] true) + simpLocation ctx (simprocs := simprocs) discharge? <| + (loc.map expandLocation).getD (.targets #[] true) let stx ← mkSimpCallStx stx usedSimps TryThis.addSuggestion tk stx (origSpan? := ← getRef) diff --git a/Std/Tactic/Simpa.lean b/Std/Tactic/Simpa.lean index 0ae674f872..e2a6bb1fdd 100644 --- a/Std/Tactic/Simpa.lean +++ b/Std/Tactic/Simpa.lean @@ -53,59 +53,6 @@ syntax (name := simpa) "simpa" "?"? "!"? simpaArgsRest : tactic open private useImplicitLambda from Lean.Elab.Term --- FIXME: remove when lean4#1862 lands -open TSyntax.Compat in -/-- -If `stx` is the syntax of a `simp`, `simp_all` or `dsimp` tactic invocation, and -`usedSimps` is the set of simp lemmas used by this invocation, then `mkSimpOnly` -creates the syntax of an equivalent `simp only`, `simp_all only` or `dsimp only` -invocation. --/ -private def mkSimpOnly (stx : Syntax) (usedSimps : UsedSimps) : MetaM Syntax.Tactic := do - let isSimpAll := stx[0].getAtomVal == "simp_all" - let mut stx := stx - if stx[3].isNone then - stx := stx.setArg 3 (mkNullNode #[mkAtom "only"]) - let mut args := #[] - let mut localsOrStar := some #[] - let lctx ← getLCtx - let env ← getEnv - for (thm, _) in usedSimps.toArray.qsort (·.2 < ·.2) do - match thm with - | .decl declName inv => -- global definitions in the environment - if env.contains declName && !simpOnlyBuiltins.contains declName then - args := args.push (← if inv then - `(Parser.Tactic.simpLemma| ← $(mkIdent (← unresolveNameGlobal declName)):ident) - else - `(Parser.Tactic.simpLemma| $(mkIdent (← unresolveNameGlobal declName)):ident)) - | .fvar fvarId => -- local hypotheses in the context - if isSimpAll then - continue - -- `simp_all` uses all hypotheses anyway, so we do not need to include - -- them in the arguments. In fact, it would be harmful to do so: - -- `simp_all only [h]`, where `h` is a hypothesis, simplifies `h` to - -- `True` and subsequenly removes it from the context, whereas - -- `simp_all` does not. So to get behavior equivalent to `simp_all`, we - -- must omit `h`. - if let some ldecl := lctx.find? fvarId then - localsOrStar := localsOrStar.bind fun locals => - if !ldecl.userName.isInaccessibleUserName && - (lctx.findFromUserName? ldecl.userName).get!.fvarId == ldecl.fvarId then - some (locals.push ldecl.userName) - else - none - -- Note: the `if let` can fail for `simp (config := {contextual := true})` when - -- rewriting with a variable that was introduced in a scope. In that case we just ignore. - | .stx _ thmStx => -- simp theorems provided in the local invocation - args := args.push thmStx - | .other _ => -- Ignore "special" simp lemmas such as constructed by `simp_all`. - pure () -- We can't display them anyway. - if let some locals := localsOrStar then - args := args ++ (← locals.mapM fun id => `(Parser.Tactic.simpLemma| $(mkIdent id):ident)) - else - args := args.push (← `(Parser.Tactic.simpStar| *)) - let argsStx := if args.isEmpty then #[] else #[mkAtom "[", (mkAtom ",").mkSep args, mkAtom "]"] - return stx.setArg 4 (mkNullNode argsStx) /-- Gets the value of the `linter.unnecessarySimpa` option. -/ def getLinterUnnecessarySimpa (o : Options) : Bool := @@ -117,13 +64,14 @@ elab_rules : tactic | `(tactic| simpa%$tk $[?%$squeeze]? $[!%$unfold]? $(cfg)? $(disch)? $[only%$only]? $[[$args,*]]? $[using $usingArg]?) => Elab.Tactic.focus do let stx ← `(tactic| simp $(cfg)? $(disch)? $[only%$only]? $[[$args,*]]?) - let { ctx, dischargeWrapper } ← withMainContext <| mkSimpContext stx (eraseLocal := false) + let { ctx, simprocs, dischargeWrapper } ← + withMainContext <| mkSimpContext stx (eraseLocal := false) let ctx := if unfold.isSome then { ctx with config.autoUnfold := true } else ctx -- TODO: have `simpa` fail if it doesn't use `simp`. let ctx := { ctx with config := { ctx.config with failIfUnchanged := false } } dischargeWrapper.with fun discharge? => do - let (some (_, g), usedSimps) ← - simpGoal (← getMainGoal) ctx (simplifyTarget := true) (discharge? := discharge?) + let (some (_, g), usedSimps) ← simpGoal (← getMainGoal) ctx (simprocs := simprocs) + (simplifyTarget := true) (discharge? := discharge?) | if getLinterUnnecessarySimpa (← getOptions) then logLint linter.unnecessarySimpa (← getRef) "try 'simp' instead of 'simpa'" g.withContext do @@ -135,7 +83,7 @@ elab_rules : tactic pure (h, g) else (← g.assert `h (← inferType e) e).intro1 - let (result?, usedSimps) ← simpGoal g ctx (fvarIdsToSimp := #[h]) + let (result?, usedSimps) ← simpGoal g ctx (simprocs := simprocs) (fvarIdsToSimp := #[h]) (simplifyTarget := false) (usedSimps := usedSimps) (discharge? := discharge?) match result? with | some (xs, g) => @@ -151,8 +99,9 @@ elab_rules : tactic m!"try 'simp at {Expr.fvar h}' instead of 'simpa using {Expr.fvar h}'" pure usedSimps else if let some ldecl := (← getLCtx).findFromUserName? `this then - if let (some (_, g), usedSimps) ← simpGoal g ctx (fvarIdsToSimp := #[ldecl.fvarId]) - (simplifyTarget := false) (usedSimps := usedSimps) (discharge? := discharge?) then + if let (some (_, g), usedSimps) ← simpGoal g ctx (simprocs := simprocs) + (fvarIdsToSimp := #[ldecl.fvarId]) (simplifyTarget := false) (usedSimps := usedSimps) + (discharge? := discharge?) then g.assumption; pure usedSimps else pure usedSimps diff --git a/Std/Tactic/SolveByElim.lean b/Std/Tactic/SolveByElim.lean index 773472abed..7983129a3b 100644 --- a/Std/Tactic/SolveByElim.lean +++ b/Std/Tactic/SolveByElim.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, David Renshaw -/ +import Lean.Elab.Tactic.Config import Std.Data.Sum.Basic import Std.Tactic.LabelAttr import Std.Tactic.Relation.Symm @@ -112,6 +113,10 @@ structure Config extends ApplyRulesConfig where /-- Enable backtracking search. -/ backtracking : Bool := true maxDepth := 6 + /-- Trying calling `intro` if no lemmas apply. -/ + intro : Bool := true + /-- Try calling `constructor` if no lemmas apply. -/ + constructor : Bool := true instance : Coe Config BacktrackConfig := ⟨(·.toApplyRulesConfig.toBacktrackConfig)⟩ @@ -178,6 +183,10 @@ def withDischarge (cfg : Config := {}) (discharge : MVarId → MetaM (Option (Li def introsAfter (cfg : Config := {}) : Config := cfg.withDischarge fun g => do pure [(← g.intro1P).2] +/-- Call `constructor` when no lemmas apply. -/ +def constructorAfter (cfg : Config := {}) : Config := + cfg.withDischarge fun g => g.constructor {newGoals := .all} + /-- Create or modify a `Config` which calls `synthInstance` on any goal for which no lemma applies. -/ def synthInstanceAfter (cfg : Config := {}) : Config := @@ -212,6 +221,14 @@ def requireUsingAll (cfg : Config := {}) (use : List Expr) : Config := cfg.testSolutions fun sols => do pure <| use.all fun e => sols.any fun s => e.occurs s +/-- +Process the `intro` and `constructor` options by implementing the `discharger` tactic. +-/ +def processOptions (cfg : Config) : Config := + let cfg := if cfg.intro then introsAfter { cfg with intro := false } else cfg + let cfg := if cfg.constructor then constructorAfter { cfg with constructor := false } else cfg + cfg + end Config /-- @@ -225,18 +242,12 @@ def elabContextLemmas (g : MVarId) (lemmas : List (TermElabM Expr)) (ctx : TermE /-- Returns the list of tactics corresponding to applying the available lemmas to the goal. -/ def applyLemmas (cfg : Config) (lemmas : List (TermElabM Expr)) (ctx : TermElabM (List Expr)) (g : MVarId) : Nondet MetaM (List MVarId) := Nondet.squash fun _ => do - -- We handle `cfg.symm` by saturating hypotheses of all goals using `symm`. - -- This has better performance that the mathlib3 approach. - let g ← if cfg.symm then g.symmSaturate else pure g let es ← elabContextLemmas g lemmas ctx return applyTactics cfg.toApplyConfig cfg.transparency es g /-- Applies the first possible lemma to the goal. -/ def applyFirstLemma (cfg : Config) (lemmas : List (TermElabM Expr)) (ctx : TermElabM (List Expr)) (g : MVarId) : MetaM (List MVarId) := do --- We handle `cfg.symm` by saturating hypotheses of all goals using `symm`. --- This has better performance that the mathlib3 approach. -let g ← if cfg.symm then g.symmSaturate else pure g let es ← elabContextLemmas g lemmas ctx applyFirst cfg.toApplyConfig cfg.transparency es g @@ -258,20 +269,27 @@ Custom wrappers (e.g. `apply_assumption` and `apply_rules`) may modify this beha -/ def solveByElim (cfg : Config) (lemmas : List (TermElabM Expr)) (ctx : TermElabM (List Expr)) (goals : List MVarId) : MetaM (List MVarId) := do + let cfg := cfg.processOptions + -- We handle `cfg.symm` by saturating hypotheses of all goals using `symm`. + -- This has better performance that the mathlib3 approach. + let preprocessedGoals ← if cfg.symm then + goals.mapM fun g => g.symmSaturate + else + pure goals try - run goals + run cfg preprocessedGoals catch e => do -- Implementation note: as with `cfg.symm`, this is different from the mathlib3 approach, -- for (not as severe) performance reasons. - match goals, cfg.exfalso with + match preprocessedGoals, cfg.exfalso with | [g], true => withTraceNode `Meta.Tactic.solveByElim (fun _ => return m!"⏮️ starting over using `exfalso`") do - run [← g.exfalso] + run cfg [← g.exfalso] | _, _ => throw e where /-- Run either backtracking search, or repeated application, on the list of goals. -/ - run : List MVarId → MetaM (List MVarId) := + run (cfg : Config) : List MVarId → MetaM (List MVarId) := if cfg.backtracking then backtrack cfg `Meta.Tactic.solveByElim (applyLemmas cfg lemmas ctx) else diff --git a/Std/Tactic/SqueezeScope.lean b/Std/Tactic/SqueezeScope.lean index d967f5b719..520fa5b2e8 100644 --- a/Std/Tactic/SqueezeScope.lean +++ b/Std/Tactic/SqueezeScope.lean @@ -95,13 +95,14 @@ elab_rules : tactic let stx := tac.raw let usedSimps ← match stx.getKind with | ``Parser.Tactic.simp => do - let { ctx, dischargeWrapper } ← withMainContext <| mkSimpContext stx (eraseLocal := false) + let { ctx, simprocs, dischargeWrapper } ← + withMainContext <| mkSimpContext stx (eraseLocal := false) dischargeWrapper.with fun discharge? => - simpLocation ctx discharge? (expandOptLocation stx[5]) + simpLocation ctx simprocs discharge? (expandOptLocation stx[5]) | ``Parser.Tactic.simpAll => do - let { ctx, .. } ← mkSimpContext stx + let { ctx, simprocs, .. } ← mkSimpContext stx (eraseLocal := true) (kind := .simpAll) (ignoreStarArg := true) - let (result?, usedSimps) ← simpAll (← getMainGoal) ctx + let (result?, usedSimps) ← simpAll (← getMainGoal) ctx simprocs match result? with | none => replaceMainGoal [] | some mvarId => replaceMainGoal [mvarId] diff --git a/Std/Tactic/TryThis.lean b/Std/Tactic/TryThis.lean index 76d5d9f32d..752a9e3d02 100644 --- a/Std/Tactic/TryThis.lean +++ b/Std/Tactic/TryThis.lean @@ -9,6 +9,7 @@ import Std.Lean.Name import Std.Lean.Format import Std.Lean.Position import Std.Data.Json +import Std.Lean.Syntax /-! # "Try this" support @@ -523,3 +524,61 @@ def addTermSuggestions (ref : Syntax) (es : Array Expr) (codeActionPrefix? : Option String := none) : MetaM Unit := do addSuggestions ref (← es.mapM delabToRefinableSuggestion) (origSpan? := origSpan?) (header := header) (codeActionPrefix? := codeActionPrefix?) + +open Lean Elab Elab.Tactic PrettyPrinter Meta Std.Tactic.TryThis + +/-- Add a suggestion for `have h : t := e`. -/ +def addHaveSuggestion (ref : Syntax) (h? : Option Name) (t? : Option Expr) (e : Expr) + (origSpan? : Option Syntax := none) : TermElabM Unit := do + let estx ← delabToRefinableSyntax e + let prop ← isProp (← inferType e) + let tac ← if let some t := t? then + let tstx ← delabToRefinableSyntax t + if prop then + match h? with + | some h => `(tactic| have $(mkIdent h) : $tstx := $estx) + | none => `(tactic| have : $tstx := $estx) + else + `(tactic| let $(mkIdent (h?.getD `_)) : $tstx := $estx) + else + if prop then + match h? with + | some h => `(tactic| have $(mkIdent h) := $estx) + | none => `(tactic| have := $estx) + else + `(tactic| let $(mkIdent (h?.getD `_)) := $estx) + addSuggestion ref tac origSpan? + +open Lean.Parser.Tactic +open Lean.Syntax + +/-- Add a suggestion for `rw [h₁, ← h₂] at loc`. -/ +def addRewriteSuggestion (ref : Syntax) (rules : List (Expr × Bool)) + (type? : Option Expr := none) (loc? : Option Expr := none) + (origSpan? : Option Syntax := none) : + TermElabM Unit := do + let rules_stx := TSepArray.ofElems <| ← rules.toArray.mapM fun ⟨e, symm⟩ => do + let t ← delabToRefinableSyntax e + if symm then `(rwRule| ← $t:term) else `(rwRule| $t:term) + let tac ← do + let loc ← loc?.mapM fun loc => do `(location| at $(← delab loc):term) + `(tactic| rw [$rules_stx,*] $(loc)?) + + -- We don't simply write `let mut tacMsg := m!"{tac}"` here + -- but instead rebuild it, so that there are embedded `Expr`s in the message, + -- thus giving more information in the hovers. + -- Perhaps in future we will have a better way to attach elaboration information to + -- `Syntax` embedded in a `MessageData`. + let mut tacMsg := + let rulesMsg := MessageData.sbracket <| MessageData.joinSep + (rules.map fun ⟨e, symm⟩ => (if symm then "← " else "") ++ m!"{e}") ", " + if let some loc := loc? then + m!"rw {rulesMsg} at {loc}" + else + m!"rw {rulesMsg}" + let mut extraMsg := "" + if let some type := type? then + tacMsg := tacMsg ++ m!"\n-- {type}" + extraMsg := extraMsg ++ s!"\n-- {← PrettyPrinter.ppExpr type}" + addSuggestion ref (s := { suggestion := tac, postInfo? := extraMsg, messageData? := tacMsg }) + origSpan? diff --git a/Std/Test/Internal/DummyLabelAttr.lean b/Std/Test/Internal/DummyLabelAttr.lean index 7203be2557..2104976c00 100644 --- a/Std/Test/Internal/DummyLabelAttr.lean +++ b/Std/Test/Internal/DummyLabelAttr.lean @@ -1,3 +1,8 @@ +/- +Copyright (c) 2023 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison +-/ import Std.Tactic.LabelAttr /-- A dummy label attribute, which can be used for testing. -/ diff --git a/Std/Util/Cache.lean b/Std/Util/Cache.lean index 5ec7abc82f..1aa9979b1f 100644 --- a/Std/Util/Cache.lean +++ b/Std/Util/Cache.lean @@ -146,7 +146,7 @@ def DiscrTreeCache.mk [BEq α] (profilingName : String) IO (DiscrTreeCache α) := let updateTree := fun name constInfo tree => do return (← processDecl name constInfo).foldl (init := tree) fun t (k, v) => - t.insertCore k v config + t.insertCore k v let addDecl := fun name constInfo (tree₁, tree₂) => return (← updateTree name constInfo tree₁, tree₂) let addLibraryDecl := fun name constInfo (tree₁, tree₂) => diff --git a/Std/Util/CheckTactic.lean b/Std/Util/CheckTactic.lean new file mode 100644 index 0000000000..76cfedf473 --- /dev/null +++ b/Std/Util/CheckTactic.lean @@ -0,0 +1,77 @@ +/- +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 + +/- +This file is the home for commands to tactics behave as expected. + +It currently includes two tactixs: + +#check_tactic t ~> res + + +only a #check_simp command that applies simp +IT +-/ + +namespace Std.Tactic + +open Lean +open 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 + +/-- +`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 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}" + 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 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 `try simp` reduces `t` to `r`. +-/ +macro "#check_simp " t:term "~>" exp:term : command => + `(command|#check_tactic $t ~> $exp by try simp) diff --git a/Std/WF.lean b/Std/WF.lean index 010b9072e6..d672a79d00 100644 --- a/Std/WF.lean +++ b/Std/WF.lean @@ -51,7 +51,7 @@ instance wfRel {r : α → α → Prop} : WellFoundedRelation { val // Acc r val ((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)) -termination_by _ a h => Subtype.mk a h +termination_by Subtype.mk a t private theorem recC_intro {motive : (a : α) → Acc r a → Sort v} (intro : (x : α) → (h : ∀ (y : α), r y x → Acc r y) → @@ -97,7 +97,7 @@ than the one inferred by default: ``` def otherWF : WellFounded Nat := … def foo (n : Nat) := … -termination_by foo n => otherWF.wrap n +termination_by otherWF.wrap n ``` -/ def wrap {α : Sort u} {r : α → α → Prop} (h : WellFounded r) (x : α) : {x : α // Acc r x} := @@ -118,7 +118,7 @@ Workaround until Lean has native support for this. -/ @[specialize] private def fixC {α : Sort u} {C : α → Sort v} {r : α → α → Prop} (hwf : WellFounded r) (F : ∀ x, (∀ y, r y x → C y) → C x) (x : α) : C x := F x (fun y _ => fixC hwf F y) -termination_by _ x => hwf.wrap x +termination_by hwf.wrap x @[csimp] private theorem fix_eq_fixC : @fix = @fixC := rfl diff --git a/lake-manifest.json b/lake-manifest.json new file mode 100644 index 0000000000..00280abb6f --- /dev/null +++ b/lake-manifest.json @@ -0,0 +1,50 @@ +{"version": 7, + "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"}], + "name": "std", + "lakeDir": ".lake"} diff --git a/lean-toolchain b/lean-toolchain index 3f21e50bd4..cfcdd3277d 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.5.0-rc1 +leanprover/lean4:v4.6.0-rc1 diff --git a/test/array.lean b/test/array.lean new file mode 100644 index 0000000000..487ac93581 --- /dev/null +++ b/test/array.lean @@ -0,0 +1,21 @@ +import Std.Util.CheckTactic +import Std.Data.Array + +section +variable {α : Type _} +variable [Inhabited α] +variable (a : Array α) +variable (i : Nat) +variable (v d : α) +variable (g : i < (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).getD i d ~> if i < a.size then v else d +#check_simp (a.set! i v)[i] ~> v + +-- This doesn't work currently. +-- It will be address in the comprehensive overhaul of array lemmas. +-- #check_simp (a.set! i v)[i]? ~> .some v + +end diff --git a/test/classical.lean b/test/classical.lean new file mode 100644 index 0000000000..a9c527eb75 --- /dev/null +++ b/test/classical.lean @@ -0,0 +1,32 @@ +import Std.Tactic.Classical +import Std.Tactic.PermuteGoals +import Std.Tactic.GuardExpr + +noncomputable example : Bool := by + fail_if_success have := ∀ p, decide p -- no classical in scope + classical! + have := ∀ p, decide p -- uses the classical instance + -- uses the classical instance even though `0 < 1` is decidable + guard_expr decide (0 < 1) = @decide (0 < 1) (‹(a : Prop) → Decidable a› _) + exact decide (0 < 1) + +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) -- uses 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) -- uses the decidable instance diff --git a/test/decidability.lean b/test/decidability.lean new file mode 100644 index 0000000000..4b14856db7 --- /dev/null +++ b/test/decidability.lean @@ -0,0 +1,7 @@ +import Std.Data.Nat.Lemmas + +-- Prior to leanprover/lean4#2552 there was a performance trap +-- depending on the implementation details in `decidableBallLT`. +-- We keep this example (which would have gone over maxHeartbeats) +-- as a regression test for the instance. +example : ∀ m, m < 25 → ∀ n, n < 25 → ∀ c, c < 25 → m ^ 2 + n ^ 2 + c ^ 2 ≠ 7 := by decide diff --git a/test/library_search/basic.lean b/test/library_search/basic.lean index 138e4d2086..16ba656dd9 100644 --- a/test/library_search/basic.lean +++ b/test/library_search/basic.lean @@ -6,13 +6,25 @@ set_option autoImplicit true -- And this option to trace all candidate lemmas before application. -- set_option trace.Tactic.stdLibrarySearch.lemmas true +-- Many of the tests here are quite volatile, +-- and when changes are made to `solve_by_elim` or `exact?`, +-- or the library itself, the printed messages change. +-- Hence many of the tests here use `#guard_msgs (drop info)`, +-- and do not actually verify the particular output, just that `exact?` succeeds. +-- We keep the most recent output as a comment +-- (not a doc-comment: so `#guard_msgs` doesn't check it) +-- for reference. +-- If you find further tests failing please: +-- 1. update the comment using the code action on `#guard_msgs` +-- 2. (optional) add `(drop info)` after `#guard_msgs` and change the doc-comment to a comment + noncomputable section -/-- info: Try this: exact Nat.lt.base x -/ +/-- info: Try this: exact Nat.le.refl -/ #guard_msgs in example (x : Nat) : x ≠ x.succ := Nat.ne_of_lt (by std_apply?) -/-- info: Try this: exact Nat.zero_lt_succ 1 -/ +/-- info: Try this: exact Nat.le.step Nat.le.refl -/ #guard_msgs in example : 0 ≠ 1 + 1 := Nat.ne_of_lt (by std_apply?) @@ -22,7 +34,7 @@ example : 0 ≠ 1 + 1 := Nat.ne_of_lt (by exact Fin.size_pos') #guard_msgs in example (x y : Nat) : x + y = y + x := by std_apply? -/-- info: Try this: exact fun a => Nat.add_le_add_right a k -/ +/-- info: Try this: exact fun a => Nat.add_le_add a Nat.le.refl -/ #guard_msgs in example (n m k : Nat) : n ≤ m → n + k ≤ m + k := by std_apply? @@ -34,7 +46,7 @@ example (ha : a > 0) (w : b ∣ c) : a * b ∣ a * c := by std_apply? #guard_msgs (drop info) in example : Int := by std_apply? -/-- info: Try this: Nat.lt.base x -/ +/-- info: Try this: Nat.le.refl -/ #guard_msgs in example : x < x + 1 := std_exact?% @@ -144,7 +156,7 @@ end synonym example : ∀ P : Prop, ¬(P ↔ ¬P) := by std_apply? -- We even find `iff` results: -/-- info: Try this: exact Iff.mpr (Nat.dvd_add_iff_left h₁) h₂ -/ +/-- info: Try this: exact (Nat.dvd_add_iff_left h₁).mpr h₂ -/ #guard_msgs in example {a b c : Nat} (h₁ : a ∣ c) (h₂ : a ∣ b + c) : a ∣ b := by std_apply? @@ -156,11 +168,11 @@ example {a b c : Nat} (h₁ : a ∣ c) (h₂ : a ∣ b + c) : a ∣ b := by std_ opaque f : Nat → Nat axiom F (a b : Nat) : f a ≤ f b ↔ a ≤ b -/-- info: Try this: exact Iff.mpr (F a b) h -/ +/-- info: Try this: exact (F a b).mpr h -/ #guard_msgs in example (a b : Nat) (h : a ≤ b) : f a ≤ f b := by std_apply? -/-- info: Try this: exact List.join L -/ +/-- info: Try this: exact List.findIdxs (fun a => false) L -/ #guard_msgs in example (L : List (List Nat)) : List Nat := by std_apply? using L @@ -222,3 +234,19 @@ theorem Bool_eq_iff2 {A B : Bool} : (A = B) = (A ↔ B) := by -- first -- | exact? says exact le_antisymm hxy hyx -- | exact? says exact ge_antisymm hyx hxy + +/-- +info: Try this: refine Int.mul_ne_zero ?a0 h +--- +warning: declaration uses 'sorry' +-/ +#guard_msgs in +example {x : Int} (h : x ≠ 0) : 2 * x ≠ 0 := by + std_apply? using h + +-- Check that adding `with_reducible` prevents expensive kernel reductions. +-- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/.60exact.3F.60.20failure.3A.20.22maximum.20recursion.20depth.20has.20been.20reached.22/near/417649319 +/-- info: Try this: exact Nat.add_comm n m -/ +#guard_msgs in +example (_h : List.range 10000 = List.range 10000) (n m : Nat) : n + m = m + n := by + with_reducible std_exact? diff --git a/test/omega/benchmark.lean b/test/omega/benchmark.lean new file mode 100644 index 0000000000..0e6f2b2d15 --- /dev/null +++ b/test/omega/benchmark.lean @@ -0,0 +1,341 @@ +import Std.Tactic.Omega.Frontend + +/-! +# Benchmarking the `omega` tactic + +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; +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, +with one test removed (in which a test-for-failure succeeds with today's `omega`). + +The benchmark consists of `lake build && hyperfine "lake env lean test/omega/benchmark.lean"` +run on a freshly rebooted machine! + +2024-02-06 feat: omega uses Lean.HashMap instead of Std.Data.HashMap (#588) +kim@carica std4 % lake build && hyperfine "lake env lean test/omega/benchmark.lean" +Benchmark 1: lake env lean test/omega/benchmark.lean + Time (mean ± σ): 2.530 s ± 0.008 s [User: 2.249 s, System: 0.276 s] + Range (min … max): 2.513 s … 2.542 s 10 runs + +2024-02-03 feat: omega handles min, max, if (#575) +kim@carica std4 % lake build && hyperfine "lake env lean test/omega/benchmark.lean" +Benchmark 1: lake env lean test/omega/benchmark.lean + Time (mean ± σ): 2.526 s ± 0.009 s [User: 2.250 s, System: 0.272 s] + Range (min … max): 2.513 s … 2.542 s 10 runs + +2024-02-02 fix: revert OmegaM state when not multiplying out (#570) +kim@carica std4 % lake build && hyperfine "lake env lean test/omega/benchmark.lean" +Benchmark 1: lake env lean test/omega/benchmark.lean + Time (mean ± σ): 2.569 s ± 0.004 s [User: 2.291 s, System: 0.273 s] + Range (min … max): 2.563 s … 2.574 s 10 runs + +2024-01-12 feat: omega handles double negation and implication hypotheses (#522) +kim@carica std4 % lake build && hyperfine "lake env lean test/omega/benchmark.lean" +Benchmark 1: lake env lean test/omega/benchmark.lean + Time (mean ± σ): 2.575 s ± 0.004 s [User: 2.302 s, System: 0.268 s] + Range (min … max): 2.570 s … 2.581 s 10 runs + +2024-01-10 feat: omega understands Prod.Lex (#511) +kim@carica std4 % lake build && hyperfine "lake env lean test/omega/benchmark.lean" +Benchmark 1: lake env lean test/omega/benchmark.lean + Time (mean ± σ): 2.567 s ± 0.006 s [User: 2.295 s, System: 0.268 s] + Range (min … max): 2.559 s … 2.576 s 10 runs + +2024-01-10 feat: omega handles iff and implications (#503) +kim@carica std4 % lake build && hyperfine "lake env lean test/omega/benchmark.lean" +Benchmark 1: lake env lean test/omega/benchmark.lean + Time (mean ± σ): 2.348 s ± 0.007 s [User: 2.060 s, System: 0.282 s] + Range (min … max): 2.335 s … 2.356 s 10 runs + +2023-12-21 feat: omega (#463) +kim@carica std4 % lake build && hyperfine "lake env lean test/omega/benchmark.lean" +Benchmark 1: lake env lean test/omega/benchmark.lean + Time (mean ± σ): 2.362 s ± 0.008 s [User: 2.080 s, System: 0.277 s] + Range (min … max): 2.349 s … 2.372 s 10 runs + +-/ + +open Std.Tactic.Omega + +example : True := by + fail_if_success omega + trivial + +-- set_option trace.omega true +example (_ : (1 : Int) < (0 : Int)) : False := by omega + +example (_ : (0 : Int) < (0 : Int)) : False := by omega +example (_ : (0 : Int) < (1 : Int)) : True := by (fail_if_success omega); trivial + +example {x : Int} (_ : 0 ≤ x) (_ : x ≤ 1) : True := by (fail_if_success omega); trivial +example {x : Int} (_ : 0 ≤ x) (_ : x ≤ -1) : False := by omega + +example {x : Int} (_ : x % 2 < x - 2 * (x / 2)) : False := by omega +example {x : Int} (_ : x % 2 > 5) : False := by omega + +example {x : Int} (_ : 2 * (x / 2) > x) : False := by omega +example {x : Int} (_ : 2 * (x / 2) ≤ x - 2) : False := by omega + +example {x : Nat} : x / 0 = 0 := by omega +example {x : Int} : x / 0 = 0 := by omega + +example {x : Int} : x / 2 + x / (-2) = 0 := by omega + +example (_ : 7 < 3) : False := by omega +example (_ : 0 < 0) : False := by omega + +example {x : Nat} (_ : x > 7) (_ : x < 3) : False := by omega +example {x : Nat} (_ : x ≥ 7) (_ : x ≤ 3) : False := by omega + +example {x y : Nat} (_ : x + y > 10) (_ : x < 5) (_ : y < 5) : False := by omega + +example {x y : Int} (_ : x + y > 10) (_ : 2 * x < 11) (_ : y < 5) : False := by omega +example {x y : Nat} (_ : x + y > 10) (_ : 2 * x < 11) (_ : y < 5) : False := by omega + +example {x y : Int} (_ : 2 * x + 4 * y = 5) : False := by omega +example {x y : Nat} (_ : 2 * x + 4 * y = 5) : False := by omega + +example {x y : Int} (_ : 6 * x + 7 * y = 5) : True := by (fail_if_success omega); trivial + +example {x y : Nat} (_ : 6 * x + 7 * y = 5) : False := by omega + +example {x : Nat} (_ : x < 0) : False := by omega + +example {x y z : Int} (_ : x + y > z) (_ : x < 0) (_ : y < 0) (_ : z > 0) : False := by omega + +example {x y : Nat} (_ : x - y = 0) (_ : x > y) : False := by + fail_if_success omega (config := { splitNatSub := false }) + omega + +example {x y z : Int} (_ : x - y - z = 0) (_ : x > y + z) : False := by omega + +example {x y z : Nat} (_ : x - y - z = 0) (_ : x > y + z) : False := by omega + +example {a b c d e f : Nat} (_ : a - b - c - d - e - f = 0) (_ : a > b + c + d + e + f) : + False := by + omega + +example {x y : Nat} (h₁ : x - y ≤ 0) (h₂ : y < x) : False := by omega + +example {x y : Int} (_ : x / 2 - y / 3 < 1) (_ : 3 * x ≥ 2 * y + 6) : False := by omega + +example {x y : Nat} (_ : x / 2 - y / 3 < 1) (_ : 3 * x ≥ 2 * y + 6) : False := by omega + +example {x y : Nat} (_ : x / 2 - y / 3 < 1) (_ : 3 * x ≥ 2 * y + 4) : False := by omega + +example {x y : Nat} (_ : x / 2 - y / 3 < x % 2) (_ : 3 * x ≥ 2 * y + 4) : False := by omega + +example {x : Int} (h₁ : 5 ≤ x) (h₂ : x ≤ 4) : False := by omega + +example {x : Nat} (h₁ : 5 ≤ x) (h₂ : x ≤ 4) : False := by omega + +example {x : Nat} (h₁ : x / 3 ≥ 2) (h₂ : x < 6) : False := by omega + +example {x : Int} {y : Nat} (_ : 0 < x) (_ : x + y ≤ 0) : False := by omega + +example {a b c : Nat} (_ : a - (b - c) ≤ 5) (_ : b ≥ c + 3) (_ : a + c ≥ b + 6) : False := by omega + +example {x : Nat} : 1 < (1 + ((x + 1 : Nat) : Int) + 2) / 2 := by omega + +example {x : Nat} : (x + 4) / 2 ≤ x + 2 := by omega + +example {x : Int} {m : Nat} (_ : 0 < m) (_ : ¬x % ↑m < (↑m + 1) / 2) : -↑m / 2 ≤ x % ↑m - ↑m := by + omega + + +example (h : (7 : Int) = 0) : False := by omega + +example (h : (7 : Int) ≤ 0) : False := by omega + +example (h : (-7 : Int) + 14 = 0) : False := by omega + +example (h : (-7 : Int) + 14 ≤ 0) : False := by omega + +example (h : (1 : Int) + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 = 0) : False := by + omega + +example (h : (7 : Int) - 14 = 0) : False := by omega + +example (h : (14 : Int) - 7 ≤ 0) : False := by omega + +example (h : (1 : Int) - 1 + 1 - 1 + 1 - 1 + 1 - 1 + 1 - 1 + 1 - 1 + 1 - 1 + 1 = 0) : False := by + omega + +example (h : -(7 : Int) = 0) : False := by omega + +example (h : -(-7 : Int) ≤ 0) : False := by omega + +example (h : 2 * (7 : Int) = 0) : False := by omega + +example (h : (7 : Int) < 0) : False := by omega + +example {x : Int} (h : x + x + 1 = 0) : False := by omega + +example {x : Int} (h : 2 * x + 1 = 0) : False := by omega + +example {x y : Int} (h : x + x + y + y + 1 = 0) : False := by omega + +example {x y : Int} (h : 2 * x + 2 * y + 1 = 0) : False := by omega + +example {x : Int} (h₁ : 0 ≤ -7 + x) (h₂ : 0 ≤ 3 - x) : False := by omega + +example {x : Int} (h₁ : 0 ≤ -7 + x) (h₂ : 0 < 4 - x) : False := by omega + +example {x : Int} (h₁ : 0 ≤ 2 * x + 1) (h₂ : 2 * x + 1 ≤ 0) : False := by omega + +example {x : Int} (h₁ : 0 < 2 * x + 2) (h₂ : 2 * x + 1 ≤ 0) : False := by omega + +example {x y : Int} (h₁ : 0 ≤ 2 * x + 1) (h₂ : x = y) (h₃ : 2 * y + 1 ≤ 0) : False := by omega + +example {x y z : Int} (h₁ : 0 ≤ 2 * x + 1) (h₂ : x = y) (h₃ : y = z) (h₄ : 2 * z + 1 ≤ 0) : + False := by omega + +example {x1 x2 x3 x4 x5 x6 : Int} (h : 0 ≤ 2 * x1 + 1) (h : x1 = x2) (h : x2 = x3) (h : x3 = x4) + (h : x4 = x5) (h : x5 = x6) (h : 2 * x6 + 1 ≤ 0) : False := by omega + +example {x : Int} (_ : 1 ≤ -3 * x) (_ : 1 ≤ 2 * x) : False := by omega + +example {x y : Int} (_ : 2 * x + 3 * y = 0) (_ : 1 ≤ x) (_ : 1 ≤ y) : False := by omega + +example {x y z : Int} (_ : 2 * x + 3 * y = 0) (_ : 3 * y + 4 * z = 0) (_ : 1 ≤ x) (_ : 1 ≤ -z) : + False := by omega + +example {x y z : Int} (_ : 2 * x + 3 * y + 4 * z = 0) (_ : 1 ≤ x + y) (_ : 1 ≤ y + z) + (_ : 1 ≤ x + z) : False := by omega + +example {x y : Int} (_ : 1 ≤ 3 * x) (_ : y ≤ 2) (_ : 6 * x - 2 ≤ y) : False := by omega + +example {x y : Int} (_ : y = x) (_ : 0 ≤ x - 2 * y) (_ : x - 2 * y ≤ 1) (_ : 1 ≤ x) : False := by + omega +example {x y : Int} (_ : y = x) (_ : 0 ≤ x - 2 * y) (_ : x - 2 * y ≤ 1) (_ : x ≥ 1) : False := by + omega +example {x y : Int} (_ : y = x) (_ : 0 ≤ x - 2 * y) (_ : x - 2 * y ≤ 1) (_ : 0 < x) : False := by + omega +example {x y : Int} (_ : y = x) (_ : 0 ≤ x - 2 * y) (_ : x - 2 * y ≤ 1) (_ : x > 0) : False := by + omega + +example {x : Nat} (_ : 10 ∣ x) (_ : ¬ 5 ∣ x) : False := by omega +example {x y : Nat} (_ : 5 ∣ x) (_ : ¬ 10 ∣ x) (_ : y = 7) (_ : x - y ≤ 2) (_ : x ≥ 6) : False := by + omega + +example (x : Nat) : x % 4 - x % 8 = 0 := by omega + +example {n : Nat} (_ : n > 0) : (2*n - 1) % 2 = 1 := by omega + +example (x : Int) (_ : x > 0 ∧ x < -1) : False := by omega +example (x : Int) (_ : x > 7) : x < 0 ∨ x > 3 := by omega + +example (_ : ∃ n : Nat, n < 0) : False := by omega +example (_ : { x : Int // x < 0 ∧ x > 0 }) : False := by omega +example {x y : Int} (_ : x < y) (z : { z : Int // y ≤ z ∧ z ≤ x }) : False := by omega + +example (a b c d e : Int) + (ha : 2 * a + b + c + d + e = 4) + (hb : a + 2 * b + c + d + e = 5) + (hc : a + b + 2 * c + d + e = 6) + (hd : a + b + c + 2 * d + e = 7) + (he : a + b + c + d + 2 * e = 8) : e = 3 := by omega + +example (a b c d e : Int) + (ha : 2 * a + b + c + d + e = 4) + (hb : a + 2 * b + c + d + e = 5) + (hc : a + b + 2 * c + d + e = 6) + (hd : a + b + c + 2 * d + e = 7) + (he : a + b + c + d + 2 * e = 8 ∨ e = 3) : e = 3 := by + fail_if_success omega (config := { splitDisjunctions := false }) + omega + +example {a b : Int} (h : a < b) (w : b < a) : False := by omega + +example (_e b c a v0 v1 : Int) (_h1 : v0 = 5 * a) (_h2 : v1 = 3 * b) (h3 : v0 + v1 + c = 10) : + v0 + 5 + (v1 - 3) + (c - 2) = 10 := by omega + +example (h : (1 : Int) < 0) (_ : ¬ (37 : Int) < 42) (_ : True) (_ : (-7 : Int) < 5) : + (3 : Int) < 7 := by omega + +example (A B : Int) (h : 0 < A * B) : 0 < 8 * (A * B) := by omega + +example (A B : Nat) (h : 7 < A * B) : 0 < A*B/8 := by omega +example (A B : Int) (h : 7 < A * B) : 0 < A*B/8 := by omega + +example (ε : Int) (h1 : ε > 0) : ε / 2 + ε / 3 + ε / 7 < ε := by omega + +example (x y z : Int) (h1 : 2*x < 3*y) (h2 : -4*x + z/2 < 0) (h3 : 12*y - z < 0) : False := by omega + +example (ε : Int) (h1 : ε > 0) : ε / 2 < ε := by omega + +example (ε : Int) (_ : ε > 0) : ε - 2 ≤ ε / 3 + ε / 3 + ε / 3 := by omega +example (ε : Int) (_ : ε > 0) : ε / 3 + ε / 3 + ε / 3 ≤ ε := by omega +example (ε : Int) (_ : ε > 0) : ε - 2 ≤ ε / 3 + ε / 3 + ε / 3 ∧ ε / 3 + ε / 3 + ε / 3 ≤ ε := by + omega + +example (x : Int) (h : 0 < x) : 0 < x / 1 := by omega + +example (x : Int) (h : 5 < x) : 0 < x/2/3 := by omega + +example (_a b _c : Nat) (h2 : b + 2 > 3 + b) : False := by omega +example (_a b _c : Int) (h2 : b + 2 > 3 + b) : False := by omega + +example (g v V c h : Int) (_ : h = 0) (_ : v = V) (_ : V > 0) (_ : g > 0) + (_ : 0 ≤ c) (_ : c < 1) : v ≤ V := by omega + +example (x y z : Int) (h1 : 2 * x < 3 * y) (h2 : -4 * x + 2 * z < 0) (h3 : 12 * y - 4 * z < 0) : + False := by + omega + +example (x y z : Int) (h1 : 2 * x < 3 * y) (h2 : -4 * x + 2 * z < 0) (_h3 : x * y < 5) + (h3 : 12 * y - 4 * z < 0) : False := by omega + +example (a b c : Int) (h1 : a > 0) (h2 : b > 5) (h3 : c < -10) (h4 : a + b - c < 3) : False := by + omega + +example (_ b _ : Int) (h2 : b > 0) (h3 : ¬ b ≥ 0) : False := by + omega + +example (x y z : Int) (hx : x ≤ 3 * y) (h2 : y ≤ 2 * z) (h3 : x ≥ 6 * z) : x = 3 * y := by + omega + +example (x y z : Int) (h1 : 2 * x < 3 * y) (h2 : -4 * x + 2 * z < 0) (_h3 : x * y < 5) : + ¬ 12 * y - 4 * z < 0 := by + omega + +example (x y z : Int) (hx : ¬ x > 3 * y) (h2 : ¬ y > 2 * z) (h3 : x ≥ 6 * z) : x = 3 * y := by + omega + +example (x y : Int) (h : 6 + ((x + 4) * x + (6 + 3 * y) * y) = 3) (h' : (x + 4) * x ≥ 0) + (h'' : (6 + 3 * y) * y ≥ 0) : False := by omega + +example (a : Int) (ha : 0 ≤ a) : 0 * 0 ≤ 2 * a := by omega + +example (x y : Int) (h : x < y) : x ≠ y := by omega + +example (x y : Int) (h : x < y) : ¬ x = y := by omega + +example (x : Int) : id x ≥ x := by omega + +example (prime : Nat → Prop) (x y z : Int) (h1 : 2 * x + ((-3) * y) < 0) (h2 : (-4) * x + 2* z < 0) + (h3 : 12 * y + (-4) * z < 0) (_ : prime 7) : False := by omega + +example (i n : Nat) (h : (2 : Int) ^ i ≤ 2 ^ n) : (0 : Int) ≤ 2 ^ n - 2 ^ i := by omega + +-- Check we use `exfalso` on non-comparison goals. +example (prime : Nat → Prop) (_ b _ : Nat) (h2 : b > 0) (h3 : b < 0) : prime 10 := by + omega + +example (a b c : Nat) (h2 : (2 : Nat) > 3) : a + b - c ≥ 3 := by omega + +-- Verify that we split conjunctions in hypotheses. +example (x y : Int) + (h : 6 + ((x + 4) * x + (6 + 3 * y) * y) = 3 ∧ (x + 4) * x ≥ 0 ∧ (6 + 3 * y) * y ≥ 0) : + False := by omega + +example (mess : Nat → Nat) (S n : Nat) : + mess S + (n * mess S + n * 2 + 1) < n * mess S + mess S + (n * 2 + 2) := by omega + +example (p n p' n' : Nat) (h : p + n' = p' + n) : n + p' = n' + p := by + omega + +example (a b c : Int) (h1 : 32 / a < b) (h2 : b < c) : 32 / a < c := by omega diff --git a/test/omega/examples.lean b/test/omega/examples.lean index 552f9ec1b2..3f530b0e19 100644 --- a/test/omega/examples.lean +++ b/test/omega/examples.lean @@ -2,7 +2,7 @@ import Std.Tactic.Omega.Frontend -- Turn on `trace.omega` to get detailed information about the processing of hypotheses, -- and the justification of the contradiction found. -set_option trace.omega true +-- set_option trace.omega true -- Inequalities example {x y : Nat} (_ : x + y > 10) (_ : x < 5) (_ : y < 5) : False := by omega diff --git a/test/omega/test.lean b/test/omega/test.lean index 45c49753e6..095e8f110a 100644 --- a/test/omega/test.lean +++ b/test/omega/test.lean @@ -198,8 +198,12 @@ example (a b c d e : Int) fail_if_success omega (config := { splitDisjunctions := false }) omega -example {x y : Int} (_ : (x - y).natAbs < 3) (_ : x < 5) (_ : y > 15) : False := by +example {x : Int} (h : x = 7) : x.natAbs = 7 := by fail_if_success omega (config := { splitNatAbs := false }) + fail_if_success omega (config := { splitDisjunctions := false }) + omega + +example {x y : Int} (_ : (x - y).natAbs < 3) (_ : x < 5) (_ : y > 15) : False := by omega example {a b : Int} (h : a < b) (w : b < a) : False := by omega @@ -315,8 +319,8 @@ def List.permutationsAux.rec' {C : List α → List α → Sort v} (H0 : ∀ is, | [], is => H0 is | t :: ts, is => H1 t ts is (permutationsAux.rec' H0 H1 ts (t :: is)) (permutationsAux.rec' H0 H1 is []) - termination_by _ ts is => (length ts + length is, length ts) - decreasing_by simp_wf; omega + termination_by ts is => (length ts + length is, length ts) + decreasing_by all_goals simp_wf; omega example {x y w z : Nat} (h : Prod.Lex (· < ·) (· < ·) (x + 1, y + 1) (w, z)) : Prod.Lex (· < ·) (· < ·) (x, y) (w, z) := by omega @@ -343,3 +347,39 @@ example (a : Nat) (h : a < 0) : Nat → Nat := by omega example {a₁ a₂ p₁ p₂ : Nat} (h₁ : a₁ = a₂ → ¬p₁ = p₂) : (a₁ < a₂ ∨ a₁ = a₂ ∧ p₁ < p₂) ∨ a₂ < a₁ ∨ a₂ = a₁ ∧ p₂ < p₁ := by omega + +-- From https://github.com/leanprover/std4/issues/562 +example {i : Nat} (h1 : i < 330) (_h2 : 7 ∣ (660 + i) * (1319 - i)) : 1319 - i < 1979 := by + omega + +example {a : Int} (_ : a < min a b) : False := by omega (config := { splitMinMax := false }) +example {a : Int} (_ : max a b < b) : False := by omega (config := { splitMinMax := false }) +example {a : Nat} (_ : a < min a b) : False := by omega (config := { splitMinMax := false }) +example {a : Nat} (_ : max a b < b) : False := by omega (config := { splitMinMax := false }) + +example {a b : Nat} (_ : a = 7) (_ : b = 3) : min a b = 3 := by + fail_if_success omega (config := { splitMinMax := false }) + omega + +example {a b : Nat} (_ : a + b = 9) : (min a b) % 2 + (max a b) % 2 = 1 := by + fail_if_success omega (config := { splitMinMax := false }) + omega + +example {a : Int} (_ : a < if a ≤ b then a else b) : False := by omega +example {a b : Int} : (if a < b then a else b - 1) ≤ b := by omega + +-- Check that we use local values. +example (i j : Nat) (p : i ≥ j) : True := by + let l := j - 1 + have _ : i ≥ l := by omega + trivial + +example (i j : Nat) (p : i ≥ j) : True := by + let l := j - 1 + let k := l + have _ : i ≥ k := by omega + trivial + +example (i : Fin 7) : (i : Nat) < 8 := by omega + +example (x y z i : Nat) (hz : z ≤ 1) : x % 2 ^ i + y % 2 ^ i + z < 2 * 2^ i := by omega diff --git a/test/on_goal.lean b/test/on_goal.lean index e669a624f7..b22d628583 100644 --- a/test/on_goal.lean +++ b/test/on_goal.lean @@ -1,5 +1,6 @@ import Std.Tactic.GuardExpr import Std.Tactic.PermuteGoals +import Std.Tactic.Unreachable example (p q r : Prop) : p → q → r → p ∧ q ∧ r := by intros diff --git a/test/rcases.lean b/test/rcases.lean index 14e33afe30..f1c4039205 100644 --- a/test/rcases.lean +++ b/test/rcases.lean @@ -188,6 +188,11 @@ example (h : a ≤ 2 ∨ 2 < a) : True := by · guard_hyp ha1 : a ≤ 2; trivial · guard_hyp ha2 : 3 ≤ a; trivial +example (a : Nat) : True := by + rcases h : a with _ | n + · guard_hyp h : a = 0; trivial + · guard_hyp h : a = n + 1; trivial + inductive BaseType : Type where | one diff --git a/test/solve_by_elim.lean b/test/solve_by_elim.lean index dad0ff45fd..779a971c3b 100644 --- a/test/solve_by_elim.lean +++ b/test/solve_by_elim.lean @@ -5,6 +5,7 @@ Authors: Scott Morrison -/ import Std.Tactic.RCases import Std.Tactic.SolveByElim +import Std.Tactic.PermuteGoals import Std.Test.Internal.DummyLabelAttr set_option autoImplicit true @@ -41,6 +42,7 @@ example {α β γ : Type} (_f : α → β) (g : β → γ) (b : β) : γ := by s example {α : Nat → Type} (f : (n : Nat) → α n → α (n+1)) (a : α 0) : α 4 := by solve_by_elim only [f, a] +set_option linter.unusedVariables false in example (h₁ h₂ : False) : True := by -- 'It doesn't make sense to remove local hypotheses when using `only` without `*`.' fail_if_success solve_by_elim only [-h₁] @@ -54,7 +56,7 @@ example (P₁ P₂ : α → Prop) (f : ∀ (a : α), P₁ a → P₂ a → β) solve_by_elim example {X : Type} (x : X) : x = x := by - fail_if_success solve_by_elim only -- needs the `rfl` lemma + fail_if_success solve_by_elim (config := {constructor := false}) only -- needs the `rfl` lemma solve_by_elim -- Needs to apply `rfl` twice, with different implicit arguments each time. @@ -62,7 +64,7 @@ example {X : Type} (x : X) : x = x := by example {X : Type} (x y : X) (p : Prop) (h : x = x → y = y → p) : p := by solve_by_elim example : True := by - fail_if_success solve_by_elim only -- needs the `trivial` lemma + fail_if_success solve_by_elim (config := {constructor := false}) only -- needs the `trivial` lemma solve_by_elim -- Requires backtracking. @@ -118,20 +120,18 @@ example : 6 = 6 ∧ [7] = [7] := by fconstructor solve_by_elim* only [@rfl _] --- TODO: restore this test after #241 is merged --- -- Test that `solve_by_elim*`, which works on multiple goals, --- -- successfully uses the relevant local hypotheses for each goal. --- example (f g : Nat → Prop) : (∃ k : Nat, f k) ∨ (∃ k : Nat, g k) ↔ ∃ k : Nat, f k ∨ g k := by --- fconstructor --- rintro (⟨n, fn⟩ | ⟨n, gn⟩) --- pick_goal 3 --- rintro ⟨n, hf | hg⟩ --- solve_by_elim* (config := {maxDepth := 13}) [Or.inl, Or.inr, Exists.intro] +-- Test that `solve_by_elim*`, which works on multiple goals, +-- successfully uses the relevant local hypotheses for each goal. +example (f g : Nat → Prop) : (∃ k : Nat, f k) ∨ (∃ k : Nat, g k) ↔ ∃ k : Nat, f k ∨ g k := by + fconstructor + rintro (⟨n, fn⟩ | ⟨n, gn⟩) + on_goal 3 => rintro ⟨n, hf | hg⟩ + solve_by_elim* (config := {maxDepth := 13}) [Or.inl, Or.inr, Exists.intro] -- Test that `Config.intros` causes `solve_by_elim` to call `intro` on intermediate goals. example (P : Prop) : P → P := by - fail_if_success solve_by_elim - solve_by_elim (config := .intros) + fail_if_success solve_by_elim (config := {intros := false}) + solve_by_elim -- This worked in mathlib3 without the `@`, but now goes into a loop. -- If someone wants to diagnose this, please do! @@ -183,3 +183,8 @@ example : 5 ≤ 7 := by exact mySorry end issue1581 + +example (x : (α × (β × γ))) : (α × β) × γ := by + rcases x with ⟨a, b, c⟩ + fail_if_success solve_by_elim (config := {constructor := false}) + solve_by_elim