Issue3027b.agda 1.2 KB

12345678910111213141516171819202122232425262728293031323334353637383940414243444546474849
  1. -- Same as Issue3027, but add the constraint in a different order.
  2. open import Agda.Builtin.Nat
  3. open import Agda.Builtin.Bool
  4. open import Agda.Builtin.Equality
  5. F : Bool → Set
  6. F false = Bool
  7. F true = Nat
  8. f : ∀ b → F b → Nat
  9. f false false = 0
  10. f false true = 1
  11. f true x = 2
  12. data D : Nat → Set where
  13. mkD : (b : Bool) (x : F b) → D (f b x)
  14. mutual
  15. ?X : Nat → Set
  16. ?X = _
  17. ?b : Nat → Bool
  18. ?b = _
  19. -- Add a constraints on ?X. This one simplifies (once ?X is solved) to
  20. -- f (?b 0) 1 ≡ 0 (*)
  21. constr₁ : ?X 1 ≡ D 0
  22. constr₁ = refl
  23. -- Here we should check
  24. -- (n : Nat) → ?X n == (x : F (?b 0)) → D (f (?b 0) x)
  25. -- and get stuck on comparing the domains, but special inference
  26. -- for constructors is overeager and compares the target types,
  27. -- solving
  28. -- ?X : Nat → Set
  29. -- ?X x := D (f (?b 0) x)
  30. -- Note that the call to f is not well-typed unless we solve the
  31. -- (as yet unsolved) constraint Nat == F (?b 0).
  32. constr₂ : (n : Nat) → ?X n
  33. constr₂ = mkD (?b 0)
  34. -- Finally, we pick the wrong solution for ?b, causing (*) to become
  35. -- f false 1 ≡ 0
  36. -- which crashes with an impossible when we try to reduce the call to f.
  37. solve-b : ?b ≡ λ _ → false
  38. solve-b = refl