{-|
This module contains examples of terms and reductions in λ○▷pat.
-}

{-# OPTIONS --safe #-}
{-# OPTIONS --with-K #-}
module Pat.Examples where

open import Data.Nat using (; zero; suc; _≥‴_; ≤‴-refl; ≤‴-step)
open import Function using (id; _∘_)
open import Relation.Binary.PropositionalEquality using (_≡_; refl)

open import Pat.Context
open import Pat.Term
open import Pat.Reduction

-- | Apply a function to a value n times.
iterate : {A : Set}    (A  A)  A  A
iterate zero f x = x
iterate (suc n) f x = f (iterate n f x)

-- | Run a single step of reduction.
step : Term 0  A  Term 0  A
step e with progress e
... | p-step {e′ = e′} _ = e′
... | p-value (`λ⟨ Δ  e) = `λ⟨ Δ  e
... | p-value `true = `true
... | p-value `false = `false
... | p-value  e  =  e 
... | p-value (`wrap Δ e ve) = `wrap Δ e

-- | Evaluate a term by repeatedly applying the step function.
eval : Term 0  A  Term 0  A
eval = iterate 1000 step

-- | (if true then true else false) evaluates to true.
test₁ : eval (`if `true `then `true `else `false)  `true
test₁ = refl

-- | let$ x = <true> in < if x then true else false > evaluates to < if true then true else false >.
test₂ : eval (`let⟨    `true   `if (` # 0 `with ) `then `true `else `false )   `if `true `then `true `else `false 
test₂ = refl

{-|
let$ y : (x : [z : bool ⊢ bool] ⊢ bool)  =
  let$ z = <true> in < x with z = z >
in <true>

evaluates to <true>
-}
test₃ : eval (`let⟨  ,[ ( , `Bool ^ 1)  `Bool ^ ≤‴-refl ]  (`let⟨     `true   ` # 1 `with ( , var Z) )  `true )   `true 
test₃ = refl

{-|
let$ y : (x : [z : bool ⊢ bool] ⊢ bool)  =
  let$ z = <true> in < x with z = z >
in
  let$ x : (z : bool ⊢ bool) = < if z then true else false > in
  let$ z = <false> in
  <y with x = x>

evaluates to <if true then true else false>
-}
test₄ : eval (`let⟨  ,[ ( , `Bool ^ 1)  `Bool ^ ≤‴-refl ]  (`let⟨    `true   ` # 1 `with ( , var Z) ) (`let⟨  , `Bool ^ 1   `if ` # 0 `with  `then `true `else `false  (`let⟨    `false   ` # 2 `with ( , var (# 1)))))   `if `true `then `true `else `false 
test₄ = refl

{-|
let$ y : (x : bool ⊢ bool) = <if x then true else false> in
let$ z = <true> in
<y with x = z>

evaluates to <if true then true else false>
-}
test₅ : eval (`let⟨  , `Bool ^ 1   `if (` # 0 `with ) `then `true `else `false  (`let⟨    `true   ` # 1 `with ( , var (# 0))  ))   `if `true `then `true `else `false 
test₅ = refl

-- | The aif term described in section 2.4.
`aif :   [    `Bool ]⇒ ([  , `Bool ^ 1   `Bool ]⇒ ([  , `Bool ^ 1   `Bool ]⇒  `Bool)) ^ 0
`aif = `λ⟨   (`λ⟨  , `Bool ^ 1  (`λ⟨  , `Bool ^ 1  
  `let⟨   (` L (# 2) `with )
  (`let⟨  , `Bool ^ 1  (` (# 3) `with ( , var (# 0)))
  (`let⟨  , `Bool ^ 1  (` (# 3) `with ( , var (# 0)))
     (`λ⟨   `if (` # 0 `with ) `then ` # 2 `with ( , var (# 0)) `else ` # 1 `with ( , var (# 0))) · ` # 2 `with  ))))

-- | The wrapToArr term described in section 3.
`wrap-to-arr :  , (Δ  A)  B  ^ n  [ Δ  A ]⇒ B ^ n
`wrap-to-arr {Δ = Δ} = `λ⟨ Δ  ` # 1 `with  · `wrap Δ (` L (# 0) `with lift R)

-- | The arrToWrap term described in section 3.
`arr-to-wrap :  , [ Δ  A ]⇒ B ^ n  Δ  A  B ^ n
`arr-to-wrap {Δ = Δ} = `λ⟨   ` # 1 `with  · `letwrap Δ (` L (# 0) `with ) (` (# 0) `with lift (S  R))

-- | Pushing a ◯ into a ▷.
◯▷⇒▷◯ :  Δ   ,  (Δ  A) ^ n  inject₁ Δ   A ^ n
◯▷⇒▷◯ Δ = `let⟨   (` # 0 `with ) (`wrap (inject₁ Δ)  `letwrap Δ (` L (# 0) `with ) (` # 0 `with lift (S  R  ∋-↾⁺  ∋-inject₁⁺)) )

-- | Pulling a ◯ out of a ▷.
▷◯⇒◯▷ :  Δ   , inject₁ Δ   A ^ n   (Δ  A) ^ n
▷◯⇒◯▷ Δ = `letwrap (inject₁ Δ) (` Z `with ) (`let⟨ inject₁ Δ  (` L (# 0) `with lift R)  `wrap Δ (` L (# 0)  `with lift R)  )

-- | Converting a later-stage substitution into a function between unhygienic values.
subst-to-arr : (σ : Subst (inject₁ Γ) (inject₁ Γ′))    [ Γ  A ]⇒ (Γ′  A) ^ n
subst-to-arr {Γ = Γ} {Γ′ = Γ′} σ = `λ⟨ Γ  `wrap Γ′ (` L (# 0) `with renameSubst R σ)

{-|
A variant of the let-wrap rule that allows introducing additional dependencies
in e₁ is admissible:

Γ , Δⁿ⁺¹ ⊢ⁿ e₁ : Δ′ ▷ A
Γ , x : [ Δ, Δ′ ⊢ⁿ A ] ⊢ⁿ e₂ : B
------------------
Γ ⊢ⁿ let_Δ_wrap_Δ′ x = e₁ in e₂ : B
-}
`letΔwrap-admissible :  Δ
     (e₁ : Γ ++ inject₁ Δ  Δ′  A ^ n)
     (e₂ : Γ ,[ inject₁ Δ ++ inject₁ Δ′  A ^ ≤‴-refl ]  B ^ n)
      --------------------------------------
     Γ  B ^ n
`letΔwrap-admissible {Δ′ = Δ′} Δ e₁ e₂ =
  `letwrap (Δ ++ Δ′)
    (`wrap (Δ ++ Δ′)
      (`letwrap Δ′
        (rename ρ₁ e₁)
        (` # 0 `with lift (S  ρ₂))))
    (rename ρ₃ e₂)
  where
  ρ₁ : Rename (Γ ++ inject₁ Δ) (Γ ++ inject₁ (Δ ++ Δ′))
  ρ₁ (L x) = L x
  ρ₁ (R x) = R (L x)
  ρ₂ : Rename (inject₁ Δ′) (Γ ++ inject₁ (Δ ++ Δ′))
  ρ₂ x = R (R x)
  ρ₃ : Rename (Γ ,[ inject₁ Δ ++ inject₁ Δ′  A ^ ≤‴-refl ]) (Γ ,[ inject₁ (Δ ++ Δ′)  A ^ ≤‴-refl ])
  ρ₃ x = x