fix(reserved_notation): lower binding power of 'iff'

This commit is contained in:
Floris van Doorn 2015-04-22 15:13:34 -04:00 committed by Leonardo de Moura
parent b86ee9dfa6
commit 9d805437f0
5 changed files with 13 additions and 15 deletions

View file

@ -5,7 +5,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
Module: init.reserved_notation
Authors: Leonardo de Moura
-/
prelude
import init.datatypes
@ -18,11 +17,9 @@ notation `take` binders `,` r:(scoped f, f) := r
If a module reassigns these, it will be incompatible with other modules that adhere to these
conventions.
When hovering over a symbol, use "C-u C-x =" to see how to input it.
When hovering over a symbol, use "C-c C-k" to see how to input it.
-/
/- Logical operations and relations -/
definition std.prec.max : num := 1024 -- the strength of application, identifiers, (, [, etc.
definition std.prec.arrow : num := 25
@ -37,22 +34,23 @@ num.succ (num.succ (num.succ (num.succ (num.succ (num.succ (num.succ (num.succ (
/- Logical operations and relations -/
reserve prefix `¬`:40
reserve prefix `~`:40
reserve infixr `∧`:35
reserve infixr `/\`:35
reserve infixr `\/`:30
reserve infixr ``:30
reserve infix `<->`:25
reserve infix `↔`:25
reserve infix `<->`:20
reserve infix `↔`:20
reserve infix `=`:50
reserve infix `≠`:50
reserve infix `≈`:50
reserve infix ``:50
reserve infix `≡`:50
reserve infixr `∘`:60 -- input with \comp
reserve postfix `⁻¹`:std.prec.max_plus -- input with \sy or \-1 or \inv
reserve infixl `⬝`:75
reserve infixr `▸`:75

View file

@ -689,7 +689,7 @@ section port_algebra
algebra.sub_add_eq_sub_sub_swap
theorem sub_eq_iff_eq_add : ∀a b c : , a - b = c ↔ a = c + b := algebra.sub_eq_iff_eq_add
theorem eq_sub_iff_add_eq : ∀a b c : , a = b - c ↔ a + c = b := algebra.eq_sub_iff_add_eq
theorem eq_iff_eq_of_sub_eq_sub : ∀{a b c d : }, a - b = c - d → a = b ↔ c = d :=
theorem eq_iff_eq_of_sub_eq_sub : ∀{a b c d : }, a - b = c - d → (a = b ↔ c = d) :=
@algebra.eq_iff_eq_of_sub_eq_sub _ _
theorem eq_sub_of_add_eq : ∀{a b c : }, a + c = b → a = b - c := @algebra.eq_sub_of_add_eq _ _
theorem sub_eq_of_eq_add : ∀{a b c : }, a = c + b → a - b = c := @algebra.sub_eq_of_eq_add _ _

View file

@ -326,7 +326,7 @@ section port_algebra
theorem add_neg_of_nonpos_of_neg : ∀{a b : }, a ≤ 0 → b < 0 → a + b < 0 :=
@algebra.add_neg_of_nonpos_of_neg _ _
theorem add_eq_zero_iff_eq_zero_and_eq_zero_of_nonneg_of_nonneg : ∀{a b : },
0 ≤ a → 0 ≤ b → a + b = 0 ↔ a = 0 ∧ b = 0 :=
0 ≤ a → 0 ≤ b → (a + b = 0 ↔ a = 0 ∧ b = 0) :=
@algebra.add_eq_zero_iff_eq_zero_and_eq_zero_of_nonneg_of_nonneg _ _
theorem le_add_of_nonneg_of_le : ∀{a b c : }, 0 ≤ a → b ≤ c → b ≤ a + c :=
@ -416,9 +416,9 @@ section port_algebra
algebra.lt_add_iff_sub_lt_left
theorem lt_add_iff_sub_lt_right : ∀a b c : , a < b + c ↔ a - c < b :=
algebra.lt_add_iff_sub_lt_right
theorem le_iff_le_of_sub_eq_sub : ∀{a b c d : }, a - b = c - d → a ≤ b ↔ c ≤ d :=
theorem le_iff_le_of_sub_eq_sub : ∀{a b c d : }, a - b = c - d → (a ≤ b ↔ c ≤ d) :=
@algebra.le_iff_le_of_sub_eq_sub _ _
theorem lt_iff_lt_of_sub_eq_sub : ∀{a b c d : }, a - b = c - d → a < b ↔ c < d :=
theorem lt_iff_lt_of_sub_eq_sub : ∀{a b c d : }, a - b = c - d → (a < b ↔ c < d) :=
@algebra.lt_iff_lt_of_sub_eq_sub _ _
theorem sub_le_sub_left : ∀{a b : }, a ≤ b → ∀c : , c - b ≤ c - a :=
@algebra.sub_le_sub_left _ _

View file

@ -13,7 +13,7 @@ open sigma.ops setoid
constant quot.{l} : Π {A : Type.{l}}, setoid A → Type.{l}
-- Remark: if we do not use propext here, then we would need a quot.lift for propositions.
constant propext {a b : Prop} : a ↔ b → a = b
constant propext {a b : Prop} : (a ↔ b) → a = b
namespace quot
constant mk : Π {A : Type} [s : setoid A], A → quot s

View file

@ -17,7 +17,7 @@ notation `take` binders `,` r:(scoped f, f) := r
If a module reassigns these, it will be incompatible with other modules that adhere to these
conventions.
When hovering over a symbol, use "C-u C-x =" to see how to input it.
When hovering over a symbol, use "C-c C-k" to see how to input it.
-/
definition std.prec.max : num := 1024 -- the strength of application, identifiers, (, [, etc.
@ -40,8 +40,8 @@ reserve infixr `∧`:35
reserve infixr `/\`:35
reserve infixr `\/`:30
reserve infixr ``:30
reserve infix `<->`:25
reserve infix `↔`:25
reserve infix `<->`:20
reserve infix `↔`:20
reserve infix `=`:50
reserve infix `≠`:50
reserve infix `≈`:50