Skip to content

Commit

Permalink
feat: (wip) impl Adequacy, pt. 5
Browse files Browse the repository at this point in the history
  • Loading branch information
rami3l committed Jul 28, 2023
1 parent a5e0455 commit 009228f
Show file tree
Hide file tree
Showing 2 changed files with 57 additions and 7 deletions.
2 changes: 1 addition & 1 deletion Plfl/Untyped/Denotational.lean
Original file line number Diff line number Diff line change
Expand Up @@ -347,7 +347,7 @@ lemma fn_elem (i : v β‡Ύ w βŠ† u) : v β‡Ύ w ∈ u := i rfl

-- https://plfa.github.io/Denotational/#function-values
/-- `IsFn u` means that `u` is a function value. -/
inductive IsFn (u : Value) : Prop | isFn (h : u = v β‡Ύ w)
inductive IsFn (u : Value) : Prop where | isFn (h : u = v β‡Ύ w)

/-- `AllFn v` means that all elements of `v` are function values. -/
def AllFn (v : Value) : Prop := βˆ€ {u}, u ∈ v β†’ IsFn u
Expand Down
62 changes: 56 additions & 6 deletions Plfl/Untyped/Denotational/Adequacy.lean
Original file line number Diff line number Diff line change
Expand Up @@ -10,6 +10,7 @@ open Untyped Untyped.Notation
open Untyped.Subst
open BigStep BigStep.Notation
open Denotational Denotational.Notation
open Soundness (soundness)

-- https://plfa.github.io/Adequacy/#the-property-of-being-greater-or-equal-to-a-function
/-- `GtFn u` means that it is "greater than" a certain function value. -/
Expand Down Expand Up @@ -57,29 +58,78 @@ theorem GtFn.dec (v : Value) : Decidable (GtFn v) := by induction v with

-- https://plfa.github.io/Adequacy/#relating-values-to-closures
mutual
/--
`𝕍 v c` will hold when:
- `c` is in WHNF (i.e. is a Ξ»-abstraction);
- `v` is a function;
- `c`'s body evaluates according to `v`.
-/
def 𝕍 : Value β†’ Clos β†’ Prop
| _, .clos (` _) _ => βŠ₯
| _, .clos (_ β–‘ _) _ => βŠ₯
| βŠ₯, .clos (Ζ› _) _ => ⊀
| v β‡Ύ w, .clos (Ζ› n) Ξ³ =>
have : sizeOf w < 1 + sizeOf v + sizeOf w := by simp_arith; apply Nat.zero_le
βˆ€ c, 𝔼 v c β†’ GtFn w β†’ βˆƒ c', (Ξ³β€š' c ⊒ n ⇓ c') ∧ 𝕍 w c'
βˆ€ {c}, 𝔼 v c β†’ GtFn w β†’ βˆƒ c', (Ξ³β€š' c ⊒ n ⇓ c') ∧ 𝕍 w c'
| .conj u v, c@(.clos (Ζ› _) _) =>
have : sizeOf v < 1 + sizeOf u + sizeOf v := by simp_arith; apply Nat.zero_le
𝕍 u c ∧ 𝕍 v c

/--
`𝔼 v c` will hold when:
- `v` is greater than a function value;
- `c` evaluates to a closure `c'` in WHNF;
- `𝕍 v c` holds.
-/
def 𝔼 : Value β†’ Clos β†’ Prop
| v, .clos m Ξ³' => GtFn v β†’ βˆƒ c, (Ξ³' ⊒ m ⇓ c) ∧ 𝕍 v c
end
-- https://leanprover.zulipchat.com/#narrow/stream/113489-new-members/topic/Termination.20of.20mutual.20recursive.20defs.20with.20a.20.22shorthand.22.3F/near/378733953
-- https://leanprover.zulipchat.com/#narrow/stream/113489-new-members/topic/.E2.9C.94.20Termination.20of.20mutual.20recursive.20defs.20with.20a.20.22shorthand.22.3F/near/378733953
termination_by
𝕍 v c => (sizeOf v, 0)
𝔼 v c => (sizeOf v, 1)

/-- `𝔾` relates `Ξ³` to `Ξ³'` if the corresponding values and closures are related by `𝔼` -/
def 𝔾 (Ξ³ : Env Ξ“) (Ξ³' : ClosEnv Ξ“) : Prop := βˆ€ {x : Ξ“ βˆ‹ ✢}, 𝔼 (Ξ³ x) (Ξ³' x)

-- namespace Notation
-- end Notation
/-- The proof of a term being in Weak-Head Normal Form. -/
def WHNF (t : Ξ“ ⊒ a) : Prop := βˆƒ n : Ξ“β€š ✢ ⊒ ✢, t = (Ζ› n)

-- open Notation
/-- A closure in a 𝕍 relation must be in WHNF. -/
lemma WHNF.of_𝕍 (vc : 𝕍 v (.clos m Ξ³)) : WHNF m := by
cases m with (simp [𝕍] at vc; try contradiction) | lam n => exists n

-- open Soundness (soundness)
lemma 𝕍.conj (uc : 𝕍 u c) (vc : 𝕍 v c) : 𝕍 (u βŠ” v) c := by
let .clos m Ξ³ := c; cases m with (simp [𝕍] at *; try contradiction)
| lam => unfold 𝕍; exact ⟨uc, vc⟩

lemma 𝕍.of_not_gtFn (nf : Β¬ GtFn v) : 𝕍 v (.clos (Ζ› n) Ξ³') := by induction v with unfold 𝕍
| bot => triv
| fn v w => exfalso; apply nf; exists v, w
| conj _ _ ih ih' => exact not_gtFn_conj_inv nf |>.imp ih ih'

mutual
lemma 𝕍.sub (vvc : 𝕍 v c) (lt : v' βŠ‘ v) : 𝕍 v' c := by
let .clos m Ξ³ := c; cases m with (simp [𝕍] at *; try contradiction) | lam m =>
rename_i Ξ“; induction lt generalizing Ξ“ with
| bot => triv
| conjL lt lt' ih ih' => unfold 𝕍; exact ⟨ih _ _ _ vvc, ih' _ _ _ vvc⟩
| conjR₁ lt ih => apply ih; unfold 𝕍 at vvc; exact vvc.1
| conjRβ‚‚ lt ih => apply ih; unfold 𝕍 at vvc; exact vvc.2
| trans lt lt' ih ih' => apply_rules [ih, ih']
| fn lt lt' ih ih' =>
unfold 𝕍; intro c evc gtw; unfold 𝕍 at vvc; rename_i vβ‚‚ v₁ _ _
have : sizeOf v₁ + sizeOf vβ‚‚ < sizeOf v + sizeOf v' := sorry
have ⟨c', ec', h⟩ := vvc (evc.sub lt) (gtw.sub lt'); exists c', ec'
let .clos m Ξ³ := c'; have ⟨m', h'⟩ := WHNF.of_𝕍 h; subst h'; exact ih' _ _ _ h
| dist => sorry

lemma 𝔼.sub (evc : 𝔼 v c) (lt : v' βŠ‘ v) : 𝔼 v' c := by
let .clos m Ξ³ := c; simp only [𝔼] at *; intro gtv'
have ⟨c, ec, vvc⟩ := evc <| gtv'.sub lt; exists c, ec; exact vvc.sub lt
end
termination_by
𝕍.sub => (sizeOf v + sizeOf v', 0)
𝔼.sub => (sizeOf v + sizeOf v', 1)

#print Value.fn.sizeOf_spec

0 comments on commit 009228f

Please sign in to comment.