Documentation

Init.Grind.Order

Helper theorems to assert constraints

theorem Lean.Grind.Order.eq_mp {p q : Prop} (h₁ : p = q) (h₂ : p) :
q
theorem Lean.Grind.Order.eq_mp_not {p q : Prop} (h₁ : p = q) (h₂ : ¬p) :
theorem Lean.Grind.Order.eq_trans_true {p q : Prop} (h₁ : p = q) (h₂ : q = True) :
theorem Lean.Grind.Order.eq_trans_false {p q : Prop} (h₁ : p = q) (h₂ : q = False) :
theorem Lean.Grind.Order.eq_trans_true' {p q : Prop} (h₁ : p = q) (h₂ : p = True) :
theorem Lean.Grind.Order.eq_trans_false' {p q : Prop} (h₁ : p = q) (h₂ : p = False) :
theorem Lean.Grind.Order.le_of_eq_1 {α : Type u_1} [LE α] [Std.IsPreorder α] {a b : α} :
a = ba b
theorem Lean.Grind.Order.le_of_eq_2 {α : Type u_1} [LE α] [Std.IsPreorder α] {a b : α} :
a = bb a
theorem Lean.Grind.Order.le_of_eq_1_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] {a b : α} :
a = ba b + 0
theorem Lean.Grind.Order.le_of_eq_2_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] {a b : α} :
a = bb a + 0
theorem Lean.Grind.Order.le_of_not_le {α : Type u_1} [LE α] [Std.IsLinearPreorder α] {a b : α} :
¬a bb a
theorem Lean.Grind.Order.lt_of_not_le {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsLinearPreorder α] {a b : α} :
¬a bb < a
theorem Lean.Grind.Order.le_of_not_lt {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsLinearPreorder α] {a b : α} :
¬a < bb a
theorem Lean.Grind.Order.le_of_not_lt_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsLinearPreorder α] [Ring α] [OrderedRing α] {a b : α} {k k' : Int} :
k'.beq' (-k) = true¬a < b + kb a + k'
theorem Lean.Grind.Order.lt_of_not_le_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsLinearPreorder α] [Ring α] [OrderedRing α] {a b : α} {k k' : Int} :
k'.beq' (-k) = true¬a b + kb < a + k'
theorem Lean.Grind.Order.int_lt {x y k k' : Int} :
k'.beq' (k - 1) = truex < y + kx y + k'

Helper theorem for equality propagation

theorem Lean.Grind.Order.eq_of_le_of_le {α : Type u_1} [LE α] [Std.IsPartialOrder α] {a b : α} :
a bb aa = b

Transitivity

theorem Lean.Grind.Order.le_trans {α : Type u_1} [LE α] [Std.IsPreorder α] {a b c : α} (h₁ : a b) (h₂ : b c) :
a c
theorem Lean.Grind.Order.lt_trans {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] {a b c : α} (h₁ : a < b) (h₂ : b < c) :
a < c
theorem Lean.Grind.Order.le_lt_trans {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] {a b c : α} (h₁ : a b) (h₂ : b < c) :
a < c
theorem Lean.Grind.Order.lt_le_trans {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] {a b c : α} (h₁ : a < b) (h₂ : b c) :
a < c
theorem Lean.Grind.Order.lt_unsat {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] (a : α) :
a < aFalse

Transitivity with offsets

theorem Lean.Grind.Order.le_trans_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] {a b c : α} {k₁ k₂ : Int} (k : Int) (h₁ : a b + k₁) (h₂ : b c + k₂) :
(k == k₂ + k₁) = truea c + k
theorem Lean.Grind.Order.lt_trans_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] {a b c : α} {k₁ k₂ : Int} (k : Int) (h₁ : a < b + k₁) (h₂ : b < c + k₂) :
(k == k₂ + k₁) = truea < c + k
theorem Lean.Grind.Order.le_lt_trans_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] {a b c : α} {k₁ k₂ : Int} (k : Int) (h₁ : a b + k₁) (h₂ : b < c + k₂) :
(k == k₂ + k₁) = truea < c + k
theorem Lean.Grind.Order.lt_le_trans_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] {a b c : α} {k₁ k₂ : Int} (k : Int) (h₁ : a < b + k₁) (h₂ : b c + k₂) :
(k == k₂ + k₁) = truea < c + k

Unsat detection

theorem Lean.Grind.Order.le_unsat_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] {a : α} {k : Int} :
k.blt' 0 = truea a + kFalse
theorem Lean.Grind.Order.lt_unsat_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] {a : α} {k : Int} :
k.ble' 0 = truea < a + kFalse

Helper theorems

Theorems for propagating constraints to True

theorem Lean.Grind.Order.le_eq_true_of_lt {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] {a b : α} :
a < b → (a b) = True
theorem Lean.Grind.Order.le_eq_true_of_le_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] {a b : α} {k₁ k₂ : Int} :
k₁.ble' k₂ = truea b + k₁ → (a b + k₂) = True
theorem Lean.Grind.Order.le_eq_true_of_lt_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] {a b : α} {k₁ k₂ : Int} :
k₁.ble' k₂ = truea < b + k₁ → (a b + k₂) = True
theorem Lean.Grind.Order.lt_eq_true_of_lt_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] {a b : α} {k₁ k₂ : Int} :
k₁.ble' k₂ = truea < b + k₁ → (a < b + k₂) = True
theorem Lean.Grind.Order.lt_eq_true_of_le_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] {a b : α} {k₁ k₂ : Int} :
k₁.blt' k₂ = truea b + k₁ → (a < b + k₂) = True

Theorems for propagating constraints to False

theorem Lean.Grind.Order.le_eq_false_of_lt {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] {a b : α} :
a < b → (b a) = False
theorem Lean.Grind.Order.lt_eq_false_of_lt {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] {a b : α} :
a < b → (b < a) = False
theorem Lean.Grind.Order.lt_eq_false_of_le {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] {a b : α} :
a b → (b < a) = False
theorem Lean.Grind.Order.le_eq_false_of_le_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] {a b : α} {k₁ k₂ : Int} :
(k₂ + k₁).blt' 0 = truea b + k₁ → (b a + k₂) = False
theorem Lean.Grind.Order.lt_eq_false_of_le_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] (a b : α) (k₁ k₂ : Int) :
(k₂ + k₁).ble' 0 = truea b + k₁ → (b < a + k₂) = False
theorem Lean.Grind.Order.lt_eq_false_of_lt_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] (a b : α) (k₁ k₂ : Int) :
(k₂ + k₁).ble' 0 = truea < b + k₁ → (b < a + k₂) = False
theorem Lean.Grind.Order.le_eq_false_of_lt_k {α : Type u_1} [LE α] [LT α] [Std.LawfulOrderLT α] [Std.IsPreorder α] [Ring α] [OrderedRing α] (a b : α) (k₁ k₂ : Int) :
(k₂ + k₁).ble' 0 = truea < b + k₁ → (b a + k₂) = False