Description: Inference quantifying both antecedent and consequent. (Contributed by NM, 1-Feb-2005)
| Ref | Expression | ||
|---|---|---|---|
| Hypothesis | ralimdv2.1 | ⊢ ( 𝜑 → ( ( 𝑥 ∈ 𝐴 → 𝜓 ) → ( 𝑥 ∈ 𝐵 → 𝜒 ) ) ) | |
| Assertion | ralimdv2 | ⊢ ( 𝜑 → ( ∀ 𝑥 ∈ 𝐴 𝜓 → ∀ 𝑥 ∈ 𝐵 𝜒 ) ) |
| Step | Hyp | Ref | Expression |
|---|---|---|---|
| 1 | ralimdv2.1 | ⊢ ( 𝜑 → ( ( 𝑥 ∈ 𝐴 → 𝜓 ) → ( 𝑥 ∈ 𝐵 → 𝜒 ) ) ) | |
| 2 | 1 | alimdv | ⊢ ( 𝜑 → ( ∀ 𝑥 ( 𝑥 ∈ 𝐴 → 𝜓 ) → ∀ 𝑥 ( 𝑥 ∈ 𝐵 → 𝜒 ) ) ) |
| 3 | df-ral | ⊢ ( ∀ 𝑥 ∈ 𝐴 𝜓 ↔ ∀ 𝑥 ( 𝑥 ∈ 𝐴 → 𝜓 ) ) | |
| 4 | df-ral | ⊢ ( ∀ 𝑥 ∈ 𝐵 𝜒 ↔ ∀ 𝑥 ( 𝑥 ∈ 𝐵 → 𝜒 ) ) | |
| 5 | 2 3 4 | 3imtr4g | ⊢ ( 𝜑 → ( ∀ 𝑥 ∈ 𝐴 𝜓 → ∀ 𝑥 ∈ 𝐵 𝜒 ) ) |