Not working, yet
This commit is contained in:
29
mambo.agda
29
mambo.agda
@@ -8,13 +8,14 @@ data Bottom : Type where
|
|||||||
¬ : Type → Type
|
¬ : Type → Type
|
||||||
¬ A = A → Bottom
|
¬ A = A → Bottom
|
||||||
|
|
||||||
Rel : Type → Type₁
|
Rel : Type → Type → Type₁
|
||||||
Rel A = A → A → Type
|
Rel A B = A → B → Type
|
||||||
|
|
||||||
record _×_ (A B : Type) : Type where
|
data Σ (A : Type) (B : A -> Type) : Type where
|
||||||
field
|
pair : (a : A) -> (b : B a) -> Σ A B
|
||||||
fst : A
|
|
||||||
snd : B
|
∃ : {A : Type} (B : A → Type) → Type
|
||||||
|
∃ {A} B = Σ A B
|
||||||
|
|
||||||
data Formula : Type where
|
data Formula : Type where
|
||||||
⊥ : Formula
|
⊥ : Formula
|
||||||
@@ -50,6 +51,7 @@ data _∈_ {A : Type} : A → List A → Type where
|
|||||||
zero : (x : A) (xs : List A) → x ∈ (x ∷ xs)
|
zero : (x : A) (xs : List A) → x ∈ (x ∷ xs)
|
||||||
succ : {y : A} (x : A) (xs : List A) → (x ∈ xs) → (x ∈ (y ∷ xs))
|
succ : {y : A} (x : A) (xs : List A) → (x ∈ xs) → (x ∈ (y ∷ xs))
|
||||||
|
|
||||||
|
|
||||||
infixr 2 _⊢_
|
infixr 2 _⊢_
|
||||||
data _⊢_ : Context → Formula → Type where
|
data _⊢_ : Context → Formula → Type where
|
||||||
var : X ∈ Γ → Γ ⊢ X
|
var : X ∈ Γ → Γ ⊢ X
|
||||||
@@ -63,17 +65,22 @@ data _⊢_ : Context → Formula → Type where
|
|||||||
⇒ᵢ : Γ , X ⊢ Y → Γ ⊢ X ⇒ Y
|
⇒ᵢ : Γ , X ⊢ Y → Γ ⊢ X ⇒ Y
|
||||||
¬ᵢ : Γ , X ⊢ ⊥ → Γ ⊢ ∼ X
|
¬ᵢ : Γ , X ⊢ ⊥ → Γ ⊢ ∼ X
|
||||||
¬ₑ : Γ ⊢ X → Γ ⊢ ∼ X → Γ ⊢ ⊥
|
¬ₑ : Γ ⊢ X → Γ ⊢ ∼ X → Γ ⊢ ⊥
|
||||||
-- ¬¬ₑ : Γ ⊢ ¬ ¬ X → Γ ⊢ X
|
|
||||||
-- Unsure if i want classical logic
|
|
||||||
-- TODO: entailments for ◇ and □
|
-- TODO: entailments for ◇ and □
|
||||||
|
|
||||||
record M (W : Type) : Type₁ where
|
record M (W : Type) : Type₁ where
|
||||||
field
|
field
|
||||||
R : Rel (W × W)
|
R : Rel W W
|
||||||
L : W → List Nat
|
L : W → List Nat
|
||||||
|
|
||||||
infixr 2 _,_⊩_
|
infixr 2 _,_⊩_
|
||||||
data _,_⊩_ {W : Type} (Model : M W) (x : W) : Formula → Type where
|
data _,_⊩_ {W : Type} (Model : M W) (x : W) : (p : Formula) → Type where
|
||||||
top : Model , x ⊩ ⊤
|
top : Model , x ⊩ ⊤
|
||||||
atom : {p : Nat} → p ∈ (M.L Model x) → Model , x ⊩ atom p
|
atom : {p : Nat} → p ∈ (M.L Model x) → Model , x ⊩ atom p
|
||||||
-- TODO: The rest :D
|
neg : {p : Formula} → ¬ (Model , x ⊩ p) → Model , x ⊩ ∼ p
|
||||||
|
and : {p q : Formula} → Model , x ⊩ p → Model , x ⊩ q → Model , x ⊩ p ∧ q
|
||||||
|
orl : {p q : Formula} → Model , x ⊩ p → Model , x ⊩ p ∨ q
|
||||||
|
orr : {p q : Formula} → Model , x ⊩ q → Model , x ⊩ p ∨ q
|
||||||
|
impl : {p q : Formula} → Model , x ⊩ p → Model , x ⊩ q → Model , x ⊩ p ⇒ q
|
||||||
|
-- impln : {p q : Formula} → ¬ (Model , x ⊩ p) → Model , x ⊩ p ⇒ q
|
||||||
|
box : {p : Formula} → (∀ (y : W) → M.R Model x y → Model , y ⊩ p) → Model , x ⊩ □ p
|
||||||
|
dia : {p : Formula} → (∃ λ (y : W) → M.R Model x y → Model , y ⊩ p) → Model , x ⊩ ◇ p
|
||||||
|
|||||||
Reference in New Issue
Block a user