Documentation

Mathlib.Tactic.CC.Lemmas

Lemmas use by the congruence closure module

theorem Mathlib.Tactic.CC.iff_eq_of_eq_true_left {a : Prop} {b : Prop} (h : a = True) :
(a b) = b
theorem Mathlib.Tactic.CC.iff_eq_of_eq_true_right {a : Prop} {b : Prop} (h : b = True) :
(a b) = a
theorem Mathlib.Tactic.CC.iff_eq_true_of_eq {a : Prop} {b : Prop} (h : a = b) :
(a b) = True
theorem Mathlib.Tactic.CC.and_eq_of_eq_true_left {a : Prop} {b : Prop} (h : a = True) :
(a b) = b
theorem Mathlib.Tactic.CC.and_eq_of_eq_true_right {a : Prop} {b : Prop} (h : b = True) :
(a b) = a
theorem Mathlib.Tactic.CC.and_eq_of_eq {a : Prop} {b : Prop} (h : a = b) :
(a b) = a
theorem Mathlib.Tactic.CC.or_eq_of_eq_false_left {a : Prop} {b : Prop} (h : a = False) :
(a b) = b
theorem Mathlib.Tactic.CC.or_eq_of_eq {a : Prop} {b : Prop} (h : a = b) :
(a b) = a
theorem Mathlib.Tactic.CC.imp_eq_of_eq_true_left {a : Prop} {b : Prop} (h : a = True) :
(ab) = b
theorem Mathlib.Tactic.CC.imp_eq_of_eq_true_right {a : Prop} {b : Prop} (h : b = True) :
(ab) = True
theorem Mathlib.Tactic.CC.imp_eq_of_eq_false_left {a : Prop} {b : Prop} (h : a = False) :
(ab) = True
theorem Mathlib.Tactic.CC.imp_eq_of_eq_false_right {a : Prop} {b : Prop} (h : b = False) :
(ab) = ¬a
theorem Mathlib.Tactic.CC.not_imp_eq_of_eq_false_right {a : Prop} {b : Prop} (h : b = False) :
(¬ab) = a
theorem Mathlib.Tactic.CC.imp_eq_true_of_eq {a : Prop} {b : Prop} (h : a = b) :
(ab) = True
theorem Mathlib.Tactic.CC.if_eq_of_eq_true {c : Prop} [d : Decidable c] {α : Sort u} (t : α) (e : α) (h : c = True) :
(if c then t else e) = t
theorem Mathlib.Tactic.CC.if_eq_of_eq_false {c : Prop} [d : Decidable c] {α : Sort u} (t : α) (e : α) (h : c = False) :
(if c then t else e) = e
theorem Mathlib.Tactic.CC.if_eq_of_eq (c : Prop) [d : Decidable c] {α : Sort u} {t : α} {e : α} (h : t = e) :
(if c then t else e) = t