Metamath Proof Explorer
		
		
		
		Description:  The composition of two sets is a set.  (Contributed by SN, 7-Feb-2025)
		
			
				
					|  |  | Ref | Expression | 
					
						|  | Hypotheses | coexd.1 | ⊢ ( 𝜑  →  𝐴  ∈  𝑉 ) | 
					
						|  |  | coexd.2 | ⊢ ( 𝜑  →  𝐵  ∈  𝑊 ) | 
				
					|  | Assertion | coexd | ⊢  ( 𝜑  →  ( 𝐴  ∘  𝐵 )  ∈  V ) | 
			
		
		
			
				Proof
				
					
						| Step | Hyp | Ref | Expression | 
						
							| 1 |  | coexd.1 | ⊢ ( 𝜑  →  𝐴  ∈  𝑉 ) | 
						
							| 2 |  | coexd.2 | ⊢ ( 𝜑  →  𝐵  ∈  𝑊 ) | 
						
							| 3 |  | coexg | ⊢ ( ( 𝐴  ∈  𝑉  ∧  𝐵  ∈  𝑊 )  →  ( 𝐴  ∘  𝐵 )  ∈  V ) | 
						
							| 4 | 1 2 3 | syl2anc | ⊢ ( 𝜑  →  ( 𝐴  ∘  𝐵 )  ∈  V ) |