Fixed binding issues with existential syntax.
This commit is contained in:
parent
a92e1a08c4
commit
7345c36d80
5 changed files with 21 additions and 21 deletions
|
@ -982,15 +982,15 @@ Remember to indent all code by two spaces.
|
|||
```
|
||||
ext∋ : ∀ {Γ B x y}
|
||||
→ x ≢ y
|
||||
→ ¬ (∃[ A ]( Γ ∋ x ⦂ A ))
|
||||
→ ¬ ∃[ A ]( Γ ∋ x ⦂ A )
|
||||
------------------------------
|
||||
→ ¬ (∃[ A ]( Γ , y ⦂ B ∋ x ⦂ A ))
|
||||
→ ¬ ∃[ A ]( Γ , y ⦂ B ∋ x ⦂ A )
|
||||
ext∋ x≢y _ ⟨ A , Z ⟩ = x≢y refl
|
||||
ext∋ _ ¬∃ ⟨ A , S _ ⊢x ⟩ = ¬∃ ⟨ A , ⊢x ⟩
|
||||
|
||||
lookup : ∀ (Γ : Context) (x : Id)
|
||||
------------------------
|
||||
→ Dec (∃[ A ](Γ ∋ x ⦂ A))
|
||||
→ Dec (∃[ A ]( Γ ∋ x ⦂ A ))
|
||||
lookup ∅ x = no (λ ())
|
||||
lookup (Γ , y ⦂ B) x with x ≟ y
|
||||
... | yes refl = yes ⟨ B , Z ⟩
|
||||
|
@ -1006,7 +1006,7 @@ Remember to indent all code by two spaces.
|
|||
→ Γ ⊢ L ↑ A ⇒ B
|
||||
→ ¬ Γ ⊢ M ↓ A
|
||||
----------------------------
|
||||
→ ¬ (∃[ B′ ]( Γ ⊢ L · M ↑ B′ ))
|
||||
→ ¬ ∃[ B′ ]( Γ ⊢ L · M ↑ B′ )
|
||||
¬arg ⊢L ¬⊢M ⟨ B′ , ⊢L′ · ⊢M′ ⟩ rewrite dom≡ (uniq-↑ ⊢L ⊢L′) = ¬⊢M ⊢M′
|
||||
|
||||
¬switch : ∀ {Γ M A B}
|
||||
|
|
|
@ -589,9 +589,9 @@ module Problem3 where
|
|||
```
|
||||
ext∋ : ∀ {Γ B x y}
|
||||
→ x ≢ y
|
||||
→ ¬ ( ∃[ A ] Γ ∋ x ⦂ A )
|
||||
→ ¬ ∃[ A ]( Γ ∋ x ⦂ A )
|
||||
-----------------------------
|
||||
→ ¬ ( ∃[ A ] Γ , y ⦂ B ∋ x ⦂ A )
|
||||
→ ¬ ∃[ A ]( Γ , y ⦂ B ∋ x ⦂ A )
|
||||
ext∋ x≢y _ ⟨ A , Z ⟩ = x≢y refl
|
||||
ext∋ _ ¬∃ ⟨ A , S _ ⊢x ⟩ = ¬∃ ⟨ A , ⊢x ⟩
|
||||
|
||||
|
@ -614,7 +614,7 @@ module Problem3 where
|
|||
→ Γ ⊢ L ↑ A ⇒ B
|
||||
→ ¬ Γ ⊢ M ↓ A
|
||||
----------------------------
|
||||
→ ¬ ( ∃[ B′ ] Γ ⊢ L · M ↑ B′ )
|
||||
→ ¬ ∃[ B′ ]( Γ ⊢ L · M ↑ B′ )
|
||||
¬arg ⊢L ¬⊢M ⟨ B′ , ⊢L′ · ⊢M′ ⟩ rewrite dom≡ (uniq-↑ ⊢L ⊢L′) = ¬⊢M ⊢M′
|
||||
|
||||
¬switch : ∀ {Γ M A B}
|
||||
|
@ -631,7 +631,7 @@ module Problem3 where
|
|||
```
|
||||
synthesize : ∀ (Γ : Context) (M : Term⁺)
|
||||
-----------------------
|
||||
→ Dec (∃[ A ](Γ ⊢ M ↑ A))
|
||||
→ Dec (∃[ A ]( Γ ⊢ M ↑ A ))
|
||||
|
||||
inherit : ∀ (Γ : Context) (M : Term⁻) (A : Type)
|
||||
---------------
|
||||
|
|
|
@ -208,11 +208,11 @@ _decidable_:
|
|||
|
||||
synthesize : ∀ (Γ : Context) (M : Term⁺)
|
||||
-----------------------
|
||||
→ Dec (∃[ A ](Γ ⊢ M ↑ A))
|
||||
→ Dec (∃[ A ]( Γ ⊢ M ↑ A ))
|
||||
|
||||
inherit : ∀ (Γ : Context) (M : Term⁻) (A : Type)
|
||||
---------------
|
||||
→ Dec (Γ ⊢ M ↓ A)
|
||||
---------------
|
||||
→ Dec (Γ ⊢ M ↓ A)
|
||||
|
||||
Given context `Γ` and synthesised term `M`, we must decide whether
|
||||
there exists a type `A` such that `Γ ⊢ M ↑ A` holds, or its negation.
|
||||
|
@ -580,9 +580,9 @@ such that `Γ ∋ x ⦂ A` holds, then there is also no type `A` such that
|
|||
```
|
||||
ext∋ : ∀ {Γ B x y}
|
||||
→ x ≢ y
|
||||
→ ¬ ( ∃[ A ] Γ ∋ x ⦂ A )
|
||||
→ ¬ ∃[ A ]( Γ ∋ x ⦂ A )
|
||||
-----------------------------
|
||||
→ ¬ ( ∃[ A ] Γ , y ⦂ B ∋ x ⦂ A )
|
||||
→ ¬ ∃[ A ]( Γ , y ⦂ B ∋ x ⦂ A )
|
||||
ext∋ x≢y _ ⟨ A , Z ⟩ = x≢y refl
|
||||
ext∋ _ ¬∃ ⟨ A , S _ ∋x ⟩ = ¬∃ ⟨ A , ∋x ⟩
|
||||
```
|
||||
|
@ -597,8 +597,8 @@ there exists a type `A` such that `Γ ∋ x ⦂ A` holds, or its
|
|||
negation:
|
||||
```
|
||||
lookup : ∀ (Γ : Context) (x : Id)
|
||||
-----------------------
|
||||
→ Dec (∃[ A ](Γ ∋ x ⦂ A))
|
||||
-------------------------
|
||||
→ Dec (∃[ A ]( Γ ∋ x ⦂ A ))
|
||||
lookup ∅ x = no (λ ())
|
||||
lookup (Γ , y ⦂ B) x with x ≟ y
|
||||
... | yes refl = yes ⟨ B , Z ⟩
|
||||
|
@ -639,7 +639,7 @@ there is no term `B′` such that `Γ ⊢ L · M ↑ B′` holds:
|
|||
→ Γ ⊢ L ↑ A ⇒ B
|
||||
→ ¬ Γ ⊢ M ↓ A
|
||||
----------------------------
|
||||
→ ¬ ( ∃[ B′ ] Γ ⊢ L · M ↑ B′ )
|
||||
→ ¬ ∃[ B′ ]( Γ ⊢ L · M ↑ B′ )
|
||||
¬arg ⊢L ¬⊢M ⟨ B′ , ⊢L′ · ⊢M′ ⟩ rewrite dom≡ (uniq-↑ ⊢L ⊢L′) = ¬⊢M ⊢M′
|
||||
```
|
||||
Let `⊢L` be evidence that `Γ ⊢ L ↑ A ⇒ B` holds and `¬⊢M` be evidence
|
||||
|
@ -686,8 +686,8 @@ and a type `A` and either returns evidence that `Γ ⊢ M ↓ A`,
|
|||
or its negation:
|
||||
```
|
||||
synthesize : ∀ (Γ : Context) (M : Term⁺)
|
||||
-----------------------
|
||||
→ Dec (∃[ A ](Γ ⊢ M ↑ A))
|
||||
---------------------------
|
||||
→ Dec (∃[ A ]( Γ ⊢ M ↑ A ))
|
||||
|
||||
inherit : ∀ (Γ : Context) (M : Term⁻) (A : Type)
|
||||
---------------
|
||||
|
|
Loading…
Reference in a new issue