{-# OPTIONS --universe-polymorphism #-} module WrongMetaLeft where open import Imports.Level postulate ∃₂ : ∀ {a c : Level} {A : Set a} {B : Set} (C : A → B → Set c) → Set (a ⊔ c) proj₂ : ∀ {a c}{A : Set a}{B : Set}{C : A → B → Set c} → ∃₂ {a}{c}{A}{B} C → B postulate Position : Set Result : Set _≡_ : Result → Result → Set postulate Mono : {p : Level} (P : Position → Position → Set p) → Set p postulate Fun : ∀ {a : Level} {A : Set a} → (A → A → Set) → (A → Result) → Set a -- The problem is that the "wrong" meta is left unsolved. It's really the -- level of _ Result -> Set) (_≈_ : ∃₂ Key → ∃₂ Key → Set) (resultsEqual : Fun {_} {∃₂ {_}{_}{Mono _ proj₂ {_}{_}{_}{Result}{_} rfk)) → Set