{-# OPTIONS --cubical-compatible --safe #-}
module Relation.Nullary.Negation where
open import Effect.Monad using (RawMonad; mkRawMonad)
open import Data.Bool.Base using (Bool; false; true; if_then_else_; not)
open import Data.Empty using (⊥-elim)
open import Data.Product.Base as Prod using (_,_; Σ; Σ-syntax; ∃; curry; uncurry)
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_])
open import Function.Base using (flip; _∘_; const; _∘′_)
open import Level using (Level)
open import Relation.Nullary.Decidable.Core using (Dec; yes; no; ¬¬-excluded-middle)
open import Relation.Unary using (Universal; Pred)
private
  variable
    a b c d p w : Level
    A B C D : Set a
    P : Pred A p
    Whatever : Set w
open import Relation.Nullary.Negation.Core public
∃⟶¬∀¬ : ∃ P → ¬ (∀ x → ¬ P x)
∃⟶¬∀¬ = flip uncurry
∀⟶¬∃¬ : (∀ x → P x) → ¬ ∃ λ x → ¬ P x
∀⟶¬∃¬ ∀xPx (x , ¬Px) = ¬Px (∀xPx x)
¬∃⟶∀¬ : ¬ ∃ (λ x → P x) → ∀ x → ¬ P x
¬∃⟶∀¬ = curry
∀¬⟶¬∃ : (∀ x → ¬ P x) → ¬ ∃ (λ x → P x)
∀¬⟶¬∃ = uncurry
∃¬⟶¬∀ : ∃ (λ x → ¬ P x) → ¬ (∀ x → P x)
∃¬⟶¬∀ = flip ∀⟶¬∃¬
¬¬-Monad : RawMonad {a} DoubleNegation
¬¬-Monad = mkRawMonad
  DoubleNegation
  contradiction
  (λ x f → negated-stable (¬¬-map f x))
¬¬-push : DoubleNegation Π[ P ] → Π[ DoubleNegation ∘ P ]
¬¬-push ¬¬∀P a ¬Pa = ¬¬∀P (λ ∀P → ¬Pa (∀P a))
call/cc : ((A → Whatever) → DoubleNegation A) → DoubleNegation A
call/cc hyp ¬a = hyp (λ a → ⊥-elim (¬a a)) ¬a
independence-of-premise : A → (B → Σ A P) → DoubleNegation (Σ[ x ∈ A ] (B → P x))
independence-of-premise {A = A} {B = B} {P = P} q f = ¬¬-map helper ¬¬-excluded-middle
  where
  helper : Dec B → Σ[ x ∈ A ] (B → P x)
  helper (yes p) = Prod.map₂ const (f p)
  helper (no ¬p) = (q , ⊥-elim ∘′ ¬p)
independence-of-premise-⊎ : (A → B ⊎ C) → DoubleNegation ((A → B) ⊎ (A → C))
independence-of-premise-⊎ {A = A} {B = B} {C = C} f = ¬¬-map helper ¬¬-excluded-middle
  where
  helper : Dec A → (A → B) ⊎ (A → C)
  helper (yes p) = Sum.map const const (f p)
  helper (no ¬p) = inj₁ (⊥-elim ∘′ ¬p)
private
  
  
  
  corollary : {B C : Set b} → (A → B ⊎ C) → DoubleNegation ((A → B) ⊎ (A → C))
  corollary {A = A} {B = B} {C = C} f =
    ¬¬-map helper (independence-of-premise true ([ _,_ true , _,_ false ] ∘′ f))
    where
    helper : ∃ (λ b → A → if b then B else C) → (A → B) ⊎ (A → C)
    helper (true  , f) = inj₁ f
    helper (false , f) = inj₂ f