diff --git a/prototyping/Luau/StrictMode.agda b/prototyping/Luau/StrictMode.agda index d3c0f153..7c49a5ac 100644 --- a/prototyping/Luau/StrictMode.agda +++ b/prototyping/Luau/StrictMode.agda @@ -14,6 +14,7 @@ open import Luau.TypeCheck using (_⊢ᴮ_∈_; _⊢ᴱ_∈_; ⊢ᴴ_; ⊢ᴼ_; open import Properties.Contradiction using (¬) open import Properties.TypeCheck using (typeCheckᴮ) open import Properties.Product using (_,_) +open import Properties.ResolveOverloads using (resolve) data Warningᴱ (H : Heap yes) {Γ} : ∀ {M T} → (Γ ⊢ᴱ M ∈ T) → Set data Warningᴮ (H : Heap yes) {Γ} : ∀ {B T} → (Γ ⊢ᴮ B ∈ T) → Set diff --git a/prototyping/Luau/TypeCheck.agda b/prototyping/Luau/TypeCheck.agda index fac25296..05c90b29 100644 --- a/prototyping/Luau/TypeCheck.agda +++ b/prototyping/Luau/TypeCheck.agda @@ -14,8 +14,9 @@ open import Luau.Type using (Type; nil; unknown; number; boolean; string; _⇒_) open import Luau.VarCtxt using (VarCtxt; ∅; _⋒_; _↦_; _⊕_↦_; _⊝_) renaming (_[_] to _[_]ⱽ) open import FFI.Data.Vector using (Vector) open import FFI.Data.Maybe using (Maybe; just; nothing) -open import Properties.DecSubtyping using (dec-subtyping; resolve) +open import Properties.DecSubtyping using (dec-subtyping) open import Properties.Product using (_×_; _,_) +open import Properties.ResolveOverloads using (resolve) orUnknown : Maybe Type → Type orUnknown nothing = unknown diff --git a/prototyping/Properties/ResolveOverloads.agda b/prototyping/Properties/ResolveOverloads.agda new file mode 100644 index 00000000..8c9bf204 --- /dev/null +++ b/prototyping/Properties/ResolveOverloads.agda @@ -0,0 +1,78 @@ +{-# OPTIONS --rewriting #-} + +module Properties.ResolveOverloads where + +open import FFI.Data.Either using (Left; Right) +open import Luau.Subtyping using (_<:_; _≮:_) +open import Luau.Type using (Type ; _⇒_; unknown; never) +open import Luau.TypeSaturation using (saturate) +open import Properties.Contradiction using (CONTRADICTION) +open import Properties.DecSubtyping using (dec-subtypingⁿ) +open import Properties.Functions using (_∘_) +open import Properties.Subtyping using (<:-refl; <:-trans; <:-∩-left; <:-∩-right; <:-∩-glb; <:-impl-¬≮:; <:-unknown; function-≮:-never) +open import Properties.TypeNormalization using (Normal; FunType; normal; _⇒_; _∩_; _∪_; never; unknown; <:-normalize) +open import Properties.TypeSaturation using (Overloads; Saturated; _⊆ᵒ_; normal-saturate; saturated; defn; here; left; right) + +data ResolvedTo F G R : Set where + + yes : ∀ Sʳ Tʳ → + + Overloads F (Sʳ ⇒ Tʳ) → + (R <: Sʳ) → + (∀ {S T} → Overloads G (S ⇒ T) → (R <: S) → (Tʳ <: T)) → + -------------------------------------------- + ResolvedTo F G R + + no : + + (∀ {S T} → Overloads G (S ⇒ T) → (R ≮: S)) → + -------------------------------------------- + ResolvedTo F G R + + never : + + (F <: never) → + -------------------------------------------- + ResolvedTo F G R + +Resolved : Type → Type → Set +Resolved F R = ResolvedTo F F R + +resolveˢ : ∀ {F G R} → FunType G → Saturated F → Normal R → (G ⊆ᵒ F) → ResolvedTo F G R +resolveˢ (Sⁿ ⇒ Tⁿ) (defn sat-∩ sat-∪) Rⁿ G⊆F with dec-subtypingⁿ Rⁿ Sⁿ +resolveˢ (Sⁿ ⇒ Tⁿ) (defn sat-∩ sat-∪) Rⁿ G⊆F | Left R≮:S = no (λ { here → R≮:S }) +resolveˢ (Sⁿ ⇒ Tⁿ) (defn sat-∩ sat-∪) Rⁿ G⊆F | Right R<:S = yes _ _ (G⊆F here) R<:S (λ { here _ → <:-refl }) +resolveˢ (Gᶠ ∩ Hᶠ) (defn sat-∩ sat-∪) Rⁿ G⊆F with resolveˢ Gᶠ (defn sat-∩ sat-∪) Rⁿ (G⊆F ∘ left) | resolveˢ Hᶠ (defn sat-∩ sat-∪) Rⁿ (G⊆F ∘ right) +resolveˢ (Gᶠ ∩ Hᶠ) (defn sat-∩ sat-∪) Rⁿ G⊆F | yes S₁ T₁ o₁ R<:S₁ tgt₁ | yes S₂ T₂ o₂ R<:S₂ tgt₂ with sat-∩ o₁ o₂ +resolveˢ (Gᶠ ∩ Hᶠ) (defn sat-∩ sat-∪) Rⁿ G⊆F | yes S₁ T₁ o₁ R<:S₁ tgt₁ | yes S₂ T₂ o₂ R<:S₂ tgt₂ | defn o p₁ p₂ = + yes _ _ o (<:-trans (<:-∩-glb R<:S₁ R<:S₂) p₁) (λ { (left o) p → <:-trans p₂ (<:-trans <:-∩-left (tgt₁ o p)) ; (right o) p → <:-trans p₂ (<:-trans <:-∩-right (tgt₂ o p)) }) +resolveˢ (Gᶠ ∩ Hᶠ) (defn sat-∩ sat-∪) Rⁿ G⊆F | yes S₁ T₁ o₁ R<:S₁ tgt₁ | no src₂ = + yes _ _ o₁ R<:S₁ (λ { (left o) p → tgt₁ o p ; (right o) p → CONTRADICTION (<:-impl-¬≮: p (src₂ o)) }) +resolveˢ (Gᶠ ∩ Hᶠ) (defn sat-∩ sat-∪) Rⁿ G⊆F | no src₁ | yes S₂ T₂ o₂ R<:S₂ tgt₂ = + yes _ _ o₂ R<:S₂ (λ { (left o) p → CONTRADICTION (<:-impl-¬≮: p (src₁ o)) ; (right o) p → tgt₂ o p }) +resolveˢ (Gᶠ ∩ Hᶠ) (defn sat-∩ sat-∪) Rⁿ G⊆F | no src₁ | no src₂ = + no (λ { (left o) → src₁ o ; (right o) → src₂ o }) +resolveˢ (Gᶠ ∩ Hᶠ) (defn sat-∩ sat-∪) Rⁿ G⊆F | _ | never q = never q +resolveˢ (Gᶠ ∩ Hᶠ) (defn sat-∩ sat-∪) Rⁿ G⊆F | never p | _ = never p + +resolveᶠ : ∀ {F R} → FunType F → Normal R → Resolved (saturate F) R +resolveᶠ Fᶠ Rⁿ = resolveˢ (normal-saturate Fᶠ) (saturated Fᶠ) Rⁿ (λ o → o) + +resolveⁿ : ∀ {F R} → Normal F → Normal R → Resolved (saturate F) R +resolveⁿ (Sⁿ ⇒ Tⁿ) Rⁿ = resolveᶠ (Sⁿ ⇒ Tⁿ) Rⁿ +resolveⁿ (Fᶠ ∩ Gᶠ) Rⁿ = resolveᶠ (Fᶠ ∩ Gᶠ) Rⁿ +resolveⁿ (Sⁿ ∪ Tˢ) Rⁿ = no (λ ()) +resolveⁿ never Rⁿ = never <:-refl +resolveⁿ unknown Rⁿ = no (λ ()) + +resolve : Type → Type → Type +resolve F R with resolveⁿ (normal F) (normal R) +resolve F R | yes S T o R<:S p = T +resolve F R | no p = unknown +resolve F R | never p = never + +<:-resolve : ∀ {R S T} → T <: resolve (S ⇒ T) R +<:-resolve {R} {S} {T} with resolveⁿ (normal (S ⇒ T)) (normal R) +<:-resolve {R} {S} {T} | yes Sⁿ Tⁿ here Rⁿ<:Sʳ p = <:-normalize T +<:-resolve {R} {S} {T} | no p = <:-unknown +<:-resolve {R} {S} {T} | never p = CONTRADICTION (<:-impl-¬≮: p function-≮:-never) diff --git a/prototyping/Properties/StrictMode.agda b/prototyping/Properties/StrictMode.agda index 32e8b169..53cbbaa7 100644 --- a/prototyping/Properties/StrictMode.agda +++ b/prototyping/Properties/StrictMode.agda @@ -23,10 +23,11 @@ open import Properties.Equality using (_≢_; sym; cong; trans; subst₁) open import Properties.Dec using (Dec; yes; no) open import Properties.Contradiction using (CONTRADICTION; ¬) open import Properties.Functions using (_∘_) -open import Properties.DecSubtyping using (dec-subtyping; resolve) +open import Properties.DecSubtyping using (dec-subtyping) open import Properties.Subtyping using (unknown-≮:; ≡-trans-≮:; ≮:-trans-≡; ≮:-trans; ≮:-refl; scalar-≢-impl-≮:; function-≮:-scalar; scalar-≮:-function; function-≮:-never; unknown-≮:-scalar; scalar-≮:-never; unknown-≮:-never) open import Properties.FunctionTypes using (src-unknown-≮:; unknown-src-≮:) -open import Properties.Subtyping using (unknown-≮:; ≡-trans-≮:; ≮:-trans-≡; ≮:-trans; ≮:-refl; scalar-≢-impl-≮:; function-≮:-scalar; scalar-≮:-function; function-≮:-never; unknown-≮:-scalar; scalar-≮:-never; unknown-≮:-never) +open import Properties.ResolveOverloads using (resolve; <:-resolve) +open import Properties.Subtyping using (unknown-≮:; ≡-trans-≮:; ≮:-trans-≡; ≮:-trans; <:-trans-≮:; ≮:-refl; scalar-≢-impl-≮:; function-≮:-scalar; scalar-≮:-function; function-≮:-never; unknown-≮:-scalar; scalar-≮:-never; unknown-≮:-never) open import Properties.TypeCheck using (typeOfᴼ; typeOfᴹᴼ; typeOfⱽ; typeOfᴱ; typeOfᴮ; typeCheckᴱ; typeCheckᴮ; typeCheckᴼ; typeCheckᴴ) open import Luau.OpSem using (_⟦_⟧_⟶_; _⊢_⟶*_⊣_; _⊢_⟶ᴮ_⊣_; _⊢_⟶ᴱ_⊣_; app₁; app₂; function; beta; return; block; done; local; subst; binOp₀; binOp₁; binOp₂; refl; step; +; -; *; /; <; >; ==; ~=; <=; >=; ··) open import Luau.RuntimeError using (BinOpError; RuntimeErrorᴱ; RuntimeErrorᴮ; FunctionMismatch; BinOpMismatch₁; BinOpMismatch₂; UnboundVariable; SEGV; app₁; app₂; bin₁; bin₂; block; local; return; +; -; *; /; <; >; <=; >=; ··) @@ -148,7 +149,7 @@ reflect-subtypingᴱ H (M $ N) (app₁ s) p | Right W = Right (app₁ W) reflect-subtypingᴱ H (M $ N) (app₂ v s) p with reflect-subtypingᴱ H N s (⇒-≮:-resolve⁻¹ (heap-weakeningᴱ ∅ H M (rednᴱ⊑ s) (resolve-≮:-⇒ p))) reflect-subtypingᴱ H (M $ N) (app₂ v s) p | Left q = Left (⇒-≮:-resolve (resolve⁻¹-≮:-⇒ q)) reflect-subtypingᴱ H (M $ N) (app₂ v s) p | Right W = Right (app₂ W) -reflect-subtypingᴱ H (M $ N) {T = T} (beta (function f ⟨ var y ∈ S ⟩∈ U is B end) v refl q) p = Left (≡-trans-≮: (cong (λ F → resolve F (typeOfᴱ H ∅ N)) (cong orUnknown (cong typeOfᴹᴼ q))) p) +reflect-subtypingᴱ H (M $ N) {T = T} (beta (function f ⟨ var y ∈ S ⟩∈ U is B end) v refl q) p = Left (≡-trans-≮: (cong (λ F → resolve F (typeOfᴱ H ∅ N)) (cong orUnknown (cong typeOfᴹᴼ q))) (<:-trans-≮: (<:-resolve {typeOfᴱ H ∅ N}) p)) reflect-subtypingᴱ H (function f ⟨ var x ∈ T ⟩∈ U is B end) (function a defn) p = Left p reflect-subtypingᴱ H (block var b ∈ T is B end) (block s) p = Left p reflect-subtypingᴱ H (block var b ∈ T is return (val v) ∙ B end) (return v) p = mapR BlockMismatch (swapLR (≮:-trans p)) diff --git a/prototyping/Properties/TypeCheck.agda b/prototyping/Properties/TypeCheck.agda index c304e8bf..12385a5a 100644 --- a/prototyping/Properties/TypeCheck.agda +++ b/prototyping/Properties/TypeCheck.agda @@ -17,10 +17,10 @@ open import Luau.Var using (Var; _≡ⱽ_) open import Luau.Heap using (Heap; Object; function_is_end) renaming (_[_] to _[_]ᴴ) open import Properties.Contradiction using (CONTRADICTION) open import Properties.Dec using (yes; no) -open import Properties.DecSubtyping using (resolve) open import Properties.Equality using (_≢_; sym; trans; cong) open import Properties.Product using (_×_; _,_) open import Properties.Remember using (Remember; remember; _,_) +open import Properties.ResolveOverloads using (resolve) typeOfᴼ : Object yes → Type typeOfᴼ (function f ⟨ var x ∈ S ⟩∈ T is B end) = (S ⇒ T)