Metamath Proof Explorer
		
		
		
		Description:  A lattice ordering is asymmetric.  ( eqss analog.)  (Contributed by NM, 8-Oct-2011)
		
			
				
					|  |  | Ref | Expression | 
					
						|  | Hypotheses | latref.b | ⊢ 𝐵  =  ( Base ‘ 𝐾 ) | 
					
						|  |  | latref.l | ⊢  ≤   =  ( le ‘ 𝐾 ) | 
				
					|  | Assertion | latasym | ⊢  ( ( 𝐾  ∈  Lat  ∧  𝑋  ∈  𝐵  ∧  𝑌  ∈  𝐵 )  →  ( ( 𝑋  ≤  𝑌  ∧  𝑌  ≤  𝑋 )  →  𝑋  =  𝑌 ) ) | 
			
		
		
			
				Proof
				
					
						| Step | Hyp | Ref | Expression | 
						
							| 1 |  | latref.b | ⊢ 𝐵  =  ( Base ‘ 𝐾 ) | 
						
							| 2 |  | latref.l | ⊢  ≤   =  ( le ‘ 𝐾 ) | 
						
							| 3 | 1 2 | latasymb | ⊢ ( ( 𝐾  ∈  Lat  ∧  𝑋  ∈  𝐵  ∧  𝑌  ∈  𝐵 )  →  ( ( 𝑋  ≤  𝑌  ∧  𝑌  ≤  𝑋 )  ↔  𝑋  =  𝑌 ) ) | 
						
							| 4 | 3 | biimpd | ⊢ ( ( 𝐾  ∈  Lat  ∧  𝑋  ∈  𝐵  ∧  𝑌  ∈  𝐵 )  →  ( ( 𝑋  ≤  𝑌  ∧  𝑌  ≤  𝑋 )  →  𝑋  =  𝑌 ) ) |