mirror of
https://github.com/nmvdw/HITs-Examples
synced 2025-11-03 15:13:51 +01:00
Added Agda code for some HITs
This commit is contained in:
32
Agda-HITs/Imperative/Language.agda
Normal file
32
Agda-HITs/Imperative/Language.agda
Normal file
@@ -0,0 +1,32 @@
|
||||
{-# OPTIONS --without-K --rewriting #-}
|
||||
|
||||
open import HoTT
|
||||
open import Syntax
|
||||
|
||||
module Language where
|
||||
|
||||
data Program : Set where
|
||||
fail : Program
|
||||
exec : Syntax -> State -> Program
|
||||
|
||||
postulate
|
||||
assignp : (z : State) (x n : Nat) -> exec (x := n) z == exec skip ( z [ x :== n ])
|
||||
comp₁ : (z : State) (S : Syntax) -> exec (conc skip S) z == exec S z
|
||||
comp₂ : (z z' : State) (S₁ S₂ S₁' : Syntax) -> exec S₁ z == exec S₁' z' -> exec (conc S₁ S₂) z == exec (conc S₁' S₂) z'
|
||||
while₁ : (z : State) (x n : Nat) (S : Syntax) -> defined z x -> equals z x n -> exec (while x == n do S) z == exec (conc S (while x == n do S)) z
|
||||
while₂ : (z : State) (x n : Nat) (S : Syntax) -> defined z x -> unequals z x n -> exec (while x == n do S) z == exec skip z
|
||||
while₃ : (z : State) (x n : Nat) (S : Syntax) -> undefined z x -> exec (while x == n do S) z == fail
|
||||
|
||||
Program-elim :
|
||||
(Y : Set)
|
||||
-> (failY : Y)
|
||||
-> (execY : Syntax -> State -> Y)
|
||||
-> (assignY : (z : State) (x n : Nat) -> execY (x := n) z == execY skip ( z [ x :== n ]) )
|
||||
-> (compY₁ : (z : State) (S : Syntax) -> execY (conc skip S) z == execY S z )
|
||||
-> (compY₂ : (z z' : State) (S₁ S₂ S₁' : Syntax) -> execY S₁ z == execY S₁' z' -> execY (conc S₁ S₂) z == execY (conc S₁' S₂) z')
|
||||
-> (whileY₁ : (z : State) (x n : Nat) (S : Syntax) -> defined z x -> equals z x n -> execY (while x == n do S) z == execY (conc S (while x == n do S)) z)
|
||||
-> (whileY₂ : (z : State) (x n : Nat) (S : Syntax) -> defined z x -> unequals z x n -> execY (while x == n do S) z == execY skip z)
|
||||
-> (whileY₃ : (z : State) (x n : Nat) (S : Syntax) -> undefined z x -> execY (while x == n do S) z == failY)
|
||||
-> Program -> Y
|
||||
Program-elim _ failY _ _ _ _ _ _ _ fail = failY
|
||||
Program-elim _ _ execY _ _ _ _ _ _ (exec s z) = execY s z
|
||||
18
Agda-HITs/Imperative/Semantics.agda
Normal file
18
Agda-HITs/Imperative/Semantics.agda
Normal file
@@ -0,0 +1,18 @@
|
||||
{-# OPTIONS --without-K --rewriting #-}
|
||||
|
||||
open import HoTT
|
||||
|
||||
module Semantics where
|
||||
|
||||
data koe : Set where
|
||||
a : koe
|
||||
b : koe
|
||||
|
||||
postulate
|
||||
kek : a ↦ b
|
||||
{-# REWRITE kek #-}
|
||||
|
||||
|
||||
Y : koe -> Set
|
||||
Y a = Nat
|
||||
Y b = Bool
|
||||
66
Agda-HITs/Imperative/Syntax.agda
Normal file
66
Agda-HITs/Imperative/Syntax.agda
Normal file
@@ -0,0 +1,66 @@
|
||||
{-# OPTIONS --without-K --rewriting #-}
|
||||
|
||||
open import HoTT
|
||||
|
||||
module Syntax where
|
||||
|
||||
data Maybe (A : Set) : Set where
|
||||
Just : A -> Maybe A
|
||||
Nothing : Maybe A
|
||||
|
||||
eqN : Nat -> Nat -> Bool
|
||||
eqN 0 0 = true
|
||||
eqN 0 _ = false
|
||||
eqN (S _) 0 = false
|
||||
eqN (S n) (S m) = eqN n m
|
||||
|
||||
-- first coordinate represents the variable x_i, second the value
|
||||
State : Set
|
||||
State = List (Nat × Nat)
|
||||
|
||||
_[_:==_] : State -> Nat -> Nat -> State
|
||||
nil [ x :== n ] = (x , n) :: nil
|
||||
((y , m) :: s) [ x :== n ] =
|
||||
if eqN x y
|
||||
then (x , n) :: s
|
||||
else ((y , m) :: (s [ x :== n ]))
|
||||
|
||||
equals : State -> Nat -> Nat -> Set
|
||||
equals nil _ _ = Empty
|
||||
equals ((x , n) :: s) y m =
|
||||
if eqN x y
|
||||
then
|
||||
if eqN n m
|
||||
then Unit
|
||||
else Empty
|
||||
else equals s y m
|
||||
|
||||
unequals : State -> Nat -> Nat -> Set
|
||||
unequals nil _ _ = Unit
|
||||
unequals ((x , n) :: s) y m =
|
||||
if eqN x y
|
||||
then
|
||||
if eqN n m
|
||||
then Empty
|
||||
else Unit
|
||||
else unequals s y m
|
||||
|
||||
defined : State -> Nat -> Set
|
||||
defined nil y = Empty
|
||||
defined ((x , n) :: s) y =
|
||||
if eqN x y
|
||||
then Unit
|
||||
else defined s y
|
||||
|
||||
undefined : State -> Nat -> Set
|
||||
undefined nil y = Unit
|
||||
undefined ((x , n) :: s) y =
|
||||
if eqN x y
|
||||
then Empty
|
||||
else undefined s y
|
||||
|
||||
data Syntax : Set where
|
||||
skip : Syntax
|
||||
_:=_ : Nat -> Nat -> Syntax
|
||||
conc : Syntax -> Syntax -> Syntax
|
||||
while_==_do_ : Nat -> Nat -> Syntax -> Syntax
|
||||
Reference in New Issue
Block a user