Metamath Proof Explorer
		
		
		
		Description:  Simple relationship between <_ and >_ .  (Contributed by David A.
     Wheeler, 10-May-2015)  (New usage is discouraged.)
		
			
				
					|  |  | Ref | Expression | 
				
					|  | Assertion | gte-lte | ⊢  ( ( 𝐴  ∈  V  ∧  𝐵  ∈  V )  →  ( 𝐴  ≥  𝐵  ↔  𝐵  ≤  𝐴 ) ) | 
			
		
		
			
				Proof
				
					
						| Step | Hyp | Ref | Expression | 
						
							| 1 |  | df-gte | ⊢  ≥   =  ◡  ≤ | 
						
							| 2 | 1 | breqi | ⊢ ( 𝐴  ≥  𝐵  ↔  𝐴 ◡  ≤  𝐵 ) | 
						
							| 3 |  | brcnvg | ⊢ ( ( 𝐴  ∈  V  ∧  𝐵  ∈  V )  →  ( 𝐴 ◡  ≤  𝐵  ↔  𝐵  ≤  𝐴 ) ) | 
						
							| 4 | 2 3 | bitrid | ⊢ ( ( 𝐴  ∈  V  ∧  𝐵  ∈  V )  →  ( 𝐴  ≥  𝐵  ↔  𝐵  ≤  𝐴 ) ) |