sstill working
This commit is contained in:
parent
c0191ea593
commit
c790f7bfc3
2 changed files with 100 additions and 16 deletions
|
@ -50,7 +50,7 @@ subst (x ∙ x₁) v = (subst x v) ∙ (subst x₁ v)
|
||||||
data value-rel : type → term → Set where
|
data value-rel : type → term → Set where
|
||||||
v-`true : value-rel bool `true
|
v-`true : value-rel bool `true
|
||||||
v-`false : value-rel bool `false
|
v-`false : value-rel bool `false
|
||||||
v-`λ[_]_ : ∀ {τ e} → value-rel τ (`λ[ τ ] e)
|
v-`λ[_]_ : ∀ {τ₁ τ₂ e} → value-rel (τ₁ -→ τ₂) (`λ[ τ₁ ] e)
|
||||||
|
|
||||||
data good-subst : ctx → sub → Set where
|
data good-subst : ctx → sub → Set where
|
||||||
nil : good-subst nil nil
|
nil : good-subst nil nil
|
||||||
|
@ -60,14 +60,20 @@ data good-subst : ctx → sub → Set where
|
||||||
→ good-subst (cons ctx τ) γ
|
→ good-subst (cons ctx τ) γ
|
||||||
|
|
||||||
data step : term → term → Set where
|
data step : term → term → Set where
|
||||||
|
step-if-e : ∀ {e e' e₁ e₂} → step e e' → step (`if e then e₁ else e₂) (`if e' then e₁ else e₂)
|
||||||
step-if-1 : ∀ {e₁ e₂} → step (`if `true then e₁ else e₂) e₁
|
step-if-1 : ∀ {e₁ e₂} → step (`if `true then e₁ else e₂) e₁
|
||||||
step-if-2 : ∀ {e₁ e₂} → step (`if `false then e₁ else e₂) e₂
|
step-if-2 : ∀ {e₁ e₂} → step (`if `false then e₁ else e₂) e₂
|
||||||
step-`λ : ∀ {τ e v} → step ((`λ[ τ ] e) ∙ v) (subst e v)
|
step-`λ-1 : ∀ {e₁ e₁' e₂} → step e₁ e₁' → step (e₁ ∙ e₂) (e₁' ∙ e₂)
|
||||||
|
step-`λ-2 : ∀ {τ e₁ e₂ e₂'} → step e₂ e₂' → step ((`λ[ τ ] e₁) ∙ e₂) ((`λ[ τ ] e₁) ∙ e₂')
|
||||||
|
step-`λ-β : ∀ {τ e v} → step ((`λ[ τ ] e) ∙ v) (subst e v)
|
||||||
|
|
||||||
data steps : ℕ → term → term → Set where
|
data steps : ℕ → term → term → Set where
|
||||||
zero : ∀ {e} → steps zero e e
|
zero : ∀ {e} → steps zero e e
|
||||||
suc : ∀ {e e' e''} → (n : ℕ) → step e e' → steps n e' e'' → steps (suc n) e e''
|
suc : ∀ {e e' e''} → (n : ℕ) → step e e' → steps n e' e'' → steps (suc n) e e''
|
||||||
|
|
||||||
|
safe : (τ : type) → (e : term) → Set
|
||||||
|
safe τ e = ∀ {n} → (e' : term) → steps n e e' → (value-rel τ e') ⊎ (∃ λ e'' → step e' e'')
|
||||||
|
|
||||||
data _⊢_∶_ : ctx → term → type → Set where
|
data _⊢_∶_ : ctx → term → type → Set where
|
||||||
type-`true : ∀ {ctx} → ctx ⊢ `true ∶ bool
|
type-`true : ∀ {ctx} → ctx ⊢ `true ∶ bool
|
||||||
type-`false : ∀ {ctx} → ctx ⊢ `false ∶ bool
|
type-`false : ∀ {ctx} → ctx ⊢ `false ∶ bool
|
||||||
|
@ -79,9 +85,9 @@ data _⊢_∶_ : ctx → term → type → Set where
|
||||||
type-`x : ∀ {ctx x}
|
type-`x : ∀ {ctx x}
|
||||||
→ (p : Is-just (lookup ctx x))
|
→ (p : Is-just (lookup ctx x))
|
||||||
→ ctx ⊢ (` x) ∶ (to-witness p)
|
→ ctx ⊢ (` x) ∶ (to-witness p)
|
||||||
type-`λ : ∀ {ctx τ τ₂ e}
|
type-`λ : ∀ {ctx τ₁ τ₂ e}
|
||||||
→ (cons ctx τ) ⊢ e ∶ τ₂
|
→ (cons ctx τ₁) ⊢ e ∶ τ₂
|
||||||
→ ctx ⊢ (`λ[ τ ] e) ∶ (τ -→ τ₂)
|
→ ctx ⊢ (`λ[ τ₁ ] e) ∶ (τ₁ -→ τ₂)
|
||||||
type-∙ : ∀ {ctx τ₁ τ₂ e₁ e₂}
|
type-∙ : ∀ {ctx τ₁ τ₂ e₁ e₂}
|
||||||
→ ctx ⊢ e₁ ∶ (τ₁ -→ τ₂)
|
→ ctx ⊢ e₁ ∶ (τ₁ -→ τ₂)
|
||||||
→ ctx ⊢ e₂ ∶ τ₂
|
→ ctx ⊢ e₂ ∶ τ₂
|
||||||
|
@ -90,19 +96,81 @@ data _⊢_∶_ : ctx → term → type → Set where
|
||||||
irreducible : term → Set
|
irreducible : term → Set
|
||||||
irreducible e = ¬ (∃ λ e' → step e e')
|
irreducible e = ¬ (∃ λ e' → step e e')
|
||||||
|
|
||||||
data term-relation : type → term → Set where
|
data term-rel : type → term → Set where
|
||||||
e-term : ∀ {τ e}
|
e-term : ∀ {τ e}
|
||||||
→ (∀ {n} → (e' : term) → steps n e e' → irreducible e' → value-rel τ e')
|
→ (∀ {n} → (e' : term) → steps n e e' → irreducible e' → value-rel τ e')
|
||||||
→ term-relation τ e
|
→ term-rel τ e
|
||||||
|
|
||||||
type-sound : ∀ {Γ e τ} → Γ ⊢ e ∶ τ → Set
|
type-soundness : ∀ {Γ e τ} → Γ ⊢ e ∶ τ → Set
|
||||||
type-sound {Γ} {e} {τ} s = ∀ {n} → (e' : term) → steps n e e' → value-rel τ e' ⊎ ∃ λ e'' → step e' e''
|
type-soundness {Γ} {e} {τ} s = ∀ {n} → (e' : term) → steps n e e' → (value-rel τ e') ⊎ (∃ λ e'' → step e' e'')
|
||||||
|
|
||||||
|
type-sound : (e : term) → (τ : type) → (p : nil ⊢ e ∶ τ) → type-soundness p
|
||||||
|
type-sound .`true .bool type-`true .`true zero = inj₁ v-`true
|
||||||
|
type-sound .`false .bool type-`false .`false zero = inj₁ v-`false
|
||||||
|
type-sound .(`if _ then _ else _) τ (type-`ifthenelse p p₁ p₂) e' steps = {! !}
|
||||||
|
type-sound .(`λ[ _ ] _) .(_ -→ _) (type-`λ p) e' steps = {! !}
|
||||||
|
type-sound .(_ ∙ _) τ (type-∙ p p₁) e' steps = {! !}
|
||||||
|
|
||||||
_⊨_∶_ : (Γ : ctx) → (e : term) → (τ : type) → Set
|
_⊨_∶_ : (Γ : ctx) → (e : term) → (τ : type) → Set
|
||||||
_⊨_∶_ Γ e τ = (γ : sub) → (good-subst Γ γ) → term-relation τ e
|
_⊨_∶_ Γ e τ = (γ : sub) → (good-subst Γ γ) → term-rel τ e
|
||||||
|
|
||||||
fundamental : ∀ {Γ e τ} → (well-typed : Γ ⊢ e ∶ τ) → type-sound well-typed → Γ ⊨ e ∶ τ
|
fundamental : ∀ {Γ e τ} → (well-typed : Γ ⊢ e ∶ τ) → type-soundness well-typed → Γ ⊨ e ∶ τ
|
||||||
fundamental {Γ} {e} {τ} well-typed type-sound γ good-sub = e-term f
|
fundamental {Γ} {e} {τ} well-typed tsound γ good-sub = e-term f
|
||||||
where
|
where
|
||||||
f : {n : ℕ} (e' : term) → steps n e e' → irreducible e' → value-rel τ e'
|
f : {n : ℕ} (e' : term) → steps n e e' → irreducible e' → value-rel τ e'
|
||||||
f e' steps irred = [ id , (λ exists → ⊥-elim (irred exists)) ] (type-sound e' steps)
|
f e' steps irred = [ id , (λ exists → ⊥-elim (irred exists)) ] (tsound e' steps)
|
||||||
|
|
||||||
|
module semantic-type-soundness where
|
||||||
|
part1 : (e : term) → (τ : type) → nil ⊢ e ∶ τ → term-rel τ e
|
||||||
|
part1 (` x) _ (type-`x ())
|
||||||
|
part1 `true bool p = e-term f
|
||||||
|
where
|
||||||
|
f : {n : ℕ} (e' : term) → steps n `true e' → irreducible e' → value-rel bool e'
|
||||||
|
f {.zero} .`true zero irred = v-`true
|
||||||
|
part1 `false bool p = e-term f
|
||||||
|
where
|
||||||
|
f : {n : ℕ} (e' : term) → steps n `false e' → irreducible e' → value-rel bool e'
|
||||||
|
f {zero} .`false zero irred = v-`false
|
||||||
|
part1 (`λ[ τ₁ ] e) (.τ₁ -→ τ₂) (type-`λ p) = e-term f
|
||||||
|
where
|
||||||
|
f : {n : ℕ} (e' : term) → steps n (`λ[ τ₁ ] e) e' → irreducible e' → value-rel (τ₁ -→ τ₂) e'
|
||||||
|
f .(`λ[ τ₁ ] e) zero irred = v-`λ[_]_
|
||||||
|
part1 (`if e then e₁ else e₂) τ (type-`ifthenelse p p₁ p₂) = e-term f
|
||||||
|
where
|
||||||
|
v : ∀ {n} → (e' : term) → steps n e e' → value-rel bool e' → ∃ (step (`if e then e₁ else e₂))
|
||||||
|
v .`true zero v-`true = e₁ , step-if-1
|
||||||
|
v .`false zero v-`false = e₂ , step-if-2
|
||||||
|
v e' (suc {_} {e''} n step steps) val =
|
||||||
|
(`if e'' then e₁ else e₂) , step-if-e step
|
||||||
|
|
||||||
|
f : {n : ℕ} (e' : term) → steps n (`if e then e₁ else e₂) e' → irreducible e' → value-rel τ e'
|
||||||
|
f .(`if e then e₁ else e₂) zero irred with part1 e bool p
|
||||||
|
f .(`if e then e₁ else e₂) zero irred | e-term f' =
|
||||||
|
let
|
||||||
|
wtf = v e zero {! !}
|
||||||
|
wtf3 = {! irred ? !}
|
||||||
|
-- wtf2 = v e wtf
|
||||||
|
-- wtf3 = irred wtf2
|
||||||
|
-- in ⊥-elim wtf3
|
||||||
|
in {! !}
|
||||||
|
f e' (suc n (step-if-e step) steps) irred = {! !}
|
||||||
|
f e' (suc n step-if-1 steps) irred with part1 e₁ τ p₁
|
||||||
|
f e' (suc n step-if-1 steps) irred | e-term f' = f' e' steps irred
|
||||||
|
f e' (suc n step-if-2 steps) irred with part1 e₂ τ p₂
|
||||||
|
f e' (suc n step-if-2 steps) irred | e-term f' = f' e' steps irred
|
||||||
|
part1 (e₁ ∙ e₂) τ₂ (type-∙ {_} {τ₁} p₁ p₂) = e-term f
|
||||||
|
where
|
||||||
|
f : {n : ℕ} (e' : term) → steps n (e₁ ∙ e₂) e' → irreducible e' → value-rel τ₂ e'
|
||||||
|
f .(e₁ ∙ e₂) zero irred = {! !}
|
||||||
|
f e' (suc n (step-`λ-1 step) steps) irred with part1 e₁ (τ₁ -→ τ₂) p₁
|
||||||
|
f e' (suc _ (step-`λ-1 step) steps) irred | e-term f' = {! !}
|
||||||
|
f e' (suc n (step-`λ-2 step) steps) irred = {! !}
|
||||||
|
f e' (suc n step-`λ-β steps) irred = {! !}
|
||||||
|
|
||||||
|
part2 : (e : term) → (τ : type) → term-rel τ e → safe τ e
|
||||||
|
part2 e τ (e-term p) .e zero =
|
||||||
|
inj₁ (p e zero λ q → {! !})
|
||||||
|
where
|
||||||
|
v : ⊥
|
||||||
|
v = {! !}
|
||||||
|
part2 e τ (e-term p) e' (suc {f} {g} n step steps) = inj₂ ({! !} , {! step !})
|
|
@ -196,3 +196,19 @@ ret z$
|
||||||
You can draw a search tree of probabilities. Add up the probabilities to get the probability that a program returns a specific value.
|
You can draw a search tree of probabilities. Add up the probabilities to get the probability that a program returns a specific value.
|
||||||
|
|
||||||
You can share $z$ since it doesn't depend directly on $x$. This builds a *binary decision diagram*.
|
You can share $z$ since it doesn't depend directly on $x$. This builds a *binary decision diagram*.
|
||||||
|
|
||||||
|
=== Knowledge compilation
|
||||||
|
|
||||||
|
https://en.wikipedia.org/wiki/Knowledge_compilation
|
||||||
|
|
||||||
|
Relationship between hardness of propositional reasoning tasks and its syntax of the formula
|
||||||
|
|
||||||
|
SAT for DNF is easy. What kinds of structure enables efficient reasoning?
|
||||||
|
|
||||||
|
==== Succinctness
|
||||||
|
|
||||||
|
$cal(L)_1$ is more succinct than $cal(L)_2$ if it's efficient (polynomial-time) to translate (in a semantics-preserving way) programs written in $cal(L)_2$ to programs written in $cal(L)_1$
|
||||||
|
|
||||||
|
==== Canonicity of BDDs
|
||||||
|
|
||||||
|
There is only 1 structural BDD for any particular formula.
|
Loading…
Reference in a new issue