diff --git a/cedar-lean/Cedar/Data/Map.lean b/cedar-lean/Cedar/Data/Map.lean index ce44ac842..21e5eb3dd 100644 --- a/cedar-lean/Cedar/Data/Map.lean +++ b/cedar-lean/Cedar/Data/Map.lean @@ -88,6 +88,23 @@ def size {α β} (m : Map α β) : Nat := def mapOnValues {α β γ} [LT α] [DecidableLT α] (f : β → γ) (m : Map α β) : Map α γ := Map.mk (m.kvs.map (λ (k, v) => (k, f v))) +def mapOnValuesAttach {α β γ} [LT α] [DecidableLT α] (m : Map α β) (f : {x : β // ∃ (k : α), (k,x) ∈ m.kvs} → γ) : Map α γ := + let wrapper (prod : {x // x ∈ m.kvs}) : {x // ∃ k, (k,x) ∈ m.kvs} := + let property : ∃ k, (k, prod.val.snd) ∈ m.kvs := by + have h := prod.property + cases heq : prod.val + rename_i fst snd + exists fst + simp + rw [← heq] + assumption + { val := prod.val.snd, property := property } + let f' (prod : {x // x ∈ m.kvs}) := + let key := prod.val.fst + let value := wrapper prod + (key, f value) + Map.mk (m.kvs.map₁ (λ p => f' p)) + def mapOnKeys {α β γ} [LT γ] [DecidableLT γ] (f : α → γ) (m : Map α β) : Map γ β := Map.make (m.kvs.map (λ (k, v) => (f k, v))) @@ -99,6 +116,14 @@ def mapMOnKeys {α β γ} [LT γ] [DecidableLT γ] [Monad m] (f : α → m γ) ( let kvs ← map.kvs.mapM (λ (k, v) => f k >>= λ k' => pure (k', v)) pure (Map.make kvs) + + +def Equiv {α β} [LT α] [DecidableLT α] [BEq α] (m₁ m₂ : Map α β) : Prop := + ∀ (k : α), + m₁.find? k = m₂.find? k + +infix:50 " ≃ " => Equiv + ----- Instances ----- instance [LT (Prod α β)] : LT (Map α β) where diff --git a/cedar-lean/Cedar/Spec/Value.lean b/cedar-lean/Cedar/Spec/Value.lean index b4207f436..4de21abf3 100644 --- a/cedar-lean/Cedar/Spec/Value.lean +++ b/cedar-lean/Cedar/Spec/Value.lean @@ -127,6 +127,19 @@ instance : Coe Value (Result (Data.Set Value)) where ----- Derivations ----- +instance : LawfulBEq Attr where + rfl := by + intros a + simp + eq_of_beq := by + intros a b h + unfold Attr at a + unfold Attr at b + apply LawfulBEq.eq_of_beq + assumption + + +deriving instance Repr, DecidableEq, BEq, LT for Attr deriving instance Repr, DecidableEq, BEq for Except deriving instance Repr, DecidableEq for Error deriving instance Repr, DecidableEq, Inhabited, Lean.ToJson for Name diff --git a/cedar-lean/Cedar/Thm.lean b/cedar-lean/Cedar/Thm.lean index 6a298decf..12e0492d0 100644 --- a/cedar-lean/Cedar/Thm.lean +++ b/cedar-lean/Cedar/Thm.lean @@ -18,3 +18,5 @@ import Cedar.Thm.Authorization import Cedar.Thm.Partial import Cedar.Thm.Slicing import Cedar.Thm.Typechecking +import Cedar.Thm.Entities +import Cedar.Thm.Validation diff --git a/cedar-lean/Cedar/Thm/Data/List/Lemmas.lean b/cedar-lean/Cedar/Thm/Data/List/Lemmas.lean index 57e536ba0..6686b9036 100644 --- a/cedar-lean/Cedar/Thm/Data/List/Lemmas.lean +++ b/cedar-lean/Cedar/Thm/Data/List/Lemmas.lean @@ -1279,4 +1279,23 @@ theorem mem_pmap_subtype := by induction as <;> simp [*] +theorem in_lists_means_smaller [SizeOf α] (x : α) (list : List α) + (h : x ∈ list) : + sizeOf x < sizeOf list + := by + cases list + case nil => + cases h + case cons head tail => + cases h + case _ => + simp + omega + case _ in_tail => + have step : sizeOf x < sizeOf tail := by + apply in_lists_means_smaller + assumption + simp + omega + end List diff --git a/cedar-lean/Cedar/Thm/Data/Map.lean b/cedar-lean/Cedar/Thm/Data/Map.lean index 3ca40b50b..18a6d83e7 100644 --- a/cedar-lean/Cedar/Thm/Data/Map.lean +++ b/cedar-lean/Cedar/Thm/Data/Map.lean @@ -559,6 +559,225 @@ theorem in_kvs_in_mapOnValues [LT α] [DecidableLT α] [DecidableEq α] {f : β simp only [kvs, List.mem_map, Prod.mk.injEq] exists (k, v) +/-- + We can remove the attach for the sake of proofs +-/ +theorem mapOnValuesAttachIsMapOnValues + {α : Type u} {β γ : Type v} [LT α] [DecidableLT α] + {m : Map α β} + {f : β → γ} : + m.mapOnValues f = m.mapOnValuesAttach (λ prod => f prod.val) + := by + rw [← eq_iff_kvs_eq] + simp [mapOnValues, mapOnValuesAttach] + rw [← List.map₁_eq_map] + +theorem mapOnValues_cons + {α : Type u} {β γ : Type v} [LT α] [DecidableLT α] [BEq α ] + {f : β → γ} + {kv : α × β} + {kvs : List (α × β)} + : + (Map.mk (kv :: kvs)).mapOnValues f = + Map.mk ((kv.fst, f kv.snd) :: ((Map.mk kvs).mapOnValues f).kvs) + := by + rw [← eq_iff_kvs_eq] + simp [mapOnValues, List.map] + +/-- + Keys are not effected by mapping on values + ie: the domain of a map is unchanged by map on values +-/ +theorem mapOnValuesAttach_preservesContains + {α : Type u} {β γ : Type v} [LT α] [DecidableLT α] [BEq α ] [LawfulBEq α] [DecidableEq α] + {m : Map α β} + {f : β → γ} + {k : α} : + (m.contains k ) = (m.mapOnValuesAttach (λ v => f v.val)).contains k + := by + rw [← mapOnValuesAttachIsMapOnValues] + cases hcontains : (m.contains k) <;> cases m <;> rename_i kvs + case true => + induction kvs + case nil => + simp [contains, find?, List.find?] at hcontains + case cons head tail ih => + cases head + rename_i key value + cases heq_head_key : decide (key = k) <;> simp at heq_head_key + case _ => + have beq_false : (key == k) = false := by + rw [beq_eq_false_iff_ne] + assumption + rw [mapOnValues_cons] + simp [contains, find?, List.find?, beq_false] + simp [contains, find?, List.find?] at ih + apply ih + simp [contains, find?, List.find?, beq_false] at hcontains + assumption + case _ => + subst heq_head_key + simp [contains, find?, List.find?] + case false => + induction kvs + case nil => + simp [contains, find?, List.find?] + case cons head tail ih => + cases head + rename_i key value + cases heq_head_key : decide (key = k) <;> simp at heq_head_key + case _ => + have beq_false : (key == k) = false := by + rw [beq_eq_false_iff_ne] + assumption + rw [mapOnValues_cons] + simp [contains, find?, List.find?, beq_false] + simp [contains, find?, List.find?] at ih + apply ih + simp [contains, find?, List.find?, beq_false] at hcontains + assumption + case _ => + subst heq_head_key + simp [contains, find?, List.find?] at hcontains + +/-- + An adapter that makes the above lemma easier to apply in context +-/ +theorem mapOnValuesAttach_preservesContains_adapter + {α : Type u} {β γ : Type v} [LT α] [DecidableLT α] [BEq α ] [LawfulBEq α] [DecidableEq α] + {m : Map α β} + {f : {x // ∃ k, (k,x) ∈ m.kvs} → γ} + {k : α} + {h₁ : ∃ (f' : β → γ), f = (λ prod => f' prod.val)} : + (m.contains k ) = (m.mapOnValuesAttach f).contains k + := by + replace ⟨f', h₁⟩ := h₁ + rw [h₁] + apply mapOnValuesAttach_preservesContains + +theorem mapOnValuesAttach_preservesKeys + {α : Type u} {β γ : Type v} [LT α] [DecidableLT α] [BEq α ] [LawfulBEq α] [DecidableEq α] + {m : Map α β} + {f : β → γ} + {k : α} + {h : m.contains k = true} : + (m.mapOnValuesAttach (λ v => f v.val)).contains k = true + := by + rw [← mapOnValuesAttachIsMapOnValues] + cases m + rename_i kvs + induction kvs + case nil => + simp [mapOnValues, List.map, contains, find?, kvs] at h + case cons head tail ih => + simp [contains, find?, kvs, List.find?] at h + simp [contains, find?, kvs, List.find?] + cases heq_head : (head.fst == k) + case true => + simp + case false => + simp + cases htail : (mk tail).contains k + case true => + have hrecur : (mapOnValues f (mk tail)).contains k = true := by + apply ih + assumption + simp [mapOnValues, contains, find?, kvs ] at hrecur + assumption + case false => + exfalso + rw [heq_head] at h + simp at h + simp [contains, find?, kvs] at htail + cases h' : List.find? (fun x => x.fst == k) tail + case none => + rw [h'] at h + simp at h + case some => + rw [h'] at htail + simp at htail + +theorem mapOnValuesAttach_preservesKeys_adapter + {α : Type u} {β γ : Type v} [LT α] [DecidableLT α] [BEq α ] [LawfulBEq α] [DecidableEq α] + {m : Map α β} + {f : { x // ∃ k, (k,x) ∈ m.kvs} → γ} + {k : α} + {h₁ : m.contains k = true} + {h₂ : ∃ (f' : β → γ), f = λ val => f' val.val } : + (m.mapOnValuesAttach f).contains k = true + := by + replace ⟨f', h₂⟩ := h₂ + rw [h₂] + apply mapOnValuesAttach_preservesKeys + assumption + +theorem mapOnValues_maps + {α : Type u} {β γ : Type v} [LT α] [DecidableLT α] [BEq α] [LawfulBEq α] [DecidableEq α] + {m : Map α β} + {f : β → γ} + {k : α} + {v : β} + {h₁ : m.find? k = some v} : + (m.mapOnValuesAttach (λ prod => f prod.val)).find? k = .some (f v) + := by + rw [← mapOnValuesAttachIsMapOnValues] + cases m + rename_i kvs + induction kvs + case nil => + simp [find?, List.find?] at h₁ + case cons head tail ih => + cases head + rename_i key value + cases heq : decide (key = k) <;> simp at heq + case _ => + have beq_false : (key == k) = false := by + apply beq_false_of_ne + assumption + simp [find?, List.find?, beq_false] at h₁ + rw [mapOnValues_cons] + simp [find?, List.find?, beq_false] + apply ih + split at h₁ <;> simp at h₁ + rename_i heq + subst h₁ + simp [find?, heq] + case _ => + subst heq + rw [mapOnValues_cons] + simp [find?, List.find?] at h₁ + simp [find?, List.find?] + subst h₁ + rfl + +theorem mapOnValues_maps_adapter + {α : Type u} {β γ : Type v} + [LT α] [DecidableLT α] [BEq α] [LawfulBEq α] [DecidableEq α] + {m : Map α β} + {f : {x // ∃ k, (k,x) ∈ m.kvs} → γ} + {f': β → γ} + {k : α} + {v : β} + {h₁ : m.find? k = some v} + {h₂ : f = λ prod => f' prod.val} : + (m.mapOnValuesAttach f).find? k = .some (f' v) + := by + rw [h₂] + apply mapOnValues_maps + assumption + +theorem mapOnValuesAttachFunEq + {α : Type u} {β γ : Type v} + [LT α] [DecidableLT α] [BEq α] + {m : Map α β} + {k : α} + {f₁ : {x // ∃ k, (k,x) ∈ m.kvs} → γ} + {f₂ : {x // ∃ k, (k,x) ∈ m.kvs} → γ} + {h₁ : f₁ = f₂} : + (m.mapOnValuesAttach f₁).find? k = (m.mapOnValuesAttach f₂).find? k + := by + rw [h₁] + /-- Converse of `in_kvs_in_mapOnValues`; requires the extra preconditions that `m` is `WellFormed` and `f` is injective @@ -934,4 +1153,83 @@ theorem mapMOnValues_error_implies_exists_error [LT α] [DecidableLT α] {f : β have h_values := in_list_in_values hkv exists v +/-! ### `sizeOf` -/ + +theorem find_means_mem + {α : Type u} {β : Type v} + [LT α] [DecidableLT α] [BEq α] [LawfulBEq α] [DecidableEq α] + {m : Map α β} + {k : α} + {v : β} + (h : m.find? k = some v) : + (k,v) ∈ m.kvs + := by + cases m + rename_i kvs + induction kvs + case nil => + simp [find?, List.find?] at h + case cons head tail ih => + simp [kvs] + cases head + rename_i key value + cases heq : decide (key = k) <;> simp at heq + case _ => + have beq : (key == k) = false := by + rw [beq_eq_false_iff_ne] + assumption + apply Or.inr + simp [kvs] at ih + apply ih + simp [find?, List.find?, beq] at h + simp [find?, List.find?] + apply h + case _ => + apply Or.inl + subst heq + simp [find?, List.find?] at h + subst h + rfl + +-- If you can find a value in a map, that value is smaller than the map +theorem find_means_smaller + {α β : Type} + [LT α] [DecidableLT α] [DecidableEq α] [SizeOf α] [SizeOf β] + {m : Map α β} + {k : α} + {v : β} + {h : m.find? k = some v} : + sizeOf v < sizeOf m := by + have h₂ : (k,v) ∈ m.kvs := by + apply find?_mem_toList + assumption + have s₁ : sizeOf v < sizeOf (k,v) := by + simp + omega + have s₂ : sizeOf m.kvs < sizeOf m := by apply sizeOf_lt_of_kvs + have s₃ : sizeOf (k,v) < sizeOf m.kvs := by + apply List.sizeOf_lt_of_mem + assumption + omega + +def equiv_refl [LT α] [DecidableLT α] [BEq α] : ∀ (m₁ m₂ : Map α β), + m₁ ≃ m₂ → + m₂ ≃ m₂ := by + simp [Equiv] + +def equiv_sym [LT α] [DecidableLT α] [BEq α] : ∀ (m : Map α β), + m ≃ m := by + simp [Equiv] + +def equiv_trans [LT α] [DecidableLT α] [BEq α] : ∀ (m₁ m₂ m₃ : Map α β ), + m₁ ≃ m₂ → + m₂ ≃ m₃ → + m₁ ≃ m₃ := by + simp [Equiv] + intros m₁ m₂ m₃ h₁ h₂ + intros k + simp [h₁, h₂] + + + end Cedar.Data.Map diff --git a/cedar-lean/Cedar/Thm/Data/Set.lean b/cedar-lean/Cedar/Thm/Data/Set.lean index b7b0ccfcd..2d00f461a 100644 --- a/cedar-lean/Cedar/Thm/Data/Set.lean +++ b/cedar-lean/Cedar/Thm/Data/Set.lean @@ -635,4 +635,46 @@ theorem union_subset_eq [LT α] [DecidableLT α] [StrictLT α] [DecidableEq α] · rw [union_comm] exact subset_union _ _ + +theorem elts_subset_of_constructor [LT α] [DecidableLT α] [StrictLT α] : + ∀ (as : List α), + (Set.make as).elts ⊆ as + := by + intros as + simp [make, elts] + apply List.canonicalize_subseteq + +theorem in_constructor_in_set [LT α] [DecidableLT α] [StrictLT α] : + ∀ (as : List α) (a : α), + a ∈ (Set.make as) → + a ∈ as + := by + intros as a mem + simp [make] at mem + simp [Membership.mem] at mem + simp [elts] at mem + apply (@List.in_canonicalize_in_list α α) + assumption + +/-! ### sizeOf -/ + +theorem in_set_means_smaller [SizeOf α] (s : Set α) (x : α) + (h : x ∈ s) : + sizeOf x < sizeOf s + := by + cases s + rename_i elts + have hin : x ∈ elts := by + rw [Set.in_list_iff_in_mk] + assumption + have step₁ : sizeOf x < sizeOf elts := by + apply List.in_lists_means_smaller + assumption + simp + omega + + + + + end Cedar.Data.Set diff --git a/cedar-lean/Cedar/Thm/Entities.lean b/cedar-lean/Cedar/Thm/Entities.lean new file mode 100644 index 000000000..184b9e2ce --- /dev/null +++ b/cedar-lean/Cedar/Thm/Entities.lean @@ -0,0 +1,5 @@ + +import Batteries.Data.List +namespace Cedar.Entities + +end Cedar.Entities diff --git a/cedar-lean/Cedar/Thm/Typechecking.lean b/cedar-lean/Cedar/Thm/Typechecking.lean index 7694867d0..49f4c3a3d 100644 --- a/cedar-lean/Cedar/Thm/Typechecking.lean +++ b/cedar-lean/Cedar/Thm/Typechecking.lean @@ -25,8 +25,8 @@ namespace Cedar.Thm open Cedar.Spec open Cedar.Validation -def typecheck (policy : Policy) (env : Environment) : Except TypeError CedarType := do - let (ty, _) ← typeOf policy.toExpr ∅ env +def typecheck (policy : Policy) (env : Environment) (l : Level) : Except TypeError CedarType := do + let (ty, _) ← typeOf policy.toExpr ∅ env (l == .infinite) if ty ⊑ .bool .anyBool then .ok ty else .error (.unexpectedType ty) @@ -40,14 +40,14 @@ these errors because it has no knowledge of the entities/context that will be provided at authorization time, and it does not reason about the semantics of arithmetic operators. -/ -theorem typecheck_is_sound (policy : Policy) (env : Environment) (t : CedarType) (request : Request) (entities : Entities) : +theorem typecheck_is_sound (policy : Policy) (env : Environment) (t : CedarType) (request : Request) (entities : Entities) (l : Level) : RequestAndEntitiesMatchEnvironment env request entities → - typecheck policy env = .ok t → + typecheck policy env l = .ok t → (∃ (b : Bool), EvaluatesTo policy.toExpr request entities b) := by intro h₁ h₂ simp [typecheck] at h₂ - cases h₃ : typeOf (Policy.toExpr policy) [] env <;> simp [h₃] at h₂ + cases h₃ : typeOf (Policy.toExpr policy) [] env (l == .infinite) <;> simp [h₃] at h₂ split at h₂ <;> simp at h₂ rename_i ht have hc := empty_capabilities_invariant request entities diff --git a/cedar-lean/Cedar/Thm/Validation.lean b/cedar-lean/Cedar/Thm/Validation.lean index 55eadac38..ad19f4ba8 100644 --- a/cedar-lean/Cedar/Thm/Validation.lean +++ b/cedar-lean/Cedar/Thm/Validation.lean @@ -17,8 +17,10 @@ import Cedar.Spec import Cedar.Data import Cedar.Validation +import Cedar.Thm.Validation.WellFormed import Cedar.Thm.Validation.Validator import Cedar.Thm.Validation.RequestEntityValidation +import Cedar.Thm.Validation.Levels /-! This file contains the top-level correctness properties for the Cedar validator. @@ -41,14 +43,14 @@ either produces a boolean value, or throws an error of type `entityDoesNotExist` information. -/ -theorem validation_is_sound (policies : Policies) (schema : Schema) (request : Request) (entities : Entities) : - validate policies schema = .ok () → - validateRequest schema request = .ok () → - validateEntities schema entities = .ok () → +theorem validation_is_sound (policies : Policies) (schema : Schema) (l : Level) (request : Request) (entities : Entities) : + validate policies schema l = .ok () → + validateRequest schema request l = .ok () → + validateEntities schema entities l = .ok () → AllEvaluateToBool policies request entities := by intro h₀ h₁ h₂ - have h₁ := request_and_entities_validate_implies_match_schema schema request entities h₁ h₂ + have h₁ := request_and_entities_validate_implies_match_schema schema request entities l h₁ h₂ unfold validate at h₀ simp only [AllEvaluateToBool] cases h₃ : policies with @@ -56,10 +58,10 @@ theorem validation_is_sound (policies : Policies) (schema : Schema) (request : R | cons h' t' => intro policy pin simp only [EvaluatesToBool] - apply typecheck_policy_with_environments_is_sound policy schema.toEnvironments request entities h₁ + apply typecheck_policy_with_environments_is_sound policy (schema.toEnvironments l) request entities l h₁ subst h₃ simp only [List.forM_cons'] at h₀ - cases h₄ : (typecheckPolicyWithEnvironments h' schema.toEnvironments) <;> + cases h₄ : (typecheckPolicyWithEnvironments l h' (schema.toEnvironments l)) <;> simp only [h₄, Except.bind_err] at h₀ case ok _ => rw [List.mem_cons] at pin @@ -68,7 +70,7 @@ theorem validation_is_sound (policies : Policies) (schema : Schema) (request : R subst h₅ assumption | inr h₅ => - apply List.forM_ok_implies_all_ok t' (λ x => typecheckPolicyWithEnvironments x schema.toEnvironments) + apply List.forM_ok_implies_all_ok t' (λ x => typecheckPolicyWithEnvironments l x (schema.toEnvironments l)) repeat assumption end Cedar.Thm diff --git a/cedar-lean/Cedar/Thm/Validation/Levels.lean b/cedar-lean/Cedar/Thm/Validation/Levels.lean new file mode 100644 index 000000000..9fe0830fc --- /dev/null +++ b/cedar-lean/Cedar/Thm/Validation/Levels.lean @@ -0,0 +1,194 @@ +import Cedar.Thm.Data.Map +import Cedar.Thm.Validation.Typechecker.Basic +import Cedar.Thm.Validation.Typechecker.And +import Cedar.Thm.Validation.Typechecker.GetAttr +import Cedar.Thm.Validation.Typechecker.HasAttr +import Cedar.Thm.Validation.Typechecker.Record +import Cedar.Thm.Validation.Typechecker.Set +import Cedar.Thm.Validation.Typechecker.Call +import Cedar.Thm.Validation.Typechecker +import Cedar.Thm.Validation.Levels.Soundness +import Cedar.Thm.Validation.Levels.Util + + +namespace Cedar.Thm + +open Cedar.Data +open Cedar.Spec +open Cedar.Validation + +-- Main soundness theorem of the leveled type system +-- If you type check with an non-finite level +-- And you provide a well formed store +-- Then the only erors you can get are arithmatic/extension errors +theorem type_of_is_sound_noninfinite {e : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} : + l < .infinite → + CapabilitiesInvariant c₁ request entities → + RequestAndEntitiesMatchEnvironmentLeveled env request entities l → + typeOf e c₁ env (l == .infinite) = .ok (ty, c₂) → + GuardedCapabilitiesInvariant e c₂ request entities ∧ + ∃ (v : Value), EvaluatesToLeveled e request entities v ∧ InstanceOfType v ty + := by + intros h_level_finite h_caps h_reqs_levels h_typeof + have h_reqs := request_and_entity_match_level_implies_regular h_reqs_levels + + have h_sound := (@ type_of_is_sound e) h_caps h_reqs h_typeof + have ⟨h_sound₁, v, _, _⟩ := h_sound + clear h_sound + cases e + case lit => + apply type_of_is_sound_noninfinite_lit + repeat assumption + case var x _ => + apply type_of_is_sound_noninfinite_var x h_caps h_reqs_levels h_typeof + case ite => + apply type_of_is_sound_noninfinite_ite + repeat assumption + case _ => + prove_ih type_of_is_sound_noninfinite + case _ => + prove_ih type_of_is_sound_noninfinite + case _ => + prove_ih type_of_is_sound_noninfinite + case and => + apply type_of_is_sound_noninfinite_and + repeat assumption + case _ => + prove_ih type_of_is_sound_noninfinite + case _ => + prove_ih type_of_is_sound_noninfinite + case or => + apply type_of_is_sound_noninfinite_or + repeat assumption + case _ => + prove_ih type_of_is_sound_noninfinite + case _ => + prove_ih type_of_is_sound_noninfinite + case getAttr => + apply type_of_is_sound_noninfinite_getAttr + repeat assumption + prove_ih type_of_is_sound_noninfinite + case unaryApp => + apply type_of_is_sound_noninfinite_unary + repeat assumption + prove_ih type_of_is_sound_noninfinite + case binaryApp => + apply type_of_is_sound_noninfinite_binary + repeat assumption + prove_ih type_of_is_sound_noninfinite + prove_ih type_of_is_sound_noninfinite + case hasAttr => + apply type_of_is_sound_noninfinite_hasAttr + repeat assumption + prove_ih type_of_is_sound_noninfinite + case set => + apply type_of_is_sound_noninfinite_set + repeat assumption + intros + prove_ih type_of_is_sound_noninfinite + case record => + apply type_of_is_sound_noninfinite_record + repeat assumption + intros + prove_ih type_of_is_sound_noninfinite + case call => + apply type_of_is_sound_noninfinite_call + repeat assumption + intros + prove_ih type_of_is_sound_noninfinite +termination_by sizeOf e +decreasing_by + all_goals simp_wf + all_goals (try omega) + case _ => + rename e = (.ite _ _ _) => heq + simp [heq] + omega + case _ => + rename e = (.ite _ _ _) => heq + simp [heq] + omega + case _ => + rename e = (.ite _ _ _) => heq + simp [heq] + omega + case _ => + rename e = (.and _ _) => heq + simp [heq] + omega + case _ => + rename e = (.and _ _) => heq + simp [heq] + omega + case _ => + rename e = (.or _ _) => heq + simp [heq] + omega + case _ => + rename e = (.or _ _) => heq + simp [heq] + omega + case _ => + rename e = .unaryApp _ _ => heq + simp [heq] + omega + case _ => + rename e = .binaryApp _ _ _ => heq + simp [heq] + omega + case _ => + rename e = .binaryApp _ _ _ => heq + simp [heq] + omega + case _ => + rename e = .getAttr _ _ => heq + simp [heq] + omega + case _ => + rename e = .hasAttr _ _ => heq + simp [heq] + omega + case _ => + rename List Expr => members + rename _ = Expr.set _ => heq + rename _ ∈ members => hin + rename Expr => e' + have size₁ : sizeOf e' < sizeOf members := by + apply List.in_lists_means_smaller + assumption + have size₂ : sizeOf members < sizeOf (Expr.set members) := by + simp + omega + rw [← heq] at size₂ + omega + case _ => + rename List (Attr × Expr) => members + rename _ = Expr.record members => heq + rename Expr => e' + rename Attr => a + have step₁ : sizeOf e' < sizeOf (a,e') := by + simp + omega + have step₂ : sizeOf (a,e') < sizeOf members := by + apply List.in_lists_means_smaller + assumption + have step₃ : sizeOf members < sizeOf (Expr.record members) := by + simp + omega + rw [← heq] at step₃ + omega + case _ => + rename List Expr => args + rename _ = Expr.call _ args => heq + rename ExtFun => xfn + rename Expr => e' + have step₁ : sizeOf e' < sizeOf args := by + apply List.in_lists_means_smaller + assumption + have step₂ : sizeOf args < sizeOf (Expr.call xfn args) := by + simp + omega + rw [← heq] at step₂ + omega + +end Cedar.Thm diff --git a/cedar-lean/Cedar/Thm/Validation/Levels/Soundness.lean b/cedar-lean/Cedar/Thm/Validation/Levels/Soundness.lean new file mode 100644 index 000000000..089545485 --- /dev/null +++ b/cedar-lean/Cedar/Thm/Validation/Levels/Soundness.lean @@ -0,0 +1,1640 @@ +import Cedar.Thm.Data.Map +import Cedar.Thm.Validation.Typechecker.Basic +import Cedar.Thm.Validation.Typechecker.And +import Cedar.Thm.Validation.Typechecker.GetAttr +import Cedar.Thm.Validation.Typechecker.HasAttr +import Cedar.Thm.Validation.Typechecker.Record +import Cedar.Thm.Validation.Typechecker.Set +import Cedar.Thm.Validation.Typechecker.Call +import Cedar.Thm.Validation.Typechecker +import Cedar.Thm.Validation.Levels.Util +import Cedar.Thm.Validation.Levels.WellFormed + +namespace Cedar.Thm + +open Cedar.Data +open Cedar.Spec +open Cedar.Validation + +theorem type_of_is_sound_noninfinite_lit (p : Prim) {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.lit p) c₁ env (l == .infinite) = .ok (ty, c₂)) : + GuardedCapabilitiesInvariant (.lit p) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.lit p) request entities v ∧ InstanceOfType v ty + := by + have hsound : GuardedCapabilitiesInvariant (.lit p) c₂ request entities ∧ ∃ v, EvaluatesTo (.lit p) request entities v ∧ InstanceOfType v ty := by + type_soundness + have ⟨hsound₁, v', hsound₂, hsound₃⟩ := hsound + clear hsound + constructor <;> try assumption + exists v' + constructor <;> try assumption + cases p <;> ( + simp [EvaluatesToLeveled] + inrr + cases hsound₂ <;> rename_i hsound₂ <;> simp [evaluate] at hsound₂ + rw [← hsound₂] + simp [hsound₂, evaluate]) + +theorem type_of_is_sound_noninfinite_var (x : Var) {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.var x) c₁ env (l == .infinite) = .ok (ty, c₂)) : + GuardedCapabilitiesInvariant (.var x) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.var x) request entities v ∧ InstanceOfType v ty := by + have hsound : GuardedCapabilitiesInvariant (.var x) c₂ request entities ∧ ∃ v, EvaluatesTo (.var x) request entities v ∧ InstanceOfType v ty := by + type_soundness + have ⟨hsound₁, v', hsound₂, hsound₃⟩ := hsound + clear hsound + constructor <;> try assumption + exists v' + constructor <;> try assumption + simp [EvaluatesToLeveled] + inrr + cases hsound₂ + <;> rename_i hsound₂ + <;> simp [evaluate] at hsound₂ + <;> cases x <;> simp at hsound₂ + <;> simp [evaluate, hsound₂] + +theorem type_of_is_sound_noninfinite_ite {cond cons alt : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.ite cond cons alt) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih₁ : TypeOfIsSoundLeveled cond) + (ih₂ : TypeOfIsSoundLeveled cons) + (ih₃ : TypeOfIsSoundLeveled alt) : + GuardedCapabilitiesInvariant (.ite cond cons alt) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.ite cond cons alt) request entities v ∧ InstanceOfType v ty := by + have hsound : GuardedCapabilitiesInvariant (.ite cond cons alt) c₂ request entities ∧ ∃ v, EvaluatesTo (.ite cond cons alt) request entities v ∧ InstanceOfType v ty := by + type_soundness + have ⟨hsound₁, v', hsound₂, hsound₃⟩ := hsound + clear hsound + constructor <;> try assumption + exists v' + constructor <;> try assumption + have hinv := type_of_ite_inversion h₄ + replace ⟨bty, c₁', ty₂, c₂', ty₃, c₃, hinv⟩ := hinv + cases bty + case tt => + simp at hinv + have ⟨hinv₁, hinv₂, hinv₃, hinv₄⟩ := hinv + clear hinv + subst hinv₄ + subst hinv₃ + simp [EvaluatesToLeveled] + have hsound₁ : GuardedCapabilitiesInvariant cond c₁' request entities ∧ ∃ v_cond, EvaluatesToLeveled cond request entities v_cond ∧ InstanceOfType v_cond (.bool .tt) := by + apply ih₁ + apply h₁ + apply h₂ + apply h₃ + apply hinv₁ + have ⟨hsound₁₁, v_cond, hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₂ + case _ hsound₁₂ => + inl + simp [evaluate, hsound₁₂, Result.as] + case _ hsound₁₂ => + cases hsound₁₂ + case _ hsound₁₂ => + inrl + simp [evaluate, hsound₁₂, Result.as] + case _ hsound₁₂ => + cases hsound₁₃ + rename_i bool hsound₁₃ + cases bool <;> simp [InstanceOfBoolType] at hsound₁₃ + have hsound₂ : GuardedCapabilitiesInvariant cons c₂' request entities ∧ ∃ v_cons, EvaluatesToLeveled cons request entities v_cons ∧ InstanceOfType v_cons ty := by + apply ih₂ + apply h₁ + apply capability_union_invariant + apply h₂ + apply hsound₁₁ + apply hsound₁₂ + apply h₃ + apply hinv₂ + have ⟨_, v_cons, hsound₂₂, hsound₂₃⟩ := hsound₂ + clear hsound₂ + cases hsound₂₂ + case _ hsound₂₂ => + inl + simp [evaluate, hsound₁₂, Result.as, Coe.coe, Value.asBool, hsound₂₂] + case _ hsound₂₂ => + cases hsound₂₂ + case _ hsound₂₂ => + inrl + simp [evaluate, hsound₁₂, Result.as, Coe.coe, Value.asBool, hsound₂₂] + case _ hsound₂₂ => + inrr + simp [evaluate, hsound₁₂, Result.as, Coe.coe, Value.asBool, hsound₂₂] + have heval : evaluate (.ite cond cons alt) request entities = .ok v_cons := by + simp [evaluate, hsound₁₂, Result.as, Coe.coe, Value.asBool, hsound₂₂] + dril_to_value hsound₂ heval + rfl + case ff => + simp at hinv + have ⟨hinv₁, hinv₂, hinv₃, hinv₄⟩ := hinv + clear hinv + subst hinv₄ + subst hinv₃ + have hsound₁ : GuardedCapabilitiesInvariant cond c₁' request entities ∧ ∃ v_cond, EvaluatesToLeveled cond request entities v_cond ∧ InstanceOfType v_cond (.bool .ff) := by + apply ih₁ + apply h₁ + apply h₂ + apply h₃ + apply hinv₁ + have ⟨_, v_cond, hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + simp [EvaluatesToLeveled] + cases hsound₁₂ + case _ hsound₁₂ => + inl + simp [evaluate, hsound₁₂, Result.as] + case _ hsound₁₂ => + cases hsound₁₂ + case _ hsound₁₂ => + inrl + simp [evaluate, hsound₁₂, Result.as] + case _ hsound₁₂ => + cases hsound₁₃ + rename_i bool hsound₁₃ + cases bool <;> simp [InstanceOfBoolType] at hsound₁₃ + have hsound₂ : GuardedCapabilitiesInvariant alt c₂ request entities ∧ ∃ v, EvaluatesToLeveled alt request entities v ∧ InstanceOfType v ty := by + apply ih₃ + apply h₁ + apply h₂ + apply h₃ + apply hinv₂ + have ⟨_, v_alt, hsound₂₂, hsound₂₃⟩ := hsound₂ + clear hsound₂ + cases hsound₂₂ + case _ hsound₂₂ => + inl + simp [evaluate, hsound₁₂, hsound₂₂, Result.as, Coe.coe, Value.asBool] + case _ hsound₂₂ => + cases hsound₂₂ + case _ hsound₂₂ => + inrl + simp [evaluate, hsound₁₂, hsound₂₂, Result.as, Coe.coe, Value.asBool] + case _ hsound₂₂ => + inrr + simp [evaluate, hsound₁₂, hsound₂₂, Result.as, Coe.coe, Value.asBool] + have heval : evaluate (.ite cond cons alt) request entities = .ok v_alt := by + simp [evaluate, hsound₁₂, hsound₂₂, Result.as, Coe.coe, Value.asBool] + dril_to_value hsound₂ heval + rfl + case anyBool => + simp at hinv + have ⟨hinv₁, hinv₂, hinv₃, _, _⟩ := hinv + have hsound₁ : GuardedCapabilitiesInvariant cond c₁' request entities ∧ ∃ v, EvaluatesToLeveled cond request entities v ∧ InstanceOfType v (.bool .anyBool) := by + apply ih₁ + repeat assumption + have ⟨hsound₁₁, v_cond, hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + simp [EvaluatesToLeveled] + cases hsound₁₂ + case _ hsound₁₂ => + inl + simp [evaluate, hsound₁₂, Result.as] + case _ hsound₁₂ => + cases hsound₁₂ + case _ hsound₁₂ => + inrl + simp [evaluate, hsound₁₂, Result.as] + case _ hsound₁₂ => + cases hsound₁₃ + rename_i bool hbol + cases bool + case true => + have hsound₂ : GuardedCapabilitiesInvariant cons c₂' request entities ∧ ∃ v, EvaluatesToLeveled cons request entities v ∧ InstanceOfType v ty₂ := by + apply ih₂ + apply h₁ + apply capability_union_invariant + apply h₂ + apply hsound₁₁ + apply hsound₁₂ + apply h₃ + apply hinv₂ + have ⟨_, v_cons, hsound₂₂, hsound₂₃⟩ := hsound₂ + clear hsound₂ + cases hsound₂₂ + case _ hsound₂₂ => + inl + simp [evaluate, hsound₂₂, hsound₁₂, Result.as, Coe.coe, Value.asBool] + case _ hsound₂₂ => + cases hsound₂₂ + case _ hsound₂₂ => + inrl + simp [evaluate, hsound₂₂, hsound₁₂, Result.as, Coe.coe, Value.asBool] + case _ hsound₂₂ => + inrr + have heval : evaluate (.ite cond cons alt) request entities = .ok v_cons := by + simp [evaluate, hsound₂₂, hsound₁₂, Result.as, Coe.coe, Value.asBool] + simp [heval] + dril_to_value hsound₂ heval + rfl + case false => + have hsound₂ : GuardedCapabilitiesInvariant alt c₃ request entities ∧ ∃ v, EvaluatesToLeveled alt request entities v ∧ InstanceOfType v ty₃ := by + apply ih₃ + apply h₁ + apply h₂ + apply h₃ + apply hinv₃ + have ⟨_, v_cons, hsound₂₂, hsound₂₃⟩ := hsound₂ + clear hsound₂ + cases hsound₂₂ + case _ hsound₂₂ => + inl + simp [evaluate, hsound₂₂, hsound₁₂, Result.as, Coe.coe, Value.asBool] + case _ hsound₂₂ => + cases hsound₂₂ + case _ hsound₂₂ => + inrl + simp [evaluate, hsound₂₂, hsound₁₂, Result.as, Coe.coe, Value.asBool] + case _ hsound₂₂ => + inrr + have heval : evaluate (.ite cond cons alt) request entities = .ok v_cons := by + simp [evaluate, hsound₂₂, hsound₁₂, Result.as, Coe.coe, Value.asBool] + simp [heval] + dril_to_value hsound₂ heval + rfl + +theorem type_of_is_sound_noninfinite_and {lhs rhs : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.and lhs rhs) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih₁ : TypeOfIsSoundLeveled lhs) + (ih₂ : TypeOfIsSoundLeveled rhs) : + GuardedCapabilitiesInvariant (.and lhs rhs) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.and lhs rhs) request entities v ∧ InstanceOfType v ty + := by + have hsound : GuardedCapabilitiesInvariant (.and lhs rhs) c₂ request entities ∧ ∃ v, EvaluatesTo (.and lhs rhs) request entities v ∧ InstanceOfType v ty := by + type_soundness + replace ⟨_, v, hsound, _⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + have hinv := type_of_and_inversion h₄ + replace ⟨bty₁, c₁', hinv₁, hinv⟩ := hinv + cases bty₁ + case tt => + simp at hinv + have ⟨bty₂, c₂', hinv₂, _⟩ := hinv + clear hinv + have hsound₁ : GuardedCapabilitiesInvariant lhs c₁' request entities ∧ ∃ v, EvaluatesToLeveled lhs request entities v ∧ InstanceOfType v (.bool .tt) := by + apply ih₁ + repeat assumption + have ⟨hsound₁₁, v_lhs, hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i bool hbool + cases bool <;> simp [InstanceOfBoolType] at hbool + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ <;> simp [evaluate, h₁₂, Result.as] + have hsound₂ : GuardedCapabilitiesInvariant rhs c₂' request entities ∧ ∃ v, EvaluatesToLeveled rhs request entities v ∧ InstanceOfType v (.bool bty₂) := by + apply ih₂ + apply h₁ + apply capability_union_invariant + apply h₂ + apply hsound₁₁ + apply h₁₂ + apply h₃ + apply hinv₂ + have ⟨_, v_rhs, hsound₂₂, hsound₂₃⟩ := hsound₂ + cases hsound₂₃ + clear hsound₂ + rcases hsound₂₂ with h₂₂ | h₂₂ | h₂₂ <;> simp [evaluate, h₂₂, Result.as, Coe.coe, Value.asBool, pure, Except.pure] + rename_i bool hb + rcases hsound with hs | hs | hs | hs <;> simp [evaluate, h₁₂, h₂₂, Result.as, Coe.coe, Value.asBool, pure, Except.pure] at hs + assumption + case ff => + simp at hinv + have ⟨hinv₁, hinv₂⟩ := hinv + clear hinv + subst hinv₁ + subst hinv₂ + have hsound₁ : GuardedCapabilitiesInvariant lhs c₁' request entities ∧ ∃ v, EvaluatesToLeveled lhs request entities v ∧ InstanceOfType v (.bool .ff) := by + apply ih₁ + repeat assumption + have ⟨_, v_lhs, hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i bool hbool + simp [EvaluatesToLeveled] + cases bool <;> simp [InstanceOfBoolType] at hbool + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ <;> simp [evaluate, h₁₂, Result.as, Coe.coe, Value.asBool] + rcases hsound with hs | hs | hs | hs <;> simp [evaluate, h₁₂, Result.as, Coe.coe, Value.asBool] at hs + assumption + case anyBool => + simp at hinv + have ⟨bty₂, c₂', hinv₂, _⟩ := hinv + have hsound₁ : GuardedCapabilitiesInvariant lhs c₁' request entities ∧ ∃ v, EvaluatesToLeveled lhs request entities v ∧ InstanceOfType v (.bool .anyBool) := by + apply ih₁ + repeat assumption + have ⟨hsound₁₁, v_lhs, hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i bool hbool + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ <;> simp [h₁₂, evaluate, Result.as, Coe.coe, Value.asBool] + cases bool + case true => + simp + have hsound₂ : GuardedCapabilitiesInvariant rhs c₂' request entities ∧ ∃ v, EvaluatesToLeveled rhs request entities v ∧ InstanceOfType v (.bool bty₂) := by + apply ih₂ + apply h₁ + apply capability_union_invariant + apply h₂ + apply hsound₁₁ + repeat assumption + have ⟨_, v_rhs, hsound₂₂, hsound₂₃⟩ := hsound₂ + cases hsound₂₃ + rename_i bool hbool + clear hsound₂ + rcases hsound₂₂ with h₂₂ | h₂₂ | h₂₂ <;> simp [h₂₂, pure, Except.pure] + rcases hsound with hs | hs | hs | hs <;> simp [evaluate, h₁₂, h₂₂, Result.as, Coe.coe, Value.asBool, pure, Except.pure] at hs + assumption + case false => + simp + rcases hsound with hs | hs | hs | hs <;> simp [evaluate, h₁₂, Result.as, Coe.coe, Value.asBool, pure, Except.pure] at hs + assumption + +theorem type_of_is_sound_noninfinite_or {lhs rhs : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.or lhs rhs) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih₁ : TypeOfIsSoundLeveled lhs) + (ih₂ : TypeOfIsSoundLeveled rhs) : + GuardedCapabilitiesInvariant (.or lhs rhs) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.or lhs rhs) request entities v ∧ InstanceOfType v ty + := by + have hsound : GuardedCapabilitiesInvariant (.or lhs rhs) c₂ request entities ∧ ∃ v, EvaluatesTo (.or lhs rhs) request entities v ∧ InstanceOfType v ty := by + type_soundness + replace ⟨_, v, hsound, _⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + have hinv := type_of_or_inversion h₄ + replace ⟨bty₁, c₁', hinv₁, hinv⟩ := hinv + have hsound₁ : GuardedCapabilitiesInvariant lhs c₁' request entities ∧ ∃ v, EvaluatesToLeveled lhs request entities v ∧ InstanceOfType v (.bool bty₁) := by + apply ih₁ + repeat assumption + have ⟨_, v_lhs, hsound₁₂, hsound₁₃⟩ := hsound₁ + cases hsound₁₃ + rename_i lhs_bool hsound₁₃ + cases bty₁ + case tt => + simp at hinv + have ⟨hinv₂, hinv₃⟩ := hinv + subst hinv₂ + subst hinv₃ + clear hinv + simp [EvaluatesToLeveled] + cases lhs_bool <;> simp [InstanceOfBoolType] at hsound₁₃ + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ <;> simp [evaluate, h₁₂, Result.as, Coe.coe, Value.asBool] + rename InstanceOfType v (.bool .tt) => hinst + cases hinst + rename_i bool hinst + cases bool <;> simp [InstanceOfBoolType] at hinst + rfl + case ff => + simp at hinv + replace ⟨bty₂, c₂', hinv⟩ := hinv + subst hinv + have hsound₂ : GuardedCapabilitiesInvariant rhs c₂ request entities ∧ ∃ v, EvaluatesToLeveled rhs request entities v ∧ InstanceOfType v (.bool bty₂) := by + apply ih₂ + repeat assumption + have ⟨_, v_rhs, hsound₂₂, hsound₂₃⟩ := hsound₂ + cases hsound₂₃ + rename_i rhs_bool hsound₂₃ + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ <;> simp [evaluate, h₁₂, Result.as, Coe.coe, Value.asBool] + cases lhs_bool <;> simp [InstanceOfBoolType] at hsound₁₃ + rcases hsound₂₂ with h₂₂ | h₂₂ | h₂₂ <;> simp [h₂₂, Result.as, Coe.coe, Value.asBool, pure, Except.pure] + rcases hsound with hs | hs | hs | hs <;> simp [evaluate, h₁₂, h₂₂, Result.as, Coe.coe, Value.asBool, pure, Except.pure] at hs + assumption + case anyBool => + simp at hinv + replace ⟨bty₂, c₂', hinv₂, hinv⟩ := hinv + clear hinv + have hsound₂ : GuardedCapabilitiesInvariant rhs c₂' request entities ∧ ∃ v, EvaluatesToLeveled rhs request entities v ∧ InstanceOfType v (.bool bty₂) := by + apply ih₂ + repeat assumption + have ⟨_, v_rhs, hsound₂₂, hsound₂₃⟩ := hsound₂ + clear hsound₂ + simp [EvaluatesToLeveled] + cases hsound₂₃ + rename_i rhs_bool hsound₂₃ + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ + <;> cases lhs_bool + <;> rcases hsound₂₂ with h₂₂ | h₂₂ | h₂₂ <;> simp [evaluate, h₂₂, h₁₂, Result.as, Value.asBool, Coe.coe, pure, Except.pure] + <;> rcases hsound with hs | hs | hs | hs <;> simp [evaluate, h₂₂, h₁₂, Result.as, Value.asBool, Coe.coe, pure, Except.pure] at hs + <;> assumption + +theorem type_of_is_sound_noninfinite_getAttr {e : Expr} {a : Attr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (e.getAttr a) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih : TypeOfIsSoundLeveled e) : + GuardedCapabilitiesInvariant (e.getAttr a) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (e.getAttr a) request entities v ∧ InstanceOfType v ty := by + have hsound : GuardedCapabilitiesInvariant (e.getAttr a) c₂ request entities ∧ ∃ v, EvaluatesTo (e.getAttr a) request entities v ∧ InstanceOfType v ty := by + type_soundness + have ⟨_, v, hsound, _⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + have hinv := type_of_getAttr_inversion_levels h₄ + replace ⟨_, c₁', hinv⟩ := hinv + cases hinv + case _ hinv => + -- Entity Case + replace ⟨ety, l', hinv₁, hinv₂⟩ := hinv + clear hinv + have hsound₁ : GuardedCapabilitiesInvariant e c₁' request entities ∧ ∃ v, EvaluatesToLeveled e request entities v ∧ InstanceOfType v (.entity ety l') := by + apply ih + repeat assumption + have ⟨_, v', hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i euid hsound₁₃ + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ <;> simp [evaluate, h₁₂, Result.as] + inrr + simp [getAttr] + have wf : entities ⊢ (.prim (.entityUID euid)) : (.entity ety l') := by + apply evaluates_to_well_formed + repeat assumption + cases wf + case _ wf => + cases hinv₂ + omega + case _ => + rename_i attrs h₄ h₃ h₂ h₁ + simp [attrsOf, h₃] + rcases hsound with hs | hs | hs | hs + <;> simp [evaluate, attrsOf, h₃, h₁₂, getAttr, Map.findOrErr] at hs + <;> cases hfind : attrs.find? a + <;> simp [hfind] at hs + subst hs + simp [Map.findOrErr, hfind] + case _ hinv => + -- Record Case + replace ⟨rty, hinv⟩ := hinv + have hsound₁ : GuardedCapabilitiesInvariant e c₁' request entities ∧ ∃ v, EvaluatesToLeveled e request entities v ∧ InstanceOfType v (.record rty) := by + apply ih + repeat assumption + have ⟨_, v', hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i rv h₁ h₂ h₃ + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ <;> simp [evaluate, h₁₂, Result.as, getAttr, attrsOf] + inrr + rcases hsound with hs | hs | hs | hs <;> simp [evaluate, h₁₂, Result.as, getAttr, attrsOf] at hs + <;> cases hfind : rv.find? a + <;> simp [hfind, Map.findOrErr] at hs + subst hs + simp [Map.findOrErr, hfind] + +theorem type_of_is_sound_noninfinite_unary {e : Expr} {o : UnaryOp} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.unaryApp o e) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih : TypeOfIsSoundLeveled e) : + GuardedCapabilitiesInvariant (.unaryApp o e) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.unaryApp o e) request entities v ∧ InstanceOfType v ty + := by + have hsound : GuardedCapabilitiesInvariant (.unaryApp o e) c₂ request entities ∧ ∃ v, EvaluatesTo (.unaryApp o e) request entities v ∧ InstanceOfType v ty := by + type_soundness + replace ⟨_, v, hsound, _⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + cases o + case _ => + have hinv := type_of_not_inversion h₄ + replace ⟨_, bty, c₁', _, hinv⟩ := hinv + have hsound₁ : GuardedCapabilitiesInvariant e c₁' request entities ∧ ∃ v, EvaluatesToLeveled e request entities v ∧ InstanceOfType v (.bool bty) := by + apply ih + repeat assumption + have ⟨_, v', hsound₁₂, hsound₂₃⟩ := hsound₁ + clear hsound₁ + cases hsound₂₃ + rename_i bool hbool + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ <;> simp [evaluate, h₁₂, Result.as] + inrr + simp [apply₁] + rcases hsound with hs | hs | hs | hs <;> simp [evaluate, h₁₂, Result.as, apply₁] at hs + assumption + case _ => + have hinv := type_of_neg_inversion h₄ + + replace ⟨_, _, c₁', hinv⟩ := hinv + have hsound₁ : GuardedCapabilitiesInvariant e c₁' request entities ∧ ∃ v, EvaluatesToLeveled e request entities v ∧ InstanceOfType v (.int) := by + apply ih + repeat assumption + have ⟨_, v', hsound₁₂, hsound₂₃⟩ := hsound₁ + clear hsound₁ + cases hsound₂₃ + rename_i i + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ <;> simp [evaluate, h₁₂, Result.as] + simp [apply₁] + simp [intOrErr] + rcases hsound with hs | hs | hs | hs + <;> simp [evaluate, h₁₂, Result.as, apply₁] at hs + <;> cases hneg : i.neg? + <;> simp [hneg, intOrErr] + <;> simp [hneg, intOrErr] at hs + <;> assumption + case _ => + have hinv := type_of_like_inversion h₄ + replace ⟨_, _, c₁', hinv⟩ := hinv + have hsound₁ : GuardedCapabilitiesInvariant e c₁' request entities ∧ ∃ v, EvaluatesToLeveled e request entities v ∧ InstanceOfType v .string := by + apply ih + repeat assumption + have ⟨_, v', hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i s + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ <;> simp [evaluate, h₁₂, Result.as] + inrr + simp [apply₁] + rcases hsound with hs | hs | hs | hs <;> simp [evaluate, h₁₂, Result.as, apply₁] at hs + assumption + case _ => + rename_i ety_rhs _ + have hinv := type_of_is_inversion h₄ + replace ⟨_, _, l, c₁', hinv⟩ := hinv + rename EntityType => ety_lhs + rename InstanceOfType v ty => hinst + cases heq : decide (ety_rhs = ety_lhs) <;> simp at heq + case _ => + rw [if_neg] at hinv + have ⟨hinv₁, hinv₂⟩ := hinv + clear hinv + subst hinv₁ + have hsound₁ : GuardedCapabilitiesInvariant e c₁' request entities ∧ ∃ v, EvaluatesToLeveled e request entities v ∧ InstanceOfType v (.entity ety_lhs l) := by + apply ih + repeat assumption + have ⟨_, v', hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i euid hsound₁₃ + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ <;> simp [evaluate, h₁₂, Result.as] + inrr + simp [InstanceOfEntityType] at hsound₁₃ + simp [apply₁, heq] + rw [← hsound₁₃] + have hneq : (ety_rhs == ety_lhs) = false := by + exact beq_false_of_ne heq + simp [hneq] + cases hinst + rename_i bool hbool + cases bool <;> simp [InstanceOfBoolType] at hbool + rfl + assumption + case _ => + subst heq + simp at hinv + have ⟨hinv₁, hivnv₂⟩ := hinv + clear hinv + subst hinv₁ + have hsound₁ : GuardedCapabilitiesInvariant e c₁' request entities ∧ ∃ v, EvaluatesToLeveled e request entities v ∧ InstanceOfType v (.entity ety_rhs l) := by + apply ih + repeat assumption + have ⟨_, v', hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i euid hsound₁₃ + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ <;> simp [evaluate, h₁₂, Result.as] + inrr + simp [InstanceOfEntityType] at hsound₁₃ + simp [apply₁, hsound₁₃] + cases hinst + rename_i bool hinst + cases bool <;> simp [InstanceOfBoolType] at hinst + rfl + +macro "singleton_boolean" hinst:(ident) : tactic => + `(tactic | ( + have hinst := $hinst + cases hinst + rename_i bool hinst + cases bool <;> simp [InstanceOfBoolType] at hinst + )) + +theorem type_of_is_sound_noninfinite_binary_eq {lhs rhs : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.binaryApp .eq lhs rhs) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih₁ : TypeOfIsSoundLeveled lhs) + (ih₂ : TypeOfIsSoundLeveled rhs) : + GuardedCapabilitiesInvariant (.binaryApp .eq lhs rhs) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.binaryApp .eq lhs rhs) request entities v ∧ InstanceOfType v ty := by + have hsound : GuardedCapabilitiesInvariant (.binaryApp .eq lhs rhs) c₂ request entities ∧ ∃ v, EvaluatesTo (.binaryApp .eq lhs rhs) request entities v ∧ InstanceOfType v ty := by + type_soundness + replace ⟨_, v, hsound, _⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + have hinv := type_of_eq_inversion h₄ + replace ⟨_, hinv⟩ := hinv + split at hinv + case _ hinv => + rename_i p₁ p₂ _ _ + simp [EvaluatesToLeveled] + inrr + rename InstanceOfType v ty => hinst + cases heq : decide (p₁ = p₂) + case _ => + simp [evaluate, apply₂, heq, BEq.beq] + simp at heq + rw [if_neg] at hinv + subst hinv + cases hinst + rename_i bool hinst + cases bool <;> simp [InstanceOfBoolType] at hinst + rfl + assumption + case _ => + simp at heq + subst heq + simp [evaluate, apply₂] + simp at hinv + subst hinv + cases hinst + rename_i bool hinst + cases bool <;> simp [InstanceOfBoolType] at hinst + rfl + case _ hinv => + replace ⟨ty₁, c₁', ty₂, c₂', hinv₁, hinv₂, hinv⟩ := hinv + split at hinv + case _ _ => + have hsound₁ : GuardedCapabilitiesInvariant lhs c₁' request entities ∧ ∃ v, EvaluatesToLeveled lhs request entities v ∧ InstanceOfType v ty₁ := by + apply ih₁ + repeat assumption + have ⟨_, v_lhs, hsound₁₂, _⟩ := hsound₁ + clear hsound₁ + have hsound₂ : GuardedCapabilitiesInvariant rhs c₂' request entities ∧ ∃ v, EvaluatesToLeveled rhs request entities v ∧ InstanceOfType v ty₂ := by + apply ih₂ + repeat assumption + have ⟨_, v_rhs, hsound₂₂, _⟩ := hsound₂ + clear hsound₂ + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ + <;> rcases hsound₂₂ with h₂₂ | h₂₂ | h₂₂ + <;> simp [evaluate, h₁₂, h₂₂, Result.as, apply₂] + rcases hsound with hs | hs | hs | hs <;> simp [evaluate, h₁₂, h₂₂, Result.as, apply₂] at hs + assumption + case _ hinv => + replace ⟨hinv₃, ety₁, l₁, ety₂, l₂, hinv₄, hinv₅⟩ := hinv + subst hinv₄ + subst hinv₅ + have hsound₁ : GuardedCapabilitiesInvariant lhs c₁' request entities ∧ ∃ v, EvaluatesToLeveled lhs request entities v ∧ InstanceOfType v (.entity ety₁ l₁) := by + apply ih₁ + repeat assumption + have ⟨_, v_lhs, hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + have hsound₂ : GuardedCapabilitiesInvariant rhs c₂' request entities ∧ ∃ v, EvaluatesToLeveled rhs request entities v ∧ InstanceOfType v (.entity ety₂ l₂) := by + apply ih₂ + repeat assumption + have ⟨_, v_rhs, hsound₂₂, hsound₂₃⟩ := hsound₂ + clear hsound₂ + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ + <;> rcases hsound₂₂ with h₂₂ | h₂₂ | h₂₂ + <;> simp [evaluate, h₁₂, h₂₂, Result.as, apply₂] + have hneq : ¬ v_lhs = v_rhs := by + apply no_entity_type_lub_implies_not_eq + repeat assumption + have hnbeq : (v_lhs == v_rhs) = false := by + exact beq_false_of_ne hneq + simp [hnbeq] + rename InstanceOfType v ty => hinst + subst hinv₃ + cases hinst + rename_i bool hinst + cases bool <;> simp [InstanceOfBoolType] at hinst + rfl + +theorem empty_set_decide [DecidableEq α] (a : α) : + ((Set.empty.contains a)) = false + := by + simp [Set.empty, Set.contains] + unfold Not + intros h + simp [Set.elts] at h + +theorem map_or_err_err (set : Set Value) (e₁ e₂ : Error) + (h : Set.mapOrErr Value.asEntityUID set e₁ = .error e₂) : + e₁ = e₂ := by + simp [Set.mapOrErr] at h + cases h₁ : List.mapM Value.asEntityUID set.elts + <;> simp [h₁] at h + assumption + +theorem type_of_is_sound_noninfinite_binary_mem {lhs rhs : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.binaryApp .mem lhs rhs) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih₁ : TypeOfIsSoundLeveled lhs) + (ih₂ : TypeOfIsSoundLeveled rhs) : + GuardedCapabilitiesInvariant (.binaryApp .mem lhs rhs) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.binaryApp .mem lhs rhs) request entities v ∧ InstanceOfType v ty := by + have hsound : GuardedCapabilitiesInvariant (.binaryApp .mem lhs rhs) c₂ request entities ∧ ∃ v, EvaluatesTo (.binaryApp .mem lhs rhs) request entities v ∧ InstanceOfType v ty := by + type_soundness + replace ⟨_, v, hsound, _⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + have hinv := type_of_mem_inversion_finite h₄ + replace ⟨_, ety₁, ety₂, l₁, l₂, _, hinv₂, c₂', hinv⟩ := hinv + cases hinv + case _ hinv => -- single entity case + replace ⟨c₁', hinv₂⟩ := hinv₂ + have ⟨hinv₃, _⟩ := hinv + clear hinv + have hsound₁ : GuardedCapabilitiesInvariant lhs c₁' request entities ∧ ∃ v, EvaluatesToLeveled lhs request entities v ∧ InstanceOfType v (.entity ety₁ l₁) := by + apply ih₁ + repeat assumption + have ⟨_, v_lhs, hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i euid_lhs hsound₁₃ + simp [InstanceOfEntityType] at hsound₁₃ + have hsound₂ : GuardedCapabilitiesInvariant rhs c₂' request entities ∧ ∃ v, EvaluatesToLeveled rhs request entities v ∧ InstanceOfType v (.entity ety₂ l₂) := by + apply ih₂ + repeat assumption + have ⟨_, v_lhs, hsound₂₂, hsound₂₃⟩ := hsound₂ + clear hsound₂ + cases hsound₂₃ + rename_i euid_rhs hsound₂₃ + simp [InstanceOfEntityType] at hsound₂₃ + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ + <;> rcases hsound₂₂ with h₂₂ | h₂₂ | h₂₂ + <;> simp [evaluate, h₁₂, h₂₂, Result.as] + inrr + simp [apply₂, inₑ, Entities.ancestorsOrEmpty] + rcases hsound with hs | hs | hs | hs + <;> simp [evaluate, h₁₂, h₂₂, apply₂, inₑ, Entities.ancestorsOrEmpty, empty_set_decide] at hs + assumption + case _ hinv => -- set of entities case + replace ⟨c₁', hinv₂⟩ := hinv₂ + have ⟨hinv₃, hinv₄⟩ := hinv + clear hinv + have hsound₁ : GuardedCapabilitiesInvariant lhs c₁' request entities ∧ ∃ v, EvaluatesToLeveled lhs request entities v ∧ InstanceOfType v (.entity ety₁ l₁) := by + apply ih₁ + repeat assumption + have ⟨_, v_lhs, hsound₁₂, hsound₁₃⟩ := hsound₁ + cases hsound₁₃ + rename_i euid_lhs hsound₁₃ + simp [InstanceOfEntityType] at hsound₁₃ + subst hsound₁₃ + clear hsound₁ + have hsound₂ : GuardedCapabilitiesInvariant rhs c₂' request entities ∧ ∃ v, EvaluatesToLeveled rhs request entities v ∧ InstanceOfType v (.set (.entity ety₂ l₂)) := by + apply ih₂ + repeat assumption + have ⟨_, v_rhs, hsound₂₂, hsound₂₃⟩ := hsound₂ + clear hsound₂ + cases hsound₂₃ + rename_i set_rhs hsound₁₃ + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ + <;> rcases hsound₂₂ with h₂₂ | h₂₂ | h₂₂ + <;> simp [evaluate, h₁₂, h₂₂] + inrr + rcases hsound with hs | hs | hs | hs + case _ => + simp [evaluate, h₁₂, h₂₂, apply₂, inₛ] at hs + cases hset : Set.mapOrErr Value.asEntityUID set_rhs Error.typeError + case _ err => + have heq : .typeError = err := by + apply map_or_err_err + apply hset + subst heq + simp [hset] at hs + case _ => + simp [hset] at hs + case _ => + simp [evaluate, h₁₂, h₂₂, apply₂, inₛ] at hs + cases hset : Set.mapOrErr Value.asEntityUID set_rhs Error.typeError + case _ err => + have heq : .typeError = err := by + apply map_or_err_err + apply hset + subst heq + simp [hset] at hs + case _ => + simp [hset] at hs + case _ => + simp [evaluate, h₁₂, h₂₂, apply₂, inₛ] at hs + cases hset : Set.mapOrErr Value.asEntityUID set_rhs Error.typeError + case _ err => + have heq : .typeError = err := by + apply map_or_err_err + apply hset + subst heq + simp [hset] at hs + case _ => + simp [hset] at hs + case _ => + simp [evaluate, h₁₂, h₂₂] at hs + assumption + +theorem type_of_is_sound_noninfinite_binary_int_cmp {lhs rhs : Expr} {o : BinaryOp} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (hₒ : o = .less ∨ o = .lessEq) + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.binaryApp o lhs rhs) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih₁ : TypeOfIsSoundLeveled lhs) + (ih₂ : TypeOfIsSoundLeveled rhs) : + GuardedCapabilitiesInvariant (.binaryApp o lhs rhs) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.binaryApp o lhs rhs) request entities v ∧ InstanceOfType v ty := by + have hsound : GuardedCapabilitiesInvariant (.binaryApp o lhs rhs) c₂ request entities ∧ ∃ v, EvaluatesTo (.binaryApp o lhs rhs) request entities v ∧ InstanceOfType v ty := by + type_soundness + replace ⟨_, v, hsound, hinst⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + + have hinv := type_of_int_cmp_inversion hₒ h₄ + have ⟨_, hinv₂, hinv₃, hinv₄⟩ := hinv + clear hinv + subst hinv₂ + replace ⟨c₁', hinv₃⟩ := hinv₃ + replace ⟨c₂', hinv₄⟩ := hinv₄ + have hsound₁ : GuardedCapabilitiesInvariant lhs c₁' request entities ∧ ∃ v, EvaluatesToLeveled lhs request entities v ∧ InstanceOfType v .int := by + apply ih₁ + repeat assumption + have ⟨_, v_lhs, hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i int_lhs + have hsound₂ : GuardedCapabilitiesInvariant rhs c₂' request entities ∧ ∃ v, EvaluatesToLeveled rhs request entities v ∧ InstanceOfType v .int := by + apply ih₂ + repeat assumption + have ⟨_, v_rhs, hsound₂₂, hsound₂₃⟩ := hsound₂ + clear hsound₂ + cases hsound₂₃ + rename_i int_rhs + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ + <;> rcases hsound₂₂ with h₂₂ | h₂₂ | h₂₂ + <;> simp [evaluate, h₁₂, h₂₂] + inrr + rcases hsound with hs | hs | hs | hs + <;> cases hₒ + <;> rename_i hₒ + <;> simp [evaluate, h₁₂, h₂₂, apply₂, hₒ] at hs + <;> simp [apply₂, hₒ] + <;> assumption + +theorem type_of_is_sound_noninfinite_binary_int_arith {lhs rhs : Expr} {o : BinaryOp} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (hₒ : o = .add ∨ o = .sub ∨ o = .mul) + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.binaryApp o lhs rhs) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih₁ : TypeOfIsSoundLeveled lhs) + (ih₂ : TypeOfIsSoundLeveled rhs) : + GuardedCapabilitiesInvariant (.binaryApp o lhs rhs) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.binaryApp o lhs rhs) request entities v ∧ InstanceOfType v ty := by + have hsound : GuardedCapabilitiesInvariant (.binaryApp o lhs rhs) c₂ request entities ∧ ∃ v, EvaluatesTo (.binaryApp o lhs rhs) request entities v ∧ InstanceOfType v ty := by + type_soundness + replace ⟨_, v, hsound, _⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + have hinv := type_of_int_arith_inversion hₒ h₄ + have ⟨_, hinv₂, hinv₃, hinv₄⟩ := hinv + clear hinv + subst hinv₂ + replace ⟨c₁', hinv₃⟩ := hinv₃ + replace ⟨c₂', hinv₄⟩ := hinv₄ + have hsound₁ : GuardedCapabilitiesInvariant lhs c₁' request entities ∧ ∃ v, EvaluatesToLeveled lhs request entities v ∧ InstanceOfType v .int := by + apply ih₁ + repeat assumption + have ⟨_, v_lhs, hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i i_lhs + have hsound₂ : GuardedCapabilitiesInvariant rhs c₂' request entities ∧ ∃ v, EvaluatesToLeveled rhs request entities v ∧ InstanceOfType v .int := by + apply ih₂ + repeat assumption + have ⟨_, v_rhs, hsound₂₂, hsound₂₃⟩ := hsound₂ + clear hsound₂ + cases hsound₂₃ + rename_i i_rhs + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ + <;> rcases hsound₂₂ with h₂₂ | h₂₂ | h₂₂ + <;> rcases hₒ with hₒ | hₒ | hₒ + <;> subst hₒ + <;> simp [evaluate, h₁₂, h₂₂] + <;> simp [apply₂] + case _ => + cases h : (i_lhs.add? i_rhs) <;> simp [intOrErr] + rcases hsound with hs | hs | hs | hs + <;> simp [evaluate, h₁₂, h₂₂, apply₂, h, intOrErr] at hs + assumption + case _ => + cases h : (i_lhs.sub? i_rhs) <;> simp [intOrErr] + rcases hsound with hs | hs | hs | hs + <;> simp [evaluate, h₁₂, h₂₂, apply₂, h, intOrErr] at hs + assumption + case _ => + cases h : (i_lhs.mul? i_rhs) <;> simp [intOrErr] + rcases hsound with hs | hs | hs | hs + <;> simp [evaluate, h₁₂, h₂₂, apply₂, h, intOrErr] at hs + assumption + +theorem type_of_is_sound_noninfinite_binary_contains {lhs rhs : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.binaryApp .contains lhs rhs) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih₁ : TypeOfIsSoundLeveled lhs) + (ih₂ : TypeOfIsSoundLeveled rhs) : + GuardedCapabilitiesInvariant (.binaryApp .contains lhs rhs) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.binaryApp .contains lhs rhs) request entities v ∧ InstanceOfType v ty := by + have hsound : GuardedCapabilitiesInvariant (.binaryApp .contains lhs rhs) c₂ request entities ∧ ∃ v, EvaluatesTo (.binaryApp .contains lhs rhs) request entities v ∧ InstanceOfType v ty := by + type_soundness + replace ⟨_, v, hsound, hinst⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + have hinv := type_of_contains_inversion h₄ + have ⟨_, _, ty₁, ty₂, _, hinv₄, hinv₅⟩ := hinv + clear hinv + replace ⟨c₁', hinv₄⟩ := hinv₄ + replace ⟨c₂', hinv₅⟩ := hinv₅ + + have hsound₁ : GuardedCapabilitiesInvariant lhs c₁' request entities ∧ ∃ v, EvaluatesToLeveled lhs request entities v ∧ InstanceOfType v (.set ty₁) := by + apply ih₁ + repeat assumption + have ⟨_, v_lhs, hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i set_lhs hsound₁₃ + + have hsound₂ : GuardedCapabilitiesInvariant rhs c₂' request entities ∧ ∃ v, EvaluatesToLeveled rhs request entities v ∧ InstanceOfType v ty₂ := by + apply ih₂ + repeat assumption + have ⟨_, v_rhs, hsound₂₂, _⟩ := hsound₂ + clear hsound₂ + + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ + <;> rcases hsound₂₂ with h₂₂ | h₂₂ | h₂₂ + <;> simp [evaluate, h₁₂, h₂₂, apply₂] + <;> cases hcontains : (set_lhs.contains v_rhs) + <;> rcases hsound with hs | hs | hs | hs + <;> simp [evaluate, h₁₂, h₂₂, apply₂, hcontains] at hs + <;> assumption + +theorem type_of_is_sound_noninfinite_binary_containsA {lhs rhs : Expr} {o : BinaryOp} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (hₒ : o = .containsAll ∨ o = .containsAny) + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.binaryApp o lhs rhs) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih₁ : TypeOfIsSoundLeveled lhs) + (ih₂ : TypeOfIsSoundLeveled rhs) : + GuardedCapabilitiesInvariant (.binaryApp o lhs rhs) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.binaryApp o lhs rhs) request entities v ∧ InstanceOfType v ty := by + have hsound : GuardedCapabilitiesInvariant (.binaryApp o lhs rhs) c₂ request entities ∧ ∃ v, EvaluatesTo (.binaryApp o lhs rhs) request entities v ∧ InstanceOfType v ty := by + type_soundness + replace ⟨_, v, hsound, _⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + + have hinv := type_of_containsA_inversion hₒ h₄ + have ⟨_, _, ty₁, ty₂, _, hinv₄, hinv₅⟩ := hinv + replace ⟨c₁', hinv₄⟩ := hinv₄ + replace ⟨c₂', hinv₅⟩ := hinv₅ + + have hsound₁ : GuardedCapabilitiesInvariant lhs c₁' request entities ∧ ∃ v, EvaluatesToLeveled lhs request entities v ∧ InstanceOfType v (.set ty₁) := by + apply ih₁ + repeat assumption + have ⟨_, v_lhs, hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i set_lhs hsound₁₃ + + have hsound₂ : GuardedCapabilitiesInvariant rhs c₂' request entities ∧ ∃ v, EvaluatesToLeveled rhs request entities v ∧ InstanceOfType v (.set ty₂) := by + apply ih₂ + repeat assumption + have ⟨_, v_rhs, hsound₂₂, hsound₂₃⟩ := hsound₂ + clear hsound₂ + cases hsound₂₃ + rename_i set_rhs hsound₂₃ + + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ + <;> rcases hsound₂₂ with h₂₂ | h₂₂ | h₂₂ + <;> cases hₒ + <;> rename_i hₒ + <;> subst hₒ + <;> simp [evaluate, h₁₂, h₂₂, apply₂, ] + case _ => + cases h : set_rhs.subset set_lhs + <;>rcases hsound with hs | hs | hs | hs + <;> simp [evaluate, h₁₂, h₂₂, apply₂, h] at hs + <;> assumption + case _ => + cases h : set_rhs.intersects set_lhs + <;>rcases hsound with hs | hs | hs | hs + <;> simp [evaluate, h₁₂, h₂₂, apply₂, h] at hs + <;> assumption + +theorem type_of_is_sound_noninfinite_binary {lhs rhs : Expr} {o : BinaryOp} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.binaryApp o lhs rhs) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih₁ : TypeOfIsSoundLeveled lhs) + (ih₂ : TypeOfIsSoundLeveled rhs) : + GuardedCapabilitiesInvariant (.binaryApp o lhs rhs) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.binaryApp o lhs rhs) request entities v ∧ InstanceOfType v ty := by + cases o + case eq => + apply type_of_is_sound_noninfinite_binary_eq + repeat assumption + case mem => + apply type_of_is_sound_noninfinite_binary_mem + repeat assumption + case less => + apply type_of_is_sound_noninfinite_binary_int_cmp + simp + repeat assumption + case lessEq => + apply type_of_is_sound_noninfinite_binary_int_cmp + simp + repeat assumption + case add => + apply type_of_is_sound_noninfinite_binary_int_arith + simp + repeat assumption + case sub => + apply type_of_is_sound_noninfinite_binary_int_arith + simp + repeat assumption + case mul => + apply type_of_is_sound_noninfinite_binary_int_arith + simp + repeat assumption + case contains => + apply type_of_is_sound_noninfinite_binary_contains + repeat assumption + case containsAll => + apply type_of_is_sound_noninfinite_binary_containsA + simp + repeat assumption + case containsAny => + apply type_of_is_sound_noninfinite_binary_containsA + simp + repeat assumption + +theorem type_of_is_sound_noninfinite_hasAttr {e : Expr} {a : Attr} {c₁ c₂ : Capabilities} {request : Request} {entities : Entities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.hasAttr e a) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih : TypeOfIsSoundLeveled e) : + GuardedCapabilitiesInvariant (.hasAttr e a) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.hasAttr e a) request entities v ∧ InstanceOfType v ty := by + have hsound : GuardedCapabilitiesInvariant (.hasAttr e a) c₂ request entities ∧ ∃ v, EvaluatesTo (.hasAttr e a) request entities v ∧ InstanceOfType v ty := by + type_soundness + replace ⟨_, v, hsound, _⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + + have hinv := type_of_hasAttr_inversion h₄ + replace ⟨_, c₁', hinv⟩ := hinv + cases hinv + case _ hinv => + have ⟨ety, l', hinv₁⟩ := hinv + clear hinv + have hsound₁ : GuardedCapabilitiesInvariant e c₁' request entities ∧ ∃ v, EvaluatesToLeveled e request entities v ∧ InstanceOfType v (.entity ety l') := by + apply ih + repeat assumption + have ⟨_, v', hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i euid hsound₁₃ + simp [InstanceOfEntityType] at hsound₁₃ + subst hsound₁₃ + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ + <;> simp [evaluate, h₁₂] + inrr + simp [hasAttr, attrsOf] + rcases hsound with hs | hs | hs | hs + <;> simp [evaluate, h₁₂, hasAttr, attrsOf] at hs + assumption + case _ hinv => + replace ⟨rty, hinv⟩ := hinv + have hsound₁ : GuardedCapabilitiesInvariant e c₁' request entities ∧ ∃ v, EvaluatesToLeveled e request entities v ∧ InstanceOfType v (.record rty) := by + apply ih + repeat assumption + have ⟨_, v', hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ + <;> simp [evaluate, h₁₂, hasAttr, attrsOf] + rcases hsound with hs | hs | hs | hs + <;> simp [evaluate, h₁₂, hasAttr, attrsOf] at hs + assumption + +theorem type_of_is_sound_noninfinite_set_helper {es : List Expr} {c₁ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : ∀ (e), e ∈ es → ∃ ty' c', typeOf e c₁ env (l == .infinite) = .ok (ty', c')) + (ih : ∀ (e), e ∈ es → TypeOfIsSoundLeveled e) : + ∃ vs, + ( + es.mapM (λ pair => evaluate pair request entities) = .ok vs ∨ + es.mapM (λ pair => evaluate pair request entities) = .error Error.arithBoundsError ∨ + es.mapM (λ pair => evaluate pair request entities) = .error Error.extensionError + ) := by + + + cases es + case nil => + exists [] + simp [List.mapM, List.mapM.loop, pure, Except.pure] + case cons head tail => + rw [List.mapM_cons] + have typed : ∃ ty' c', typeOf head c₁ env (l == .infinite) = .ok (ty', c') := by + apply h₄ + simp + replace ⟨ty', c', typed⟩ := typed + + have hsound : GuardedCapabilitiesInvariant head c' request entities ∧ ∃ v, EvaluatesToLeveled head request entities v ∧ InstanceOfType v ty' := by + apply ih + simp + repeat assumption + replace ⟨_, v_head, hsound⟩ := hsound + have ih' := (@type_of_is_sound_noninfinite_set_helper tail) h₁ h₂ h₃ + (by + intros e hin + apply h₄ + simp [hin] + ) + (by + intros e hin + apply ih + simp [hin] + ) + replace ⟨vs_tail, ih'⟩ := ih' + rcases hsound with hs | hs | hs + <;> rcases ih' with ih' | ih' | ih' + <;> simp [evaluate, hs, ih', pure, Except.pure] + +theorem type_of_is_sound_noninfinite_set {es : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.set es) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih : ∀ e, e ∈ es → TypeOfIsSoundLeveled e) : + GuardedCapabilitiesInvariant (.set es) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.set es) request entities v ∧ InstanceOfType v ty := by + + have hsound : GuardedCapabilitiesInvariant (.set es) c₂ request entities ∧ ∃ v, EvaluatesTo (.set es) request entities v ∧ InstanceOfType v ty := by + type_soundness + replace ⟨_, v, hsound, _⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + + have hinv := type_of_set_inversion h₄ + have ⟨_, ty₁, _, hinv₃⟩ := hinv + clear hinv + + have helper : ∃ vs, ( + es.mapM (λ e => evaluate e request entities) = .ok vs ∨ + es.mapM (λ e => evaluate e request entities) = .error Error.arithBoundsError ∨ + es.mapM (λ e => evaluate e request entities) = .error Error.extensionError + ) := by + apply type_of_is_sound_noninfinite_set_helper + apply h₁ + apply h₂ + apply h₃ + intros e hin + have ⟨ty', c', h', _⟩ := hinv₃ e hin + exists ty' + exists c' + apply ih + replace ⟨vs, helper⟩ := helper + rcases helper with helper | helper | helper + <;> simp [EvaluatesToLeveled, evaluate, List.mapM₁, List.attach, List.attachWith, + List.mapM_pmap_subtype (λ e => evaluate e request entities), + helper + ] + <;> rcases hsound with hs | hs | hs | hs + <;> simp [EvaluatesToLeveled, evaluate, List.mapM₁, List.attach, List.attachWith, + List.mapM_pmap_subtype (λ e => evaluate e request entities), + helper + ] at hs + <;> assumption + +theorem type_of_is_sound_noninfinite_record_helper {members : List (Attr × Expr)} {c₁ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : ∀ a e, (a, e) ∈ members → ∃ ty' c', typeOf e c₁ env (l == .infinite) = .ok (ty', c')) + (ih : ∀ a e, (a, e) ∈ members → TypeOfIsSoundLeveled e) : + ∃ vs, + ( + members.mapM (λ pair => bindAttr pair.fst (evaluate pair.snd request entities)) = .ok vs ∨ + members.mapM (λ pair => bindAttr pair.fst (evaluate pair.snd request entities)) = .error Error.arithBoundsError ∨ + members.mapM (λ pair => bindAttr pair.fst (evaluate pair.snd request entities)) = .error Error.extensionError + ) + := by + cases members + case nil => + exists [] + simp [List.mapM, List.mapM.loop, pure, Except.pure] + case cons head tail => + rw [List.mapM_cons] + have ⟨attr, expr⟩ := head + have typed : ∃ ty' c', typeOf expr c₁ env (l == .infinite) = .ok (ty', c') := by + apply h₄ attr expr + simp + replace ⟨ty', c', typed⟩ := typed + have hsound : GuardedCapabilitiesInvariant expr c' request entities ∧ ∃ v, EvaluatesToLeveled expr request entities v ∧ InstanceOfType v ty' := by + apply ih + simp + inl + rfl + repeat assumption + replace ⟨_, v, hsound, _⟩ := hsound + have ih' := (@type_of_is_sound_noninfinite_record_helper tail) h₁ h₂ h₃ + (by + intros a e hin + apply h₄ + simp + apply Or.inr + assumption + ) + (by + intros a e hin + apply ih + simp + apply Or.inr + apply hin + ) + + replace ⟨vs, ih'⟩ := ih' + simp [bindAttr] at ih' + + rcases hsound with hs | hs | hs + <;> rcases ih' with ih' | ih' | ih' + <;> simp [bindAttr, evaluate, hs, ih', pure, Except.pure] + +theorem get_type (a : Attr) (e : Expr) {members : List (Attr × Expr)} {rty : List (Attr × QualifiedType)} {c₁ : Capabilities} {env : Environment} {l : Level} + (h₁ : List.Forall₂ (AttrExprHasAttrType c₁ env l) members rty) + (h₂ : (a,e) ∈ members) : + ∃ ty c, typeOf e c₁ env (l == .infinite) = .ok (ty, c) := by + cases members + case nil => + cases h₂ + case cons head tail => + cases h₁ + rename_i attrtype type_tail h_head h_tail + simp [AttrExprHasAttrType] at h_head + cases h₂ + case _ => + have ⟨ty', _, c', heq₂⟩ := h_head + exists ty' + exists c' + case _ hin => + apply get_type a e + apply h_tail + apply hin + +theorem type_of_is_sound_noninfinite_record {members: List (Attr × Expr)} {c₁ c₂ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {ty : CedarType} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.record members) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih : ∀ e a, (a,e) ∈ members → TypeOfIsSoundLeveled e) : + GuardedCapabilitiesInvariant (.record members) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.record members) request entities v ∧ InstanceOfType v ty := by + + have hsound : GuardedCapabilitiesInvariant (.record members) c₂ request entities ∧ ∃ v, EvaluatesTo (.record members) request entities v ∧ InstanceOfType v ty := by + type_soundness + replace ⟨_, v, hsound, _⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + + have hinv := type_of_record_inversion h₄ + have ⟨_, rty, _, hinv₂⟩ := hinv + clear hinv + simp [EvaluatesToLeveled, evaluate, List.mapM₂, List.attach₂, List.mapM_pmap_subtype (λ (pair : (Attr × Expr)) => bindAttr pair.fst (evaluate pair.snd request entities))] + have helper : ∃ vs, ( + members.mapM (λ pair => bindAttr pair.fst (evaluate pair.snd request entities)) = .ok vs ∨ + members.mapM (λ pair => bindAttr pair.fst (evaluate pair.snd request entities)) = .error Error.arithBoundsError ∨ + members.mapM (λ pair => bindAttr pair.fst (evaluate pair.snd request entities)) = .error Error.extensionError + ) := by + apply type_of_is_sound_noninfinite_record_helper + apply h₁ + apply h₂ + apply h₃ + intros a e hin + apply get_type a e + apply hinv₂ + apply hin + intros a e hin + apply ih + apply hin + replace ⟨vs, helper⟩ := helper + rcases helper with helper | helper | helper + <;> simp [helper] + <;> rcases hsound with hs | hs | hs | hs + <;> simp [EvaluatesToLeveled, helper, evaluate, List.mapM₂, List.attach₂, List.mapM_pmap_subtype (λ (pair : (Attr × Expr)) => bindAttr pair.fst (evaluate pair.snd request entities))] at hs + assumption + +theorem type_of_is_sound_noninfinite_call_decimal {es : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {ty : CedarType} {l : Level} + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.call .decimal es) c₁ env (l == .infinite) = .ok (ty, c₂)) : + GuardedCapabilitiesInvariant (.call .decimal es) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.call .decimal es) request entities v ∧ + InstanceOfType v ty + := by + have hsound : GuardedCapabilitiesInvariant (.call .decimal es) c₂ request entities ∧ ∃ v, EvaluatesTo (.call .decimal es) request entities v ∧ InstanceOfType v ty := by + type_soundness + replace ⟨_, v, hsound, _⟩ := hsound + + constructor <;> try assumption + exists v + constructor <;> try assumption + + have hinv := type_of_call_decimal_inversion h₄ + have ⟨_, _, decimal_src, hinv₃⟩ := hinv + simp [hinv₃, EvaluatesToLeveled, evaluate, List.mapM₁, List.attach, List.attachWith, List.mapM_pmap_subtype, call, res] + cases hdec : Ext.Decimal.decimal decimal_src <;> simp [hdec] + rcases hsound with hs | hs | hs | hs + <;> simp [hinv₃, EvaluatesToLeveled, evaluate, List.mapM₁, List.attach, List.attachWith, List.mapM_pmap_subtype, call, res, hdec] at hs + assumption + +theorem type_of_is_sound_noninfinite_call_ip {es : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {ty : CedarType} {l : Level} + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.call .ip es) c₁ env (l == .infinite) = .ok (ty, c₂)) : + GuardedCapabilitiesInvariant (.call .ip es) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.call .ip es) request entities v ∧ + InstanceOfType v ty + := by + have hsound : GuardedCapabilitiesInvariant (.call .ip es) c₂ request entities ∧ ∃ v, EvaluatesTo (.call .ip es) request entities v ∧ InstanceOfType v ty := by + type_soundness + replace ⟨_, v, hsound, _⟩ := hsound + + constructor <;> try assumption + exists v + constructor <;> try assumption + + have hinv := type_of_call_ip_inversion h₄ + have ⟨_, _, ip_src, hinv₃⟩ := hinv + simp [hinv₃, EvaluatesToLeveled, evaluate, List.mapM₁, List.attach, List.attachWith, List.mapM_pmap_subtype, call, res] + cases hdec : Ext.IPAddr.ip ip_src <;> simp [hdec] + rcases hsound with hs | hs | hs | hs + <;> simp [hinv₃, EvaluatesToLeveled, evaluate, List.mapM₁, List.attach, List.attachWith, List.mapM_pmap_subtype, call, res, hdec] at hs + assumption + +theorem type_of_is_sound_noninfinite_call_dec_compare {xfn : ExtFun} {es : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {ty : CedarType} {l : Level} + (hc : IsDecimalComparator xfn) + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.call xfn es) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih : ∀ e, e ∈ es → TypeOfIsSoundLeveled e) : + GuardedCapabilitiesInvariant (.call xfn es) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.call xfn es) request entities v ∧ + InstanceOfType v ty + := by + have hsound : GuardedCapabilitiesInvariant (.call xfn es) c₂ request entities ∧ + ∃ v, EvaluatesTo (.call xfn es) request entities v ∧ + InstanceOfType v ty + := by + type_soundness + replace ⟨_, v, hsound, _⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + + have hinv := type_of_call_decimal_comparator_inversion hc h₄ + have ⟨_, _, e₁, e₂, c₁', c₂', hinv₃, hinv₄, hinv₅⟩ := hinv + simp [hinv₃] + clear hinv + + have hsound₁ : GuardedCapabilitiesInvariant e₁ c₁' request entities ∧ ∃ v, EvaluatesToLeveled e₁ request entities v ∧ InstanceOfType v (.ext ExtType.decimal) := by + apply ih + simp [hinv₃] + repeat assumption + have ⟨_, v_lhs, hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i extval hsound₁₃ + simp [InstanceOfExtType] at hsound₁₃ + cases extval <;> simp at hsound₁₃ + + have hsound₂ : GuardedCapabilitiesInvariant e₂ c₂' request entities ∧ ∃ v, EvaluatesToLeveled e₂ request entities v ∧ InstanceOfType v (.ext ExtType.decimal) := by + apply ih + simp [hinv₃] + repeat assumption + have ⟨_, v_rhs, hsound₂₂, hsound₂₃⟩ := hsound₂ + clear hsound₂ + cases hsound₂₃ + rename_i extval hsound₂₃ + simp [InstanceOfExtType] at hsound₂₃ + cases extval <;> simp at hsound₂₃ + + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ + <;> rcases hsound₂₂ with h₂₂ | h₂₂ | h₂₂ + <;> cases xfn + <;> simp [IsDecimalComparator] at hc + <;> simp [evaluate, h₁₂, h₂₂, List.mapM₁, List.attach, List.attachWith, List.mapM_pmap_subtype (λ e => evaluate e request entities), call] + <;> rcases hsound with hs | hs | hs | hs + <;> simp [evaluate, hinv₃, h₁₂, h₂₂, List.mapM₁, List.attach, List.attachWith, List.mapM_pmap_subtype (λ e => evaluate e request entities), call] at hs + <;> assumption + +theorem type_of_is_sound_noninfinite_call_ip_recognizer {xfn : ExtFun} {es : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {ty : CedarType} {l : Level} + (hc : IsIpAddrRecognizer xfn) + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.call xfn es) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih : ∀ e, e ∈ es → TypeOfIsSoundLeveled e) : + GuardedCapabilitiesInvariant (.call xfn es) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.call xfn es) request entities v ∧ + InstanceOfType v ty + := by + have hsound : GuardedCapabilitiesInvariant (.call xfn es) c₂ request entities ∧ + ∃ v, EvaluatesTo (.call xfn es) request entities v ∧ + InstanceOfType v ty + := by + type_soundness + replace ⟨_, v, hsound, _⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + + have hinv := type_of_call_ipAddr_recognizer_inversion hc h₄ + have ⟨_, _, e₁, c₁', hinv₃, hinv₄⟩ := hinv + subst hinv₃ + clear hinv + + + have hsound₁ : GuardedCapabilitiesInvariant e₁ c₁' request entities ∧ ∃ v, EvaluatesToLeveled e₁ request entities v ∧ InstanceOfType v (CedarType.ext (ExtType.ipAddr)) := by + apply ih + simp + repeat assumption + have ⟨_, _, hsound₁₂, hsound₁₃⟩ := hsound₁ + clear hsound₁ + cases hsound₁₃ + rename_i extval hsound₁₃ + cases extval <;> simp [InstanceOfExtType] at hsound₁₃ + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ + <;> rcases hsound with hs | hs | hs | hs + <;> cases xfn + <;> simp [IsIpAddrRecognizer] at hc + <;> simp [evaluate, List.mapM₁, List.attach, List.attachWith, List.mapM_pmap_subtype, h₁₂, call] + <;> simp [evaluate, List.mapM₁, List.attach, List.attachWith, List.mapM_pmap_subtype, h₁₂, call] at hs + <;> assumption + +theorem type_of_is_sound_noninfinite_call_isInRange {es : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {ty : CedarType} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.call .isInRange es) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih : ∀ e, e ∈ es → TypeOfIsSoundLeveled e) : + GuardedCapabilitiesInvariant (.call .isInRange es) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.call .isInRange es) request entities v ∧ + InstanceOfType v ty + := by + have hsound : GuardedCapabilitiesInvariant (.call .isInRange es) c₂ request entities ∧ + ∃ v, EvaluatesTo (.call .isInRange es) request entities v ∧ + InstanceOfType v ty + := by + type_soundness + replace ⟨_, v, hsound, _⟩ := hsound + constructor <;> try assumption + exists v + constructor <;> try assumption + + have hinv := type_of_call_isInRange_inversion h₄ + have ⟨_, _, e₁, e₂, c₁', c₂', hinv₃, hinv₄, hinv₅⟩ := hinv + clear hinv + subst hinv₃ + + have hsound₁ : GuardedCapabilitiesInvariant e₁ c₁' request entities ∧ ∃ v, EvaluatesToLeveled e₁ request entities v ∧ InstanceOfType v (.ext ExtType.ipAddr) := by + apply ih + simp + repeat assumption + have ⟨_, _, hsound₁₂, hsound₁₃⟩ := hsound₁ + cases hsound₁₃ + rename_i ext hsound₁₃ + cases ext <;> simp [InstanceOfExtType] at hsound₁₃ + clear hsound₁ + + have hsound₂ : GuardedCapabilitiesInvariant e₂ c₂' request entities ∧ ∃ v, EvaluatesToLeveled e₂ request entities v ∧ InstanceOfType v (.ext ExtType.ipAddr) := by + apply ih + simp + repeat assumption + have ⟨_, _, hsound₂₂, hsound₂₃⟩ := hsound₂ + cases hsound₂₃ + rename_i ext hsound₂₃ + cases ext <;> simp [InstanceOfExtType] at hsound₂₃ + + simp [EvaluatesToLeveled] + rcases hsound₁₂ with h₁₂ | h₁₂ | h₁₂ + <;> rcases hsound₂₂ with h₂₂ | h₂₂ | h₂₂ + <;> simp [evaluate, h₁₂, h₂₂, List.mapM₁, List.attach, List.attachWith, call] + <;> rcases hsound with hs | hs | hs | hs + <;> simp [evaluate, h₁₂, h₂₂, List.mapM₁, List.attach, List.attachWith, call] at hs + assumption + +theorem type_of_is_sound_noninfinite_call {xfn : ExtFun} {es : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {ty : CedarType} {l : Level} + (h₁ : l < .infinite) + (h₂ : CapabilitiesInvariant c₁ request entities) + (h₃ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l) + (h₄ : typeOf (.call xfn es) c₁ env (l == .infinite) = .ok (ty, c₂)) + (ih : ∀ e, e ∈ es → TypeOfIsSoundLeveled e) : + GuardedCapabilitiesInvariant (.call xfn es) c₂ request entities ∧ + ∃ v, EvaluatesToLeveled (.call xfn es) request entities v ∧ + InstanceOfType v ty + := by + cases xfn + case decimal => + apply type_of_is_sound_noninfinite_call_decimal + repeat assumption + case ip => + apply type_of_is_sound_noninfinite_call_ip + repeat assumption + case lessThanOrEqual => + apply type_of_is_sound_noninfinite_call_dec_compare + simp [IsDecimalComparator] + repeat assumption + case lessThan => + apply type_of_is_sound_noninfinite_call_dec_compare + simp [IsDecimalComparator] + repeat assumption + case greaterThan => + apply type_of_is_sound_noninfinite_call_dec_compare + simp [IsDecimalComparator] + repeat assumption + case greaterThanOrEqual => + apply type_of_is_sound_noninfinite_call_dec_compare + simp [IsDecimalComparator] + repeat assumption + case isIpv4 => + apply type_of_is_sound_noninfinite_call_ip_recognizer + simp [IsIpAddrRecognizer] + repeat assumption + case isIpv6 => + apply type_of_is_sound_noninfinite_call_ip_recognizer + simp [IsIpAddrRecognizer] + repeat assumption + case isLoopback => + apply type_of_is_sound_noninfinite_call_ip_recognizer + simp [IsIpAddrRecognizer] + repeat assumption + case isMulticast => + apply type_of_is_sound_noninfinite_call_ip_recognizer + simp [IsIpAddrRecognizer] + repeat assumption + case isInRange => + apply type_of_is_sound_noninfinite_call_isInRange + repeat assumption + +macro "prove_ih" ih:(ident) : tactic => + `(tactic | ( + simp [TypeOfIsSoundLeveled] + intros + apply $ih + repeat assumption)) diff --git a/cedar-lean/Cedar/Thm/Validation/Levels/Util.lean b/cedar-lean/Cedar/Thm/Validation/Levels/Util.lean new file mode 100644 index 000000000..0250677ca --- /dev/null +++ b/cedar-lean/Cedar/Thm/Validation/Levels/Util.lean @@ -0,0 +1,298 @@ +import Cedar.Thm.Data.Map +import Cedar.Thm.Validation.Typechecker.Basic +import Cedar.Thm.Validation.Typechecker.And +import Cedar.Thm.Validation.Typechecker.GetAttr +import Cedar.Thm.Validation.Typechecker.HasAttr +import Cedar.Thm.Validation.Typechecker.Record +import Cedar.Thm.Validation.Typechecker.Set +import Cedar.Thm.Validation.Typechecker.Call +import Cedar.Thm.Validation.Typechecker + +namespace Cedar.Thm + +open Cedar.Data +open Cedar.Spec +open Cedar.Validation + +theorem level_non_infite {l : Level} : + l < .infinite → + (l == .infinite) = false := by + intros h + cases h + simp + unfold Not + intros + contradiction + +theorem entity_find {entities : Entities} {attrs : Map Attr Value} {uid : EntityUID} + (h : entities.attrs uid = Except.ok attrs) : + ∃ e, entities.find? uid = some e + := by + unfold Entities.attrs at h + unfold Map.findOrErr at h + cases hfind : (Map.find? entities uid) <;> rw [hfind] at h <;> simp at h + rename_i val + exists val + +theorem lt_trans {l₁ l₂ l₃ : Level} + (h₁ : l₁ < l₂) + (h₂ : l₂ < l₃) : + l₁ < l₃ := by + cases h₁ <;> cases h₂ + case _ n₁ n₂ h₁ n₂' h₂ => + apply LevelLT.finite₁ + omega + case _ h₁ h₂ => + apply LevelLT.finite₂ + +theorem le_trans {l₁ l₂ l₃ : Level} + (h₁ : l₁ ≤ l₂) + (h₂ : l₂ ≤ l₃) : + l₁ ≤ l₃ := by + simp [LE.le] at h₁ + simp [LE.le] at h₂ + simp [LE.le] + cases h₁ <;> cases h₂ + case _ h₁ h₂ => + simp [h₁, h₂] + case _ h₁ h₂ => + simp [h₁, h₂] + case _ h₁ h₂ => + subst h₂ + simp [h₁] + case _ h₁ h₂ => + apply Or.inr + apply lt_trans + repeat assumption + +theorem le_infinite {l : Level} : + .infinite ≤ l → + l = .infinite := by + intros h + cases l <;> try rfl + cases h + +theorem zero_le_all : ∀ (l : Level), + .finite 0 ≤ l := by + intros l + simp [LE.le] + cases l + case _ => + apply Or.inr + apply LevelLT.finite₂ + case _ n => + cases n + case _ => + simp [Level.finite] + case _ => + apply Or.inr + apply LevelLT.finite₁ + omega + +theorem le_lift : ∀ (n₁ n₂ : Nat), + n₁ ≤ n₂ → + Level.finite n₁ ≤ .finite n₂ := by + intros n₁ n₂ h + cases heq : decide (n₁ = n₂) <;> simp at heq + case _ => + have lt : n₁ < n₂ := by omega + simp [LE.le] + apply Or.inr + apply LevelLT.finite₁ + assumption + case _ => + subst heq + simp [LE.le] + +theorem all_le_infinit : ∀ (l : Level), + l ≤ .infinite := by + intros l + simp [LE.le] + cases l + case _ => + simp [Level.infinite] + case _ => + apply Or.inr + apply LevelLT.finite₂ + +theorem lub_not_record : ∀ rty₁ rty₂ ty, + (.record rty₁ ⊔ .record rty₂) = some ty → + (match ty with + | .record _ => false + | _ =>true + ) = true → + False + := by + intros rty₁ rty₂ ty h₁ h₂ + unfold lub? at h₁ + cases hlub : lubRecordType rty₁.1 rty₂.1 + case _ => + simp [hlub] at h₁ + case _ => + simp [hlub] at h₁ + subst h₁ + simp at h₂ + +-- macro "solve_var" v:(ident) l₁:(ident) l₂:(ident) : tactic => do +macro "not_record" : tactic => do + `(tactic | ( + exfalso + apply lub_not_record + assumption + simp + )) + +theorem qualitype_lub_lifts {qty₁ qty₂ qty : QualifiedType} + (h : lubQualifiedType qty₁ qty₂ = some qty) : + (qty₁.getType ⊔ qty₂.getType) = some qty.getType + := by + cases qty₁ <;> cases qty₂ + <;> simp [lubQualifiedType, Option.bind] at h + <;> rename_i ty₁ ty₂ + <;> cases hlub : (ty₁ ⊔ ty₂) <;> simp [hlub] at h + <;> subst h + <;> simp [Qualified.getType] + <;> assumption + +theorem bounded_above {l₁ l₂ : Level} : + l₁ < .infinite → + l₁ ≥ l₂ → + l₂ < .infinite + := by + intros h₁ h₂ + cases l₁ <;> cases l₂ <;> try cases h₁ + cases h₂ + apply LevelLT.finite₂ + +theorem lt_zero_gt_zero_false {l : Level} + (h₁ : l > .finite 0) + (h₂ : l ≤ .finite 0) : + False := by + cases l + case none => -- l = infinity + cases h₂ + case some n => + simp [LE.le] at h₂ + cases h₂ + case _ heq => + simp [Level.finite] at heq + subst heq + cases h₁ + omega + case _ hlt => + cases h₁ + cases hlt + omega + +theorem lt_zero_helper {l₁ l₂ : Level} + (h₁ : l₂ > .finite 0) + (h₂ : l₁ ≥ l₂) : + l₁ ≠ .finite 0 + := by + simp at h₁ + simp [LE.le] at h₂ + cases h₂ + case _ h₂ => + subst h₂ + simp + intros h + subst h + cases h₁ + omega + case _ h₂ => + cases l₁ <;> cases l₂ + case _ => + simp + contradiction + case _ n₁ => + simp + intros h + contradiction + case _ n₂ => + cases h₂ + case _ n₁ n₂ => + cases h₂ + rename_i h₂ + cases h₁ + rename_i h₁ + have hzero : n₁ ≠ 0 := by omega + simp + intros contra + apply hzero + simp [Level.finite] at contra + assumption + +theorem gt_zero_helper + (l : Level) + (h₁ : l > .finite 0) : + l ≠ .finite 0 := by + cases l + case _ => + simp [Level.finite] + case _ n => + cases h₁ + simp [Level.finite] + omega + +theorem zero_contra {l : Level} + (h₁ : l > .finite 0) + (h₂ : l ≤ .finite 0) : + False + := by + simp [LE.le] at h₂ + cases h₂ + case _ h₂ => + subst h₂ + cases h₁ + omega + case _ h₂ => + simp at h₁ + cases l + case _ => + cases h₂ + case _ n => + cases h₂ + omega + +-- We repeat this proof a lot +macro "type_soundness" : tactic => + `(tactic | ( + apply type_of_is_sound + assumption + apply request_and_entity_match_level_implies_regular + assumption + repeat assumption + )) + +-- Repeat case analysis +macro "dril_to_value" hsound₂:(ident) h₅:(ident) : tactic => + `(tactic | ( + have hsound₂ := $hsound₂ + have h₅ := $h₅ + cases hsound₂ <;> rename_i hsound₂ <;> try simp [hsound₂] at h₅ + cases hsound₂ <;> rename_i hsound₂ <;> try simp [hsound₂] at h₅ + cases hsound₂ <;> rename_i hsound₂ <;> try simp [hsound₂] at h₅ + clear hsound₂ + subst h₅ + )) + +-- Lisp style car,caar,cddr macros but for ∨ + +macro "inrr" : tactic => + `(tactic | ( + apply Or.inr + apply Or.inr + )) + +macro "inrl" : tactic => + `(tactic | ( + apply Or.inr + apply Or.inl + )) + +macro "inl" : tactic => + `(tactic | ( + apply Or.inl + )) + +end Cedar.Thm diff --git a/cedar-lean/Cedar/Thm/Validation/Levels/WellFormed.lean b/cedar-lean/Cedar/Thm/Validation/Levels/WellFormed.lean new file mode 100644 index 000000000..ef1661e93 --- /dev/null +++ b/cedar-lean/Cedar/Thm/Validation/Levels/WellFormed.lean @@ -0,0 +1,1264 @@ + +import Cedar.Thm.Validation.Typechecker.Basic +import Cedar.Thm.Validation.Typechecker.And +import Cedar.Thm.Validation.Typechecker.GetAttr +import Cedar.Thm.Validation.Typechecker.HasAttr +import Cedar.Thm.Validation.Typechecker.Record +import Cedar.Thm.Validation.Typechecker.Set +import Cedar.Thm.Validation.Typechecker.Call +import Cedar.Thm.Data.Map +import Cedar.Thm.Validation.Typechecker +import Cedar.Thm.Validation.Levels.Util + +namespace Cedar.Thm + +open Cedar.Data +open Cedar.Spec +open Cedar.Validation + +-- This is the central property we will be proving here +-- If you start from a well formed entity store and evaluate an expression +-- Then you should end up with a value/type that is well formed +def EvaluatesToWellFormed (x : Expr) : Prop := + ∀ (v : Value) (ty : CedarType) (request : Request) + (env : Environment) + (entities : Entities) (c₁ c₂ : Capabilities) (l₁ : Level), + l₁ < .infinite → + RequestAndEntitiesMatchEnvironmentLeveled env request entities l₁ → + typeOf x c₁ env (l₁ == .infinite) = Except.ok (ty, c₂) → + evaluate x request entities = Except.ok v → + CapabilitiesInvariant c₁ request entities → + (entities ⊢ v : ty) + +-- A well formed value/type is still well formed at its lub +theorem levels_lub {entities : Entities} {v : Value} {ty₁ ty₂ ty : CedarType} + (h₁ : (ty₁ ⊔ ty₂) = some ty) + (h₂ : entities ⊢ v : ty₁) : + (entities ⊢ v : ty) := by + cases ty₁ <;> cases ty₂ <;> cases ty + <;> try simp [lub?, lubBool, Option.bind] at h₁ + <;> try simp [lub?, lubBool, Option.bind] at h₂ + case _ => + rename_i bty₁ bty₂ bty + cases bty₁ <;> cases bty₂ <;> cases bty + <;> try simp [lub?, lubBool] at h₁ <;> try assumption + case _ => + cases h₂ + rename_i b h₁ + cases b <;> simp [InstanceOfBoolType] at h₁ + apply WellFormed.bool + simp [InstanceOfBoolType] + case _ => + cases h₂ + rename_i b h₁ + cases b <;> simp [InstanceOfBoolType] at h₁ + apply WellFormed.bool + simp [InstanceOfBoolType] + case _ => + cases h₂ + rename_i b h₁ + cases b <;> simp [InstanceOfBoolType] at h₁ + apply WellFormed.bool + simp [InstanceOfBoolType] + case _ => + cases h₂ + rename_i b h₁ + cases b <;> simp [InstanceOfBoolType] at h₁ + apply WellFormed.bool + simp [InstanceOfBoolType] + case _ => + assumption + case _ => + assumption + case _ => + rename_i ty₁ l₁ ty₂ l₂ ty l + replace ⟨h₁, h₃, h₄⟩ := h₁ + subst h₁ + subst h₃ + subst h₄ + cases h₂ + case entity₀ euid h₁ => + cases l₂ <;> simp [min, Option.min, Level.finite, none] + <;> apply WellFormed.entity₀ + <;> assumption + case entityₙ euid attrs heq h₂ h₃ h₁ => + apply WellFormed.entityₙ + assumption + assumption + intros k v t' hin + apply h₂ + assumption + intros k v t' hin + have hstep : level t' ≥ l₁.sub1 := by apply h₁ ; repeat assumption + cases l₁ <;> cases l₂ + case _ => + simp [Level.sub1] at hstep + have hinf := le_infinite hstep + rw [hinf] + simp [LE.le, min, Option.min, Level.sub1, Level.infinite] + case _ => + rename_i n + simp [min, Option.min] + have hinf := le_infinite hstep + rw [hinf] + apply all_le_infinit + case _ => + rename_i n + simp [min, Option.min] + simp at hstep + assumption + case _ n₁ n₂ => + cases n₁ <;> cases n₂ <;> try (simp [min, Option.min, Level.sub1] ; apply zero_le_all) + case _ n₁ n₂ => + simp [Level.sub1] at hstep + simp + have h : (min (n₁ + 1) (n₂ + 1) = n₁ + 1) ∨ (min (n₁ + 1) (n₂ + 1) = n₂ + 1) := by omega + cases h + case _ h => + simp [Level.sub1, h] + assumption + case _ h => + simp [Level.sub1, h] + have h : some n₂ ≤ some n₁ := by apply le_lift ; omega + apply le_trans + repeat assumption + case _ => + rename_i ty₁ ty₂ ty + cases hlub : (ty₁ ⊔ ty₂) <;> simp [hlub] at h₁ + case _ => + rename_i ty₁ ty₂ + cases hlub : (ty₁ ⊔ ty₂) <;> simp [hlub] at h₁ + case _ => + rename_i ty₁ ty₂ + cases hlub : (ty₁ ⊔ ty₂) <;> simp [hlub] at h₁ + case _ => + rename_i ty₁ ty₂ ety₂ l₂ + cases hlub : (ty₁ ⊔ ty₂) <;> simp [hlub] at h₁ + case _ => + rename_i ty₁ ty₂ ty + cases hlub : (ty₁ ⊔ ty₂) <;> try simp [hlub] at h₁ + subst h₁ + cases h₂ + rename_i set h₂ + apply WellFormed.set + intros v hin + apply levels_lub + apply hlub + apply h₂ + apply hin + case _ => + rename_i ty₁ ty₂ rty + cases hlub : (ty₁ ⊔ ty₂) <;> simp [hlub] at h₁ + case _ => + rename_i ty₁ ty₂ rty + cases hlub : (ty₁ ⊔ ty₂) <;> simp [hlub] at h₁ + case _ => + not_record + case _ => + not_record + case _ => + not_record + case _ => + not_record + case _ => + not_record + case _ => + rename_i rty₁ rty₂ rty + + cases h₂ + rename_i value_record h₃ h₄ h₅ + cases rty₁ + rename_i rty₁ + cases rty₂ + rename_i rty₂ + cases rty + rename_i rty + have hlub : IsLubOfRecordTypes rty rty₁ rty₂ := by + apply lubRecordType_is_lub_of_record_types + simp [lub?, Option.bind] at h₁ + cases hlubRecord : lubRecordType rty₁ rty₂ <;> simp [hlubRecord] at h₁ + subst h₁ + rfl + apply WellFormed.record + case _ => + intros k h_contains + apply lubRecord_contains_left + apply hlub + apply h₃ + apply h_contains + case _ => + intros k v qty h_value_find h_ty_find + have h := lubRecord_find_implies_find hlub h_ty_find + have ⟨qty₁, qty₂, h_find₁, _, h_lub₁⟩ := h + clear h + have hwf : entities ⊢ v : (Qualified.getType qty₁) := by + apply h₅ + repeat assumption + apply levels_lub + apply qualitype_lub_lifts + apply h_lub₁ + apply hwf + case _ => + intros k qty h_find h_isreq + have h := lubRecord_find_implies_find_left hlub h_find + have ⟨qty₁, h_find₁, h_req₁⟩ := h + apply h₄ k qty₁ + apply h_find₁ + simp [h_req₁, h_isreq] + case _ => + not_record + case _ => + rename_i xty₁ xty₂ xty + have ⟨h₂, h₃⟩ := h₁ + subst h₂ + subst h₃ + assumption +termination_by (sizeOf v) +decreasing_by + all_goals simp_wf + all_goals (try omega) + case _ => + rename Set Value => members + rename _ = (Value.set members) => heq + rw [heq] + have step₁ : sizeOf v < sizeOf (members) := by + apply Set.in_set_means_smaller + assumption + have step₂ : sizeOf members < sizeOf (Value.set members) := by + simp + omega + omega + case _ => + rename _ = Value.record _ => heq + rename Map Attr Value => members + rw [heq] + have step₁ : sizeOf v < sizeOf members := by + apply Map.find_means_smaller + assumption + simp + omega + +theorem evaluates_to_value {v v' : Value} {e : Expr} {entities : Entities} {req : Request} + (h₁ : EvaluatesTo e req entities v) + (h₂ : evaluate e req entities = Except.ok v') : + v = v' := by + cases h₁ <;> rename_i h₁ <;> try simp [h₁] at h₂ + cases h₁ <;> rename_i h₁ <;> try simp [h₁] at h₂ + cases h₁ <;> rename_i h₁ <;> try simp [h₁] at h₂ + assumption + +theorem eq_is_bool {lhs rhs : Expr} {c₁ c₂ : Capabilities} {ty : CedarType} {env : Environment} {l₁ : Level} + (h₁ : typeOf (.binaryApp .eq lhs rhs) c₁ env (l₁ == .infinite) = Except.ok (ty, c₂)) : + ∃ bty, ty = .bool bty + := by + have hinv := type_of_eq_inversion h₁ + have ⟨hinv₁, hinv₂⟩ := hinv + clear hinv + subst hinv₁ + split at hinv₂ + case _ => + split at hinv₂ + <;> subst hinv₂ + <;> try (solve | exists .tt) + exists .ff + case _ => + replace ⟨ty₁, _, ty₂, _, _, _, hinv₄ ⟩ := hinv₂ + split at hinv₄ + case _ => + exists BoolType.anyBool + case _ => + exists BoolType.ff + simp [hinv₄] + +theorem int_cmp_is_bool {o : BinaryOp} {lhs rhs : Expr} {c₁ c₂ : Capabilities} {ty : CedarType} {env : Environment} {l₁ : Level} + (h₁ : o = .less ∨ o = .lessEq) + (h₂ : typeOf (.binaryApp o rhs lhs) c₁ env (l₁ == .infinite) = Except.ok (ty, c₂)) : + ∃ bty, ty = .bool bty := by + have hinv := type_of_int_cmp_inversion h₁ h₂ + replace ⟨_, hinv, _⟩ := hinv + subst hinv + exists .anyBool + +theorem int_arith_is_int {o : BinaryOp} {lhs rhs : Expr} {c₁ c₂ : Capabilities} {ty : CedarType} {env : Environment} {l₁ : Level} + (h₁ : o = .add ∨ o = .sub ∨ o = .mul) + (h₂ : typeOf (.binaryApp o rhs lhs) c₁ env (l₁ == .infinite) = Except.ok (ty, c₂)) : + ty = .int := by + have hinv := type_of_int_arith_inversion h₁ h₂ + simp [hinv] + +theorem evaluates_to_well_formed_binary {o : BinaryOp} {lhs rhs : Expr} {v : Value} {ty: CedarType} {request : Request} {entities: Entities} {env : Environment} {c₁ c₂ : Capabilities} {l₁ : Level} + (h₂ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l₁) + (h₃ : typeOf (.binaryApp o lhs rhs) c₁ env (l₁ == .infinite) = Except.ok (ty, c₂)) + (h₅ : evaluate (.binaryApp o lhs rhs) request entities = Except.ok v) + (h_caps : CapabilitiesInvariant c₁ request entities) + : + entities ⊢ v : ty := by + have hsound := (@type_of_is_sound (.binaryApp o lhs rhs)) + (by assumption) + (by + apply request_and_entity_match_level_implies_regular + assumption + ) + (by assumption) + have ⟨hsound₁, v', hsound₂, hsound₃⟩ := hsound + have heq : v' = v := by + apply evaluates_to_value + repeat assumption + subst heq + cases o + case eq => + have heq := eq_is_bool h₃ + replace ⟨bty, heq⟩ := heq + subst heq + cases hsound₃ + apply WellFormed.bool + assumption + case mem => + have hinv := type_of_mem_inversion h₃ + have ⟨_, ety₂, l', l'', hinv₁, _, c₂', hinv₃⟩ := hinv + cases hinv₃ + <;> rename_i hinv₃ + <;> have ⟨_, hinv₄⟩ := hinv₃ + <;> simp [typeOfInₑ, typeOfInₛ] at hinv₄ + <;> split at hinv₄ + <;> simp [ok, err] at hinv₄ + <;> subst hinv₄ + <;> cases hsound₃ + <;> apply WellFormed.bool + <;> assumption + case less => + have hbty := int_cmp_is_bool (by simp) h₃ + replace ⟨bty, hbty⟩ := hbty + subst hbty + cases hsound₃ + apply WellFormed.bool + assumption + case lessEq => + have hbty := int_cmp_is_bool (by simp) h₃ + replace ⟨bty, hbty⟩ := hbty + subst hbty + cases hsound₃ + apply WellFormed.bool + assumption + case add => + have hint := int_arith_is_int (by simp) h₃ + subst hint + cases hsound₃ + apply WellFormed.int + case sub => + have hint := int_arith_is_int (by simp) h₃ + subst hint + cases hsound₃ + apply WellFormed.int + case mul => + have hint := int_arith_is_int (by simp) h₃ + subst hint + cases hsound₃ + apply WellFormed.int + case contains => + have hinv := type_of_contains_inversion h₃ + have heq : ty = CedarType.bool BoolType.anyBool := by simp [hinv] + subst heq + cases hsound₃ + apply WellFormed.bool + assumption + case containsAll => + have hinv := type_of_containsA_inversion (by simp) h₃ + have heq : ty = CedarType.bool BoolType.anyBool := by simp [hinv] + subst heq + cases hsound₃ + apply WellFormed.bool + assumption + case containsAny => + have hinv := type_of_containsA_inversion (by simp) h₃ + have heq : ty = CedarType.bool BoolType.anyBool := by simp [hinv] + subst heq + cases hsound₃ + apply WellFormed.bool + assumption + +theorem evaluates_to_well_formed_unary {o : UnaryOp} {x : Expr} {v : Value} {ty: CedarType} {request : Request} {entities: Entities} {env : Environment} {c₁ c₂ : Capabilities} {l₁ : Level} + (h₂ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l₁) + (h₃ : typeOf (.unaryApp o x) c₁ env (l₁ == .infinite) = Except.ok (ty, c₂)) + (h₅ : evaluate (.unaryApp o x) request entities = Except.ok v) + (h_caps : CapabilitiesInvariant c₁ request entities) + : + entities ⊢ v : ty := by + have hreq : RequestAndEntitiesMatchEnvironment env request entities := by + apply request_and_entity_match_level_implies_regular + assumption + have hsound := (@type_of_is_sound (.unaryApp o x)) (by assumption) (by assumption) (by assumption) + have ⟨hsound₁, v', hsound₂, hsound₃⟩ := hsound + have heq : v' = v := by + apply evaluates_to_value + repeat assumption + subst heq + clear hsound + cases o + case not => + have hinv := type_of_not_inversion h₃ + have ⟨_, bty, _, hinv₂, _⟩ := hinv + clear hinv + subst hinv₂ + cases hsound₃ + apply WellFormed.bool + assumption + case neg => + have hinv := type_of_neg_inversion h₃ + have ⟨_, hinv₂, _, _⟩ := hinv + clear hinv + subst hinv₂ + cases hsound₃ + apply WellFormed.int + case like => + have hinv := type_of_like_inversion h₃ + have ⟨_, hinv₂, _⟩ := hinv + subst hinv₂ + cases hsound₃ + apply WellFormed.bool + assumption + case is => + have hinv := type_of_is_inversion h₃ + have ⟨_, ety', _, _, hinv₂, _⟩ := hinv + clear hinv + rename_i _ + split at hinv₂ <;> ( + subst hinv₂ + cases hsound₃ + apply WellFormed.bool + assumption + ) + +theorem evaluates_to_well_formed_var {v : Var} {val : Value} {request : Request} {entities: Entities} {env : Environment} {c₁ c₂ : Capabilities} {l₁ : Level} + (h₂ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l₁) + (h₃ : typeOf (.var v) c₁ env (l₁ == .infinite) = Except.ok (ty, c₂)) + (h₅ : evaluate (.var v) request entities = Except.ok val) + : + entities ⊢ val : ty + := by + unfold RequestAndEntitiesMatchEnvironmentLeveled at h₂ + have ⟨_, _, _, _, _, _, _⟩ := h₂ + clear h₂ + cases v + <;> simp [evaluate] at h₅ + <;> subst h₅ + <;> simp [typeOf, typeOfVar, ok] at h₃ + <;> replace ⟨h₃, _⟩ := h₃ + <;> subst h₃ + <;> assumption + +theorem evaluates_to_well_formed_lit {l : Prim} {ty: CedarType} {request : Request} {entities: Entities} {env : Environment} {c₁ c₂ : Capabilities} {l₁ : Level} + (h₁ : l₁ < .infinite) + (h₂ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l₁) + (h₃ : typeOf (.lit l) c₁ env (l₁ == .infinite) = Except.ok (ty, c₂)) + (h₅ : evaluate (.lit l) request entities = Except.ok v) + : + entities ⊢ v : ty := by + cases l <;> try ( + rename_i value + cases value <;> ( + simp [evaluate] at h₅ + subst h₅ + simp [typeOf, typeOfLit, ok] at h₃ + replace ⟨h₃, _⟩ := h₃ + subst h₃ + ) + ) + case int => + apply WellFormed.int + case bool.true => + apply WellFormed.bool + simp [InstanceOfBoolType] + case bool.false => + apply WellFormed.bool + simp [InstanceOfBoolType] + case string => + apply WellFormed.string + case entityUID => + rename_i euid + cases heq : decide (l₁ == .infinite) + case _ => + simp at heq + simp [typeOf, typeOfLit] at h₃ + split at h₃ + case _ => + simp [ok] at h₃ + replace ⟨h₃, _⟩ := h₃ + subst h₃ + simp [evaluate] at h₅ + subst h₅ + apply WellFormed.entity₀ + simp [InstanceOfEntityType] + case _ => + exfalso + simp [err] at h₃ + case _ => + simp at heq + subst heq + cases h₁ + +theorem set_eval_step {members : List Expr} {values : List Value} {v : Value} {request : Request} {entities : Entities} + (h₁ : (members.mapM (λ x => evaluate x request entities)) = .ok values) + (h₂ : v ∈ values) : + ∃ e, e ∈ members ∧ evaluate e request entities = .ok v + := by + cases values + case nil => + cases h₂ + case cons head tail => + cases members + case nil => + simp [pure, Except.pure] at h₁ + case cons members_head members_tail => + cases h₂ -- Case analysis on `v ∈ (head :: tail)` + case _ => -- Case 1: `v = head` + rw [List.mapM_cons] at h₁ + cases h_eval_head : evaluate members_head request entities <;> simp [h_eval_head] at h₁ + rename_i head_value + cases h_eval_tail : List.mapM (λ x => evaluate x request entities) members_tail <;> simp [h_eval_tail] at h₁ + rename_i values_tail + simp [pure, Except.pure] at h₁ + exists members_head + simp [h₁, h_eval_head] + case _ => -- Case 2: `v ∈ tail ∧ v ≠ head`. This is the inductive case + rename_i hin + rw [List.mapM_cons] at h₁ + cases h_eval_head : evaluate members_head request entities <;> simp [h_eval_head] at h₁ + rename_i head_value + cases h_eval_tail : List.mapM (λ x => evaluate x request entities) members_tail <;> simp [h_eval_tail] at h₁ + rename_i values_tail + simp [pure, Except.pure] at h₁ + have ⟨_, h₄⟩ := h₁ + subst h₄ + have ih := @set_eval_step members_tail values_tail v request entities h_eval_tail hin + replace⟨e, ih⟩ := ih + exists e + constructor <;> simp [ih] + +macro "extn_comparator_correct" f:(ident) inv:(ident) hsound₃:(ident) hsound₂:(ident) h₃:(ident) h₅:(ident) : tactic => + `(tactic | ( + have h₃ := $h₃ + have hsound₃ := $hsound₃ + have hsound₂ := $hsound₂ + have h₅ := $h₅ + have hinv := $inv (by simp [IsDecimalComparator, IsIpAddrRecognizer]) h₃ + replace ⟨hinv, _⟩ := hinv + subst hinv + cases hsound₃ + cases hsound₂ <;> rename_i hsound₂ <;> simp [h₅] at hsound₂ + subst hsound₂ + apply $f + assumption + )) + +macro "decimal_comparator_correct" hsound₃:(ident) hsound₂:(ident) h₃:(ident) h₅:(ident) : tactic => + `(tactic | ( + extn_comparator_correct WellFormed.bool type_of_call_decimal_comparator_inversion $hsound₃ $hsound₂ $h₃ $h₅ + )) + +macro "ip_recognizer_correct" hsound₃:(ident) hsound₂:(ident) h₃:(ident) h₅:(ident) : tactic => + `(tactic | ( + extn_comparator_correct WellFormed.bool type_of_call_ipAddr_recognizer_inversion $hsound₃ $hsound₂ $h₃ $h₅ + )) + +theorem evaluates_to_well_formed_call {xfn : ExtFun} {xs : List Expr} {v : Value} {ty : CedarType} {request : Request} {entities : Entities} {env : Environment} {c₁ c₂ : Capabilities} {l : Level} + (h₂ : RequestAndEntitiesMatchEnvironment env request entities) + (h₃ : typeOf (.call xfn xs) c₁ env (l == Level.infinite) = Except.ok (ty, c₂)) + (h₅ : evaluate (.call xfn xs) request entities = Except.ok v) + (h_caps : CapabilitiesInvariant c₁ request entities) : + entities ⊢ v : ty + := by + have hsound := (@type_of_is_sound (.call xfn xs)) + (by assumption) + (by assumption) + (by assumption) + have ⟨hsound₁, v, hsound₂, hsound₃⟩ := hsound + clear hsound + cases xfn + case decimal => + have hinv := type_of_call_decimal_inversion h₃ + replace ⟨hinv, _⟩ := hinv + subst hinv + cases hsound₃ + cases hsound₂ <;> rename_i hsound₂ <;> simp [h₅] at hsound₂ + subst hsound₂ + apply WellFormed.ext + assumption + case lessThan => + decimal_comparator_correct hsound₃ hsound₂ h₃ h₅ + case lessThanOrEqual => + decimal_comparator_correct hsound₃ hsound₂ h₃ h₅ + case greaterThan => + decimal_comparator_correct hsound₃ hsound₂ h₃ h₅ + case greaterThanOrEqual => + decimal_comparator_correct hsound₃ hsound₂ h₃ h₅ + case ip => + have hinv := type_of_call_ip_inversion h₃ + replace ⟨hinv, _⟩ := hinv + subst hinv + cases hsound₃ + cases hsound₂ <;> rename_i hsound₂ <;> simp [h₅] at hsound₂ + subst hsound₂ + apply WellFormed.ext + assumption + case isIpv4 => + ip_recognizer_correct hsound₃ hsound₂ h₃ h₅ + case isIpv6 => + ip_recognizer_correct hsound₃ hsound₂ h₃ h₅ + case isLoopback => + ip_recognizer_correct hsound₃ hsound₂ h₃ h₅ + case isMulticast => + ip_recognizer_correct hsound₃ hsound₂ h₃ h₅ + case isInRange => + have hinv := type_of_call_isInRange_inversion h₃ + replace ⟨hinv, _⟩ := hinv + subst hinv + cases hsound₃ + cases hsound₂ <;> rename_i hsound₂ <;> simp [h₅] at hsound₂ + subst hsound₂ + apply WellFormed.bool + assumption + +theorem evaluates_to_well_formed_ite {cond cons alt : Expr} {v : Value} {request : Request} {entities : Entities} {c₁ c₂ : Capabilities} {l₁ : Level} + (h₁ : l₁ < .infinite) + (h₂ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l₁) + (h₃ : typeOf (.ite cond cons alt) c₁ env (l₁ == .infinite) = Except.ok (ty, c₂)) + (h₅ : evaluate (.ite cond cons alt) request entities = Except.ok v) + (h_caps : CapabilitiesInvariant c₁ request entities) + (ih₁ : EvaluatesToWellFormed cons) + (ih₂ : EvaluatesToWellFormed alt) : + entities ⊢ v : ty := by + have hsound : GuardedCapabilitiesInvariant (.ite cond cons alt) c₂ request entities ∧ ∃ v, EvaluatesTo (.ite cond cons alt) request entities v ∧ InstanceOfType v ty := by + type_soundness + have ⟨hsound₁, v', hsound₂, hsound₃⟩ := hsound + clear hsound + dril_to_value hsound₂ h₅ + have hinv := type_of_ite_inversion h₃ + replace ⟨bty, c₁', ty₂, c₂', ty₃, c₃, hinv₁, hinv⟩ := hinv + have hsound_cond : GuardedCapabilitiesInvariant cond c₁' request entities ∧ ∃ v, EvaluatesTo cond request entities v ∧ InstanceOfType v (.bool bty) := by + type_soundness + have ⟨hsound_cond₁, v_cond, hsound_cond₂, hsound_cond₃⟩ := hsound_cond + clear hsound_cond + have heval_cond : ∃ v, evaluate cond request entities = Except.ok v := by + cases hcond : evaluate cond request entities + case error => + simp [evaluate, hcond, Result.as] at h₅ + case ok v => + exists v + replace ⟨v_cond, heval_cond⟩ := heval_cond + dril_to_value hsound_cond₂ heval_cond + cases hsound_cond₃ + rename_i bool_val hsound_cond₃ + cases bty + case tt => + simp at hinv + have ⟨hinv₂, hinv₃, _⟩ := hinv + clear hinv + subst hinv₃ + cases bool_val <;> simp [InstanceOfBoolType] at hsound_cond₃ + apply ih₁ + repeat assumption + simp [evaluate, heval_cond, Result.as, Coe.coe, Value.asBool] at h₅ + assumption + apply capability_union_invariant + assumption + apply hsound_cond₁ + assumption + case ff => + simp at hinv + have ⟨hinv₂, hinv₃, hinv₄⟩ := hinv + clear hinv + subst hinv₃ + subst hinv₄ + cases bool_val <;> simp [InstanceOfBoolType] at hsound_cond₃ + apply ih₂ + repeat assumption + simp [evaluate, heval_cond, Result.as, Coe.coe, Value.asBool] at h₅ + repeat assumption + case anyBool => + simp at hinv + have ⟨hinv₁, hinv₂, hinv₃, hinv₄⟩ := hinv + clear hinv + subst hinv₄ + cases bool_val + case true => + have hsound_cons : GuardedCapabilitiesInvariant cons c₂' request entities ∧ ∃ v, EvaluatesTo cons request entities v ∧ InstanceOfType v ty₂ := by + apply type_of_is_sound + apply capability_union_invariant + assumption + apply hsound_cond₁ + assumption + apply request_and_entity_match_level_implies_regular + assumption + assumption + have ⟨_, v_cons, hsound_cons₂, _⟩ := hsound_cons + clear hsound_cons + have heval_cons : ∃ v, evaluate cons request entities = .ok v := by + cases h : evaluate cons request entities + case ok v => + exists v + case error => + simp [evaluate, heval_cond, h, Result.as, Coe.coe, Value.asBool] at h₅ + replace ⟨v_cons', heval_cons⟩ := heval_cons + dril_to_value hsound_cons₂ heval_cons + simp [heval_cons, heval_cond, evaluate, Result.as, Coe.coe, Value.asBool] at h₅ + subst h₅ + have step : (entities ⊢ v_cons : ty₂) := by + apply ih₁ + assumption + assumption + assumption + assumption + apply capability_union_invariant + assumption + apply hsound_cond₁ + assumption + apply levels_lub + repeat assumption + case false => + + have hsound_alt : GuardedCapabilitiesInvariant alt c₃ request entities ∧ ∃ v, EvaluatesTo alt request entities v ∧ InstanceOfType v ty₃ := by + type_soundness + have ⟨hsound_alt₁, v_alt, hsound_alt₂, hsound_alt₃⟩ := hsound_alt + clear hsound_alt + have heval_alt : ∃ v, evaluate alt request entities = .ok v := by + cases h : evaluate alt request entities + case ok v => + exists v + case error => + simp [evaluate, h, heval_cond, Result.as, Coe.coe, Value.asBool] at h₅ + replace ⟨v_alt', heval_alt⟩ := heval_alt + dril_to_value hsound_alt₂ heval_alt + simp [evaluate, heval_alt, heval_cond, Result.as, Coe.coe, Value.asBool] at h₅ + subst h₅ + have step : entities ⊢ v_alt : ty₃ := by + apply ih₂ + repeat assumption + apply levels_lub + rw [lub_comm] at hinv₃ + repeat assumption + +macro "bool_constant" hinv:(ident) hsound₃:(ident) : tactic => + `(tactic | ( + solve | ( + have hinv := $hinv + have hsound₃ := $hsound₃ + simp [lubBool] at hinv + have ⟨hinv₁, hinv₂, hinv₃⟩ := hinv + subst hinv₂ + cases hsound₃ + rename_i bool_val hsound₃ + cases bool_val <;> simp [InstanceOfBoolType] at hsound₃ + apply WellFormed.bool + assumption + ) + )) + +theorem evaluates_to_well_formed_and {lhs rhs : Expr} {v : Value} {request : Request} {entities : Entities} {c₁ c₂ : Capabilities} {l₁ : Level} + (h₂ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l₁) + (h₃ : typeOf (.and lhs rhs) c₁ env (l₁ == .infinite) = Except.ok (ty, c₂)) + (h₅ : evaluate (.and lhs rhs) request entities = Except.ok v) + (h_caps : CapabilitiesInvariant c₁ request entities) : + entities ⊢ v : ty := by + have hsound : GuardedCapabilitiesInvariant (.and lhs rhs) c₂ request entities ∧ ∃ v, EvaluatesTo (.and lhs rhs) request entities v ∧ InstanceOfType v ty + := by type_soundness + have ⟨_, v', hsound₂, hsound₃⟩ := hsound + clear hsound + dril_to_value hsound₂ h₅ + have hinv := type_of_and_inversion h₃ + replace ⟨bty, c₁', hinv⟩ := hinv + cases bty + case tt => + simp at hinv + replace ⟨_, bty₂, c₂', hinv⟩ := hinv + cases bty₂ <;> bool_constant hinv hsound₃ + case ff => + simp at hinv + bool_constant hinv hsound₃ + case anyBool => + simp at hinv + replace ⟨_, bty₂, c₂', hinv⟩ := hinv + cases bty₂ <;> simp at hinv <;> bool_constant hinv hsound₃ + +theorem evaluates_to_well_formed_or {lhs rhs : Expr} {v : Value} {request : Request} {entities : Entities} {c₁ c₂ : Capabilities} {l₁ : Level} + (h₂ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l₁) + (h₃ : typeOf (.or lhs rhs) c₁ env (l₁ == .infinite) = Except.ok (ty, c₂)) + (h₅ : evaluate (.or lhs rhs) request entities = Except.ok v) + (h_caps : CapabilitiesInvariant c₁ request entities) : + entities ⊢ v : ty := by + have hsound : GuardedCapabilitiesInvariant (.or lhs rhs) c₂ request entities ∧ ∃ v, EvaluatesTo (.or lhs rhs) request entities v ∧ InstanceOfType v ty := by + type_soundness + have ⟨_, v', hsound₂, hsound₃⟩ := hsound + clear hsound + dril_to_value hsound₂ h₅ + have hinv := type_of_or_inversion h₃ + replace ⟨bty₁, c₁', hinv⟩ := hinv + cases bty₁ + case tt => + bool_constant hinv hsound₃ + case ff => + simp at hinv + replace ⟨_, bty₂, hinv⟩ := hinv + cases bty₂ + case ff => + have ⟨_, hinv₂⟩ := hinv + clear hinv + subst hinv₂ + cases hsound₃ + rename_i bool hsound₃ + cases bool <;> simp [InstanceOfBoolType] at hsound₃ + apply WellFormed.bool + assumption + case tt => + have ⟨_, hinv₂⟩ := hinv + clear hinv + subst hinv₂ + cases hsound₃ + rename_i bool hsound₃ + cases bool <;> simp [InstanceOfBoolType] at hsound₃ + apply WellFormed.bool + assumption + case anyBool => + have ⟨_, hinv₂⟩ := hinv + clear hinv + subst hinv₂ + cases hsound₃ + rename_i bool hsound₃ + cases bool <;> simp [InstanceOfBoolType] at hsound₃ + <;> apply WellFormed.bool + <;> assumption + case anyBool => + simp at hinv + replace ⟨_, bty₂, c₂', hinv⟩ := hinv + cases bty₂ <;> bool_constant hinv hsound₃ + +theorem evaluates_to_well_formed_getAttr_entity {e : Expr} {attr : Attr} {v : Value} {request : Request} {entities : Entities} {env : Environment} {c₁ c₁' : Capabilities} {l₁ l₂ : Level} {ety : EntityType} + (h₁ : l₁ < .infinite) + (h₂ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l₁) + (h₄ : evaluate (.getAttr e attr) request entities = .ok v) + (h₅ : CapabilitiesInvariant c₁ request entities) + (h₆ : typeOf e c₁ env (l₁ == Level.infinite) = Except.ok (CedarType.entity ety l₂, c₁')) + (h₇ : l₂ > Level.finite 0) + (ih : EvaluatesToWellFormed e) : + (entities ⊢ v : ty) := by + have hsound : GuardedCapabilitiesInvariant e c₁' request entities ∧ ∃ v, EvaluatesTo e request entities v ∧ InstanceOfType v (.entity ety l₂) := by + type_soundness + have ⟨_, v_entity, hsound₂, hsound₃⟩ := hsound + have heval : ∃ v_entity', evaluate e request entities = .ok v_entity' := by + cases h : evaluate e request entities + case ok v => + exists v + case error => + simp [evaluate, Result.as, h] at h₄ + replace ⟨v_entity', heval⟩ := heval + dril_to_value hsound₂ heval + cases hsound₃ + rename_i euid hsound₃ + have step : entities ⊢ (.prim (.entityUID euid)) : .entity ety l₂ := by + apply ih + repeat assumption + cases step -- Two ways to have a well formed entity value + case _ => -- The level can be zero, which is impossible as we're derefing ijt + cases h₇ + omega + case _ => -- The level can be nonzero + rename_i attrs ih₁ ih₂ ih₃ ih₄ + apply ih₁ + simp [getAttr, attrsOf, ih₁, evaluate, heval, ih₂] at h₄ + apply Map.findOrErr_ok_implies_in_kvs + apply h₄ + +theorem evaluates_to_well_formed_getAttr_record {e : Expr} {attr : Attr} {v : Value} {request : Request} {entities : Entities} {env : Environment} {c₁ c₂ c₁' : Capabilities} {l₁ : Level} {rty : RecordType} + (h₁ : l₁ < .infinite) + (h₂ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l₁) + (h₃ : typeOf (.getAttr e attr) c₁ env (l₁ == .infinite) = .ok (ty, c₂)) + (h₄ : evaluate (.getAttr e attr) request entities = .ok v) + (h₅ : CapabilitiesInvariant c₁ request entities) + (h₆ : typeOf e c₁ env (l₁ == Level.infinite) = Except.ok (CedarType.record rty, c₁')) + (ih : EvaluatesToWellFormed e) : + (entities ⊢ v : ty) := by + have hsound : GuardedCapabilitiesInvariant e c₁' request entities ∧ ∃ v, EvaluatesTo e request entities v ∧ InstanceOfType v (.record rty) := by + type_soundness + have ⟨_, v_record, hsound₂, hsound₃⟩ := hsound + clear hsound + have heval : ∃ v_record', evaluate e request entities = .ok v_record' := by + cases h: evaluate e request entities + case ok v => + exists v + case error _ => + simp [h, evaluate, Result.as] at h₄ + replace ⟨v_record', heval⟩ := heval + dril_to_value hsound₂ heval + cases hsound₃ + rename_i record hsound₃₁ hsound₃₂ hsound₃₃ + have step : entities ⊢ (Value.record record) : .record rty := by + apply ih + repeat assumption + cases step + rename_i ih₁ ih₂ ih₃ + simp [evaluate, getAttr, attrsOf, heval] at h₄ + have hcontains : rty.contains attr = true := by + apply hsound₃₁ + simp [Map.contains] + cases h : record.find? attr + case a.some => + simp + case a.none => + simp [Map.findOrErr, h] at h₄ + have hqty : ∃ (qty : QualifiedType), rty.find? attr = .some qty := by + simp [Map.contains] at hcontains + cases h : Map.find? rty attr + case none => + simp [h] at hcontains + case some qty => + exists qty + replace ⟨qty, hqty⟩ := hqty + have heq : qty.getType = ty := by + simp [typeOf, h₆, typeOfGetAttr, getAttrInRecord, hqty] at h₃ + cases qty + case _ => + simp at h₃ + cases hin : decide ((e,attr) ∈ c₁) <;> simp at hin + case _ => + rw [if_neg] at h₃ + simp [err] at h₃ + assumption + case _ => + rw [if_pos] at h₃ + rename_i ty' + simp [ok] at h₃ + replace ⟨h₃, _⟩ := h₃ + subst h₃ + simp [Qualified.getType] + assumption + case _ => + rename_i ty' + simp [ok] at h₃ + replace ⟨h₃, _⟩ := h₃ + subst h₃ + simp [Qualified.getType] + have step := ih₃ attr v qty + (by + cases h : record.find? attr <;> simp [h, Map.findOrErr] at h₄ + subst h₄ + rfl + ) + hqty + rw [heq] at step + assumption + +theorem evaluates_to_well_formed_getAttr {e : Expr} {attr : Attr} {v : Value} {request : Request} {entities : Entities} {env : Environment} {c₁ c₂ : Capabilities} {l₁ : Level} + (h₁ : l₁ < .infinite) + (h₂ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l₁) + (h₃ : typeOf (.getAttr e attr) c₁ env (l₁ == .infinite) = .ok (ty, c₂)) + (h₄ : evaluate (.getAttr e attr) request entities = .ok v) + (h₅ : CapabilitiesInvariant c₁ request entities) + (ih : EvaluatesToWellFormed e) : + entities ⊢ v : ty := by + have hsound : GuardedCapabilitiesInvariant (.getAttr e attr) c₂ request entities ∧ ∃ v, EvaluatesTo (.getAttr e attr) request entities v ∧ InstanceOfType v ty := by + type_soundness + have ⟨hsound₁, v', hsound₂, _⟩ := hsound + clear hsound + dril_to_value hsound₂ h₄ + have hinv := type_of_getAttr_inversion_levels h₃ + replace ⟨hinv₁, c₁', hinv⟩ := hinv + subst hinv₁ + cases hinv <;> rename_i hinv + case _ => + replace ⟨ety, l₂, hinv₁, hinv₂⟩ := hinv + clear hinv + apply evaluates_to_well_formed_getAttr_entity + repeat assumption + case _ => + replace ⟨rty, hinv⟩ := hinv + apply evaluates_to_well_formed_getAttr_record + repeat assumption + +theorem levels_correct_hasAttr {e : Expr} {attr : Attr} {v : Value} {request : Request} {entities : Entities} {env : Environment} {c₁ c₂ : Capabilities} {l₁ : Level} + (h₂ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l₁) + (h₃ : typeOf (.hasAttr e attr) c₁ env (l₁ == .infinite) = .ok (ty, c₂)) + (h₄ : evaluate (.hasAttr e attr) request entities = .ok v) + (h₅ : CapabilitiesInvariant c₁ request entities) : + entities ⊢ v : ty := by + have hsound : GuardedCapabilitiesInvariant (.hasAttr e attr) c₂ request entities ∧ ∃ v, EvaluatesTo (.hasAttr e attr) request entities v ∧ InstanceOfType v ty := by + type_soundness + have ⟨_, v', hsound₂, hsound₃⟩ := hsound + dril_to_value hsound₂ h₄ + have hinv := type_of_hasAttr_inversion' h₃ + replace ⟨bty, hinv⟩ := hinv + subst hinv + cases bty <;> try ( + solve | ( + cases hsound₃ + rename_i bool hsound₃ + cases bool <;> simp [InstanceOfBoolType] at hsound₃ + apply WellFormed.bool + assumption + ) + ) + case anyBool => + cases hsound₃ + apply WellFormed.bool + assumption + +theorem levels_correct_set {members : List Expr} {v : Value} {request : Request} {entities : Entities} {env : Environment} {c₁ c₂ : Capabilities} {l₁ : Level} + (h₁ : l₁ < .infinite) + (h₂ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l₁) + (h₃ : typeOf (.set members) c₁ env (l₁ == .infinite) = Except.ok (ty, c₂)) + (h₄ : evaluate (.set members) request entities = .ok v) + (h₅ : CapabilitiesInvariant c₁ request entities) + (ih : ∀ e, e ∈ members → EvaluatesToWellFormed e) : + entities ⊢ v : ty := by + have hsound : GuardedCapabilitiesInvariant (.set members) c₂ request entities ∧ ∃ v, EvaluatesTo (.set members) request entities v ∧ InstanceOfType v ty := by + type_soundness + have ⟨hsound₁, v', hsound₂, hsound₃⟩ := hsound + clear hsound + dril_to_value hsound₂ h₄ + have hinv := type_of_set_inversion h₃ + have ⟨hinv₁, ty', hinv₂, hinv₃⟩ := hinv + clear hinv + subst hinv₂ + subst hinv₁ + cases hsound₃ + rename_i set_value hinst + apply WellFormed.set + intros v hin + simp [evaluate, List.mapM₁, List.attach, List.attachWith] at h₄ + rw [List.mapM_pmap_subtype (λ (e : Expr) => evaluate e request entities) members] at h₄ + cases hmembers : List.mapM (λ e => evaluate e request entities) members <;> simp [hmembers] at h₄ + rename_i member_list + have in_list : v ∈ member_list := by + apply Set.in_constructor_in_set + rw [h₄] + assumption + have hexpr := (@set_eval_step members member_list) hmembers in_list + have ⟨e, hexpr₁, hexpr₂⟩ := hexpr + clear hexpr + have hinv₄ := hinv₃ e hexpr₁ + replace ⟨tyᵢ, cᵢ, hinv₄, hinv₅⟩ := hinv₄ + have step : entities ⊢ v : tyᵢ := by + apply ih + repeat assumption + apply levels_lub + repeat assumption + +theorem record_eval_step (attrs : List (Attr × Expr)) (record : List (Attr × Value)) (request : Request) (entities : Entities) (attr : Attr) (v : Value) + (h₁ : List.mapM (λ (pair : (Attr × Expr)) => bindAttr pair.fst (evaluate pair.snd request entities)) attrs = .ok record) + (h₂ : (attr, v) ∈ record) : + ∃ e, (attr, e) ∈ attrs ∧ evaluate e request entities = .ok v := by + cases attrs <;> cases record + case _ => + cases h₂ + case _ => + simp [List.mapM, List.mapM.loop, pure, Except.pure] at h₁ + case _ => + cases h₂ + case _ attr_head attr_tail record_head record_tail => + rw [List.mapM_cons] at h₁ + cases hhead : bindAttr attr_head.fst (evaluate attr_head.snd request entities) <;> simp [hhead] at h₁ + cases htail : List.mapM (λ pair => bindAttr pair.fst (evaluate pair.snd request entities)) attr_tail <;> simp [htail] at h₁ + rename_i head tail + cases h₂ + case _ => + cases head + simp [pure, Except.pure] at h₁ + have ⟨⟨heq₁, heq₂⟩, heq₃⟩ := h₁ + subst heq₁ + subst heq₂ + subst heq₃ + rename_i attr v + clear h₁ + have ⟨attr', e⟩ := attr_head + exists e + simp [bindAttr] at hhead + cases heval : evaluate e request entities <;> simp [heval] at hhead + simp [hhead] + case _ hin => + simp [pure, Except.pure] at h₁ + simp [h₁] at htail + have ih := record_eval_step attr_tail record_tail request entities attr v htail hin + replace ⟨e, ih⟩ := ih + exists e + simp [ih] + +theorem forall₂_find_match {α β : Type} (l₁ : List α) (l₂ : List β) (p : α → β → Prop) (a : α) + (h₁ : List.Forall₂ p l₁ l₂) + (h₂ : a ∈ l₁) : + ∃ b, b ∈ l₂ ∧ p a b + := by + cases h₁ + case nil => + cases h₂ + case cons => + rename_i a_head b_head a_tail b_tail holds_on_head holds_on_tail + cases h₂ + case _ => + exists b_head + simp [holds_on_head] + case _ hin => + have ih := forall₂_find_match a_tail b_tail p a holds_on_tail hin + replace ⟨b, ih₁, ih₂⟩ := ih + exists b + simp [ih₁, ih₂] + +theorem AttrExprHasAttrType_constraints_keys (k₁ k₂ : Attr) (e : Expr) (qty : QualifiedType) (c : Capabilities) (env : Environment) (l : Level) + (h₁ : AttrExprHasAttrType c env l (k₁, e) (k₂, qty)) : + k₁ = k₂ := by + simp [AttrExprHasAttrType] at h₁ + have ⟨_, h₁⟩ := h₁ + simp [h₁] + +theorem evaluate_empty_record (request : Request) (entities : Entities) : + evaluate (.record []) request entities = .ok (Value.record (Map.make [])) := by + simp [evaluate, List.mapM₂, List.attach₂] + +theorem evaluates_to_well_formed_record {attrs : List (Attr × Expr)} {v : Value} {request : Request} {entities : Entities} {env : Environment} {c₁ c₂ : Capabilities} {l₁ : Level} + (h₁ : l₁ < .infinite) + (h₂ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l₁) + (h₃ : typeOf (.record attrs) c₁ env (l₁ == .infinite) = Except.ok (ty, c₂)) + (h₅ : evaluate (.record attrs) request entities = Except.ok v) + (h_caps : CapabilitiesInvariant c₁ request entities) + (ih : ∀ (a : Attr) (e : Expr), (a,e) ∈ attrs → EvaluatesToWellFormed e ) : + entities ⊢ v : ty := by + have hsound : GuardedCapabilitiesInvariant (.record attrs) c₂ request entities ∧ ∃ v, EvaluatesTo (.record attrs) request entities v ∧ InstanceOfType v ty := by + type_soundness + have ⟨hsound₁, v', hsound₂, hsound₃⟩ := hsound + clear hsound + dril_to_value hsound₂ h₅ + clear hsound₂ + have hinv := type_of_record_inversion h₃ + replace ⟨_, rty, hinv₁, hinv⟩ := hinv + subst hinv₁ + cases hsound₃ + rename_i attr_map hsound₃ hsound₄ hsound₅ + cases attrs + case nil => + cases hinv + simp [evaluate_empty_record, Map.make, List.canonicalize] at h₅ + rw [← h₅] + apply WellFormed.record + case _ => + intros k hcontains + simp [Map.contains, Map.find?, List.find?] at hcontains + case _ => + intros k v qty hfind_value hfind_type + simp [Map.find?, List.find?] at hfind_value + case _ => + intros k qty hfind + simp [Map.find?, List.find?] at hfind + case cons attr attrs => + sorry + +theorem evaluates_to_well_formed {x : Expr} {v : Value} {ty: CedarType} {request : Request} {entities: Entities} {env : Environment} {c₁ c₂ : Capabilities} {l₁ : Level} + (h₁ : l₁ < .infinite) + (h₂ : RequestAndEntitiesMatchEnvironmentLeveled env request entities l₁) + (h₃ : typeOf x c₁ env (l₁ == .infinite) = Except.ok (ty, c₂)) + (h₅ : evaluate x request entities = Except.ok v) + (h_caps : CapabilitiesInvariant c₁ request entities) + : + entities ⊢ v : ty + := by + have hreq : RequestAndEntitiesMatchEnvironment env request entities := by + apply request_and_entity_match_level_implies_regular + assumption + cases x + case lit p => + apply evaluates_to_well_formed_lit + repeat assumption + case var x => + apply evaluates_to_well_formed_var + repeat assumption + case ite cond cons alt => + apply evaluates_to_well_formed_ite + repeat assumption + case ih₁ => + simp [EvaluatesToWellFormed] + intros + apply evaluates_to_well_formed + repeat assumption + case ih₂ => + simp [EvaluatesToWellFormed] + intros + apply evaluates_to_well_formed + repeat assumption + case _ lhs rhs => -- and + apply evaluates_to_well_formed_and + repeat assumption + case _ lhs rhs => -- or + apply evaluates_to_well_formed_or + repeat assumption + case _ op expr => -- unary ops + apply evaluates_to_well_formed_unary + repeat assumption + case _ o lhs rhs => -- binary ops + apply evaluates_to_well_formed_binary + repeat assumption + case _ => --get attr + apply evaluates_to_well_formed_getAttr + repeat assumption + simp [EvaluatesToWellFormed] + intros + apply evaluates_to_well_formed + repeat assumption + case hasAttr expr attr => + apply levels_correct_hasAttr + repeat assumption + case set members => + apply levels_correct_set + repeat assumption + intros + simp [EvaluatesToWellFormed] + intros + apply evaluates_to_well_formed + repeat assumption + case record attrs => + apply evaluates_to_well_formed_record + repeat assumption + intros _ e _ + simp [EvaluatesToWellFormed] + intros + apply evaluates_to_well_formed + repeat assumption + case call xfn args => + apply evaluates_to_well_formed_call + repeat assumption +termination_by (sizeOf x) +decreasing_by + all_goals simp_wf + all_goals (try omega) + case _ => + rename_i heq _ _ _ _ _ _ _ _ _ _ _ _ _ + rw [heq] + simp + omega + case _ => + rename_i heq _ _ _ _ _ _ _ _ _ _ _ _ _ + rw [heq] + simp + omega + case _ => + rename_i heq _ _ _ _ _ _ _ _ _ _ _ _ _ + rw [heq] + simp + omega + case _ => + rename_i members heq e hin _ _ _ _ _ _ _ _ _ _ _ _ _ + rw [heq] + have hstep : sizeOf e < sizeOf members := by + apply List.in_lists_means_smaller e members + assumption + simp + omega + case _ => + rename_i members heq a hin _ _ _ _ _ _ _ _ _ _ _ _ _ + rw [heq] + have step1 : sizeOf e < sizeOf (a,e) := by + simp + omega + have step2 : sizeOf (a,e) < sizeOf members := by + apply List.in_lists_means_smaller + assumption + simp + omega + +end Cedar.Thm diff --git a/cedar-lean/Cedar/Thm/Validation/RequestEntityValidation.lean b/cedar-lean/Cedar/Thm/Validation/RequestEntityValidation.lean index 8979902bc..7c602e5db 100644 --- a/cedar-lean/Cedar/Thm/Validation/RequestEntityValidation.lean +++ b/cedar-lean/Cedar/Thm/Validation/RequestEntityValidation.lean @@ -60,7 +60,7 @@ theorem instance_of_type_refl (v : Value) (ty : CedarType) : contradiction | entityUID uid => cases ty - case entity ety => + case entity ety l => apply InstanceOfType.instance_of_entity uid ety apply instance_of_entity_type_refl assumption @@ -250,10 +250,10 @@ theorem request_and_entities_match_env (env : Environment) (request : Request) ( exact instance_of_entity_schema_refl entities env.ets h₂ exact instance_of_action_schema_refl entities env.acts h₁ -theorem request_and_entities_validate_implies_match_schema (schema : Schema) (request : Request) (entities : Entities) : - validateRequest schema request = .ok () → - validateEntities schema entities = .ok () → - RequestAndEntitiesMatchSchema schema request entities +theorem request_and_entities_validate_implies_match_schema (schema : Schema) (request : Request) (entities : Entities) (l : Level) : + validateRequest schema request l = .ok () → + validateEntities schema entities l = .ok () → + RequestAndEntitiesMatchSchema schema request entities l := by intro h₀ h₁ simp only [RequestAndEntitiesMatchSchema] @@ -266,5 +266,5 @@ theorem request_and_entities_validate_implies_match_schema (schema : Schema) (re exact h₀ apply request_and_entities_match_env exact h₂ - have h₃ := List.forM_ok_implies_all_ok schema.toEnvironments (fun x => entitiesMatchEnvironment x entities) h₁ env h₀ + have h₃ := List.forM_ok_implies_all_ok (schema.toEnvironments l) (fun x => entitiesMatchEnvironment x entities) h₁ env h₀ simp only [h₃] diff --git a/cedar-lean/Cedar/Thm/Validation/Typechecker.lean b/cedar-lean/Cedar/Thm/Validation/Typechecker.lean index 68afedecc..145762695 100644 --- a/cedar-lean/Cedar/Thm/Validation/Typechecker.lean +++ b/cedar-lean/Cedar/Thm/Validation/Typechecker.lean @@ -43,10 +43,10 @@ produces a value of the returned type or (2) it returns an error of type `entityDoesNotExist`, `extensionError`, or `arithBoundsError`. Both options are encoded in the `EvaluatesTo` predicate. -/ -theorem type_of_is_sound {e : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} : +theorem type_of_is_sound {e : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} : CapabilitiesInvariant c₁ request entities → RequestAndEntitiesMatchEnvironment env request entities → - typeOf e c₁ env = .ok (ty, c₂) → + typeOf e c₁ env (l == .infinite) = .ok (ty, c₂) → GuardedCapabilitiesInvariant e c₂ request entities ∧ ∃ (v : Value), EvaluatesTo e request entities v ∧ InstanceOfType v ty := by diff --git a/cedar-lean/Cedar/Thm/Validation/Typechecker/And.lean b/cedar-lean/Cedar/Thm/Validation/Typechecker/And.lean index 6848d9d29..7af9ae3a0 100644 --- a/cedar-lean/Cedar/Thm/Validation/Typechecker/And.lean +++ b/cedar-lean/Cedar/Thm/Validation/Typechecker/And.lean @@ -26,20 +26,20 @@ open Cedar.Data open Cedar.Spec open Cedar.Validation -theorem type_of_and_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.and x₁ x₂) c env = Except.ok (ty, c')) : +theorem type_of_and_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.and x₁ x₂) c env (l == Level.infinite) = Except.ok (ty, c')) : ∃ bty₁ c₁, - typeOf x₁ c env = .ok (.bool bty₁, c₁) ∧ + typeOf x₁ c env (l == Level.infinite) = .ok (.bool bty₁, c₁) ∧ if bty₁ = BoolType.ff then ty = .bool BoolType.ff ∧ c' = ∅ else ∃ bty₂ c₂, - typeOf x₂ (c ∪ c₁) env = .ok (.bool bty₂, c₂) ∧ + typeOf x₂ (c ∪ c₁) env (l == Level.infinite) = .ok (.bool bty₂, c₂) ∧ if bty₂ = BoolType.ff then ty = .bool BoolType.ff ∧ c' = ∅ else ty = .bool (lubBool bty₁ bty₂) ∧ c' = c₁ ∪ c₂ := by simp [typeOf] at h₁ - cases h₂ : typeOf x₁ c env <;> simp [h₂] at * + cases h₂ : typeOf x₁ c env (l == Level.infinite) <;> simp [h₂] at * rename_i res₁ simp [typeOfAnd] at h₁ split at h₁ <;> simp [ok, err] at h₁ @@ -55,7 +55,7 @@ theorem type_of_and_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : En have ⟨hty₁, hc₁⟩ := h₄ simp [←hty₁, ←hc₁] split ; contradiction - cases h₄ : typeOf x₂ (c ∪ res₁.snd) env <;> simp [h₄] at * + cases h₄ : typeOf x₂ (c ∪ res₁.snd) env (l == Level.infinite) <;> simp [h₄] at * rename_i res₂ split at h₁ <;> simp at h₁ <;> have ⟨hty, hc⟩ := h₁ <;> subst hty hc @@ -72,10 +72,10 @@ theorem type_of_and_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : En · simp [h₆] · rfl -theorem type_of_and_is_sound {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_and_is_sound {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l: Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.and x₁ x₂) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.and x₁ x₂) c₁ env (l == Level.infinite) = Except.ok (ty, c₂)) (ih₁ : TypeOfIsSound x₁) (ih₂ : TypeOfIsSound x₂) : GuardedCapabilitiesInvariant (Expr.and x₁ x₂) c₂ request entities ∧ diff --git a/cedar-lean/Cedar/Thm/Validation/Typechecker/Basic.lean b/cedar-lean/Cedar/Thm/Validation/Typechecker/Basic.lean index c1a579384..3c7042d6a 100644 --- a/cedar-lean/Cedar/Thm/Validation/Typechecker/Basic.lean +++ b/cedar-lean/Cedar/Thm/Validation/Typechecker/Basic.lean @@ -57,6 +57,41 @@ def EvaluatesTo (e: Expr) (request : Request) (entities : Entities) (v : Value) evaluate e request entities = .error .arithBoundsError ∨ evaluate e request entities = .ok v +/-- +The leveled type soundness property says that if the typechecker assigns a type to an +expression, then it must be the case that the expression `EvaluatesTo` a value +of that type. The `EvaluatesTo` predicate covers the (obvious) case where +evaluation has no errors, but it also allows for errors of type +`extensionError`, and `arithBoundsError`. + +The typechecker cannot protect against these errors because they depend on runtime +arithmetic overflow errors. All other errors (`attrDoesNotExist` and `typeError`) can be +prevented statically. + +_Note_: Currently, `extensionError`s can also be ruled out at validation time +because the only extension functions that can error are constructors, and all +constructors are required to be applied to string literals, meaning that they +can be fully evaluated during validation. This is not guaranteed to be the case +in the future. + +_Note_: We plan to implement a range analysis that will be able to rule out +`arithBoundsError`s. +-/ +def EvaluatesToLeveled (e : Expr) (request : Request) (entities : Entities) (v : Value) : Prop := + evaluate e request entities = .error .extensionError ∨ + evaluate e request entities = .error .arithBoundsError ∨ + evaluate e request entities = .ok v + +theorem leveldSafeImpliesRegularSafe {e : Expr} {request : Request} {entities : Entities} {v : Value} : + EvaluatesToLeveled e request entities v → + EvaluatesTo e request entities v + := by + intros h + unfold EvaluatesToLeveled at h + unfold EvaluatesTo + simp [h] + + /-- On input to the typechecking function, for any (e,k) in the Capabilities, e is a record- or entity-typed expression that has key k. @@ -73,13 +108,22 @@ def GuardedCapabilitiesInvariant (e: Expr) (c: Capabilities) (request : Request) CapabilitiesInvariant c request entities def TypeOfIsSound (x₁ : Expr) : Prop := - ∀ {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities}, + ∀ {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level}, CapabilitiesInvariant c₁ request entities → RequestAndEntitiesMatchEnvironment env request entities → - typeOf x₁ c₁ env = Except.ok (ty, c₂) → + typeOf x₁ c₁ env (l == Level.infinite) = Except.ok (ty, c₂) → GuardedCapabilitiesInvariant x₁ c₂ request entities ∧ ∃ v, EvaluatesTo x₁ request entities v ∧ InstanceOfType v ty +def TypeOfIsSoundLeveled (x₁ : Expr) : Prop := + ∀ {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level}, + l < .infinite → + CapabilitiesInvariant c₁ request entities → + RequestAndEntitiesMatchEnvironmentLeveled env request entities l → + typeOf x₁ c₁ env (l == Level.infinite) = Except.ok (ty, c₂) → + GuardedCapabilitiesInvariant x₁ c₂ request entities ∧ + ∃ v, EvaluatesToLeveled x₁ request entities v ∧ InstanceOfType v ty + ----- Capability lemmas ----- theorem empty_capabilities_invariant (request : Request) (entities : Entities) : diff --git a/cedar-lean/Cedar/Thm/Validation/Typechecker/BinaryApp.lean b/cedar-lean/Cedar/Thm/Validation/Typechecker/BinaryApp.lean index bc38da042..47649fec9 100644 --- a/cedar-lean/Cedar/Thm/Validation/Typechecker/BinaryApp.lean +++ b/cedar-lean/Cedar/Thm/Validation/Typechecker/BinaryApp.lean @@ -27,25 +27,25 @@ open Cedar.Data open Cedar.Spec open Cedar.Validation -theorem type_of_eq_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.binaryApp .eq x₁ x₂) c env = Except.ok (ty, c')) : +theorem type_of_eq_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.binaryApp .eq x₁ x₂) c env (l == Level.infinite) = Except.ok (ty, c')) : c' = ∅ ∧ match x₁, x₂ with | .lit p₁, .lit p₂ => if p₁ = p₂ then ty = (.bool .tt) else ty = (.bool .ff) | _, _ => ∃ ty₁ c₁ ty₂ c₂, - typeOf x₁ c env = Except.ok (ty₁, c₁) ∧ - typeOf x₂ c env = Except.ok (ty₂, c₂) ∧ + typeOf x₁ c env (l == Level.infinite) = Except.ok (ty₁, c₁) ∧ + typeOf x₂ c env (l == Level.infinite) = Except.ok (ty₂, c₂) ∧ match ty₁ ⊔ ty₂ with | .some _ => ty = (.bool .anyBool) | .none => ty = (.bool .ff) ∧ - ∃ ety₁ ety₂, ty₁ = .entity ety₁ ∧ ty₂ = .entity ety₂ + ∃ ety₁ l₁ ety₂ l₂, ty₁ = .entity ety₁ l₁ ∧ ty₂ = .entity ety₂ l₂ := by simp [typeOf] at h₁ ; rename_i h₁' - cases h₂ : typeOf x₁ c env <;> simp [h₂] at h₁ - cases h₃ : typeOf x₂ c env <;> simp [h₃] at h₁ + cases h₂ : typeOf x₁ c env (l == Level.infinite) <;> simp [h₂] at h₁ + cases h₃ : typeOf x₂ c env (l == Level.infinite) <;> simp [h₃] at h₁ simp [typeOfBinaryApp, typeOfEq, ok, err] at h₁ rename_i tc₁ tc₂ split at h₁ @@ -72,7 +72,7 @@ theorem type_of_eq_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Env split case h_1 p₁ p₂ _ => specialize h₄ p₁ p₂ ; simp at h₄ - case h_2 ety₁ ety₂ _ true_is_instance_of_tt _ _ _ _ => + case h_2 ety₁ l₁ ety₂ l₂ _ true_is_instance_of_tt _ _ _ _ => exists tc₁.fst constructor · exists tc₁.snd @@ -81,13 +81,13 @@ theorem type_of_eq_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Env · exists tc₂.snd · simp [h₅] constructor - · exists ety₁ - · exists ety₂ + · exists ety₁ ; exists l₁ + · exists ety₂ ; exists l₂ -theorem no_entity_type_lub_implies_not_eq {v₁ v₂ : Value} {ety₁ ety₂ : EntityType} - (h₁ : InstanceOfType v₁ (CedarType.entity ety₁)) - (h₂ : InstanceOfType v₂ (CedarType.entity ety₂)) - (h₃ : (CedarType.entity ety₁ ⊔ CedarType.entity ety₂) = none) : +theorem no_entity_type_lub_implies_not_eq {v₁ v₂ : Value} {ety₁ ety₂ : EntityType} {l₁ l₂ : Level} + (h₁ : InstanceOfType v₁ (CedarType.entity ety₁ l₁)) + (h₂ : InstanceOfType v₂ (CedarType.entity ety₂ l₂)) + (h₃ : (CedarType.entity ety₁ l₁ ⊔ CedarType.entity ety₂ l₂) = none) : ¬ v₁ = v₂ := by by_contra h₄ ; subst h₄ @@ -99,10 +99,10 @@ theorem no_entity_type_lub_implies_not_eq {v₁ v₂ : Value} {ety₁ ety₂ : E subst h₄ h₅ contradiction -theorem type_of_eq_is_sound {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_eq_is_sound {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.binaryApp .eq x₁ x₂) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.binaryApp .eq x₁ x₂) c₁ env (l == Level.infinite) = Except.ok (ty, c₂)) (ih₁ : TypeOfIsSound x₁) (ih₂ : TypeOfIsSound x₂) : GuardedCapabilitiesInvariant (Expr.binaryApp .eq x₁ x₂) c₂ request entities ∧ @@ -141,7 +141,7 @@ theorem type_of_eq_is_sound {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : rw [hty] apply bool_is_instance_of_anyBool case h_2 heq => - have ⟨hty₀, ⟨ety₁, hty₁⟩, ⟨ety₂, hty₂⟩⟩ := hty ; clear hty + have ⟨hty₀, ⟨ety₁, l₁, hty₁⟩, ⟨ety₂, l₂, hty₂⟩⟩ := hty ; clear hty subst hty₀ hty₁ hty₂ have h₆ := no_entity_type_lub_implies_not_eq ih₃ ih₄ heq cases h₇ : v₁ == v₂ <;> @@ -149,17 +149,17 @@ theorem type_of_eq_is_sound {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : case false => exact false_is_instance_of_ff case true => contradiction -theorem type_of_int_cmp_inversion {op₂ : BinaryOp} {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} +theorem type_of_int_cmp_inversion {op₂ : BinaryOp} {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} (h₁ : op₂ = .less ∨ op₂ = .lessEq) - (h₂ : typeOf (Expr.binaryApp op₂ x₁ x₂) c env = Except.ok (ty, c')) : + (h₂ : typeOf (Expr.binaryApp op₂ x₁ x₂) c env (l == Level.infinite) = Except.ok (ty, c')) : c' = ∅ ∧ ty = .bool .anyBool ∧ - (∃ c₁, typeOf x₁ c env = Except.ok (.int, c₁)) ∧ - (∃ c₂, typeOf x₂ c env = Except.ok (.int, c₂)) + (∃ c₁, typeOf x₁ c env (l == Level.infinite) = Except.ok (.int, c₁)) ∧ + (∃ c₂, typeOf x₂ c env (l == Level.infinite)= Except.ok (.int, c₂)) := by simp [typeOf] at * - cases h₃ : typeOf x₁ c env <;> simp [h₃] at h₂ - cases h₄ : typeOf x₂ c env <;> simp [h₄] at h₂ + cases h₃ : typeOf x₁ c env (l == Level.infinite) <;> simp [h₃] at h₂ + cases h₄ : typeOf x₂ c env (l == Level.infinite) <;> simp [h₄] at h₂ rcases h₁ with h₁ | h₁ all_goals { subst h₁ @@ -172,11 +172,11 @@ theorem type_of_int_cmp_inversion {op₂ : BinaryOp} {x₁ x₂ : Expr} {c c' : · exists tc₂.snd ; simp [←h₆] } -theorem type_of_int_cmp_is_sound {op₂ : BinaryOp} {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_int_cmp_is_sound {op₂ : BinaryOp} {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₀ : op₂ = .less ∨ op₂ = .lessEq) (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.binaryApp op₂ x₁ x₂) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.binaryApp op₂ x₁ x₂) c₁ env (l == Level.infinite) = Except.ok (ty, c₂)) (ih₁ : TypeOfIsSound x₁) (ih₂ : TypeOfIsSound x₂) : GuardedCapabilitiesInvariant (Expr.binaryApp op₂ x₁ x₂) c₂ request entities ∧ @@ -206,17 +206,17 @@ theorem type_of_int_cmp_is_sound {op₂ : BinaryOp} {x₁ x₂ : Expr} {c₁ c apply bool_is_instance_of_anyBool } -theorem type_of_int_arith_inversion {op₂ : BinaryOp} {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} +theorem type_of_int_arith_inversion {op₂ : BinaryOp} {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} (h₁ : op₂ = .add ∨ op₂ = .sub ∨ op₂ = .mul) - (h₂ : typeOf (Expr.binaryApp op₂ x₁ x₂) c env = Except.ok (ty, c')) : + (h₂ : typeOf (Expr.binaryApp op₂ x₁ x₂) c env (l == Level.infinite) = Except.ok (ty, c')) : c' = ∅ ∧ ty = .int ∧ - (∃ c₁, typeOf x₁ c env = Except.ok (.int, c₁)) ∧ - (∃ c₂, typeOf x₂ c env = Except.ok (.int, c₂)) + (∃ c₁, typeOf x₁ c env (l == Level.infinite) = Except.ok (.int, c₁)) ∧ + (∃ c₂, typeOf x₂ c env (l == Level.infinite) = Except.ok (.int, c₂)) := by simp [typeOf] at * - cases h₃ : typeOf x₁ c env <;> simp [h₃] at h₂ - cases h₄ : typeOf x₂ c env <;> simp [h₄] at h₂ + cases h₃ : typeOf x₁ c env (l == Level.infinite) <;> simp [h₃] at h₂ + cases h₄ : typeOf x₂ c env (l == Level.infinite) <;> simp [h₄] at h₂ rcases h₁ with h₁ | h₁ | h₁ all_goals { subst h₁ @@ -230,11 +230,11 @@ theorem type_of_int_arith_inversion {op₂ : BinaryOp} {x₁ x₂ : Expr} {c c' · exists tc₂.snd ; simp [←h₂, ←h₆] } -theorem type_of_int_arith_is_sound {op₂ : BinaryOp} {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_int_arith_is_sound {op₂ : BinaryOp} {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₀ : op₂ = .add ∨ op₂ = .sub ∨ op₂ = .mul) (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.binaryApp op₂ x₁ x₂) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.binaryApp op₂ x₁ x₂) c₁ env (l == Level.infinite) = Except.ok (ty, c₂)) (ih₁ : TypeOfIsSound x₁) (ih₂ : TypeOfIsSound x₂) : GuardedCapabilitiesInvariant (Expr.binaryApp op₂ x₁ x₂) c₂ request entities ∧ @@ -271,18 +271,18 @@ theorem type_of_int_arith_is_sound {op₂ : BinaryOp} {x₁ x₂ : Expr} {c₁ c case none => exact type_is_inhabited CedarType.int case some => simp [InstanceOfType.instance_of_int] -theorem type_of_contains_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.binaryApp .contains x₁ x₂) c env = Except.ok (ty, c')) : +theorem type_of_contains_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.binaryApp .contains x₁ x₂) c env (l == .infinite) = Except.ok (ty, c')) : c' = ∅ ∧ ty = .bool .anyBool ∧ ∃ ty₁ ty₂, (ty₁ ⊔ ty₂).isSome ∧ - (∃ c₁, typeOf x₁ c env = Except.ok (.set ty₁, c₁)) ∧ - (∃ c₂, typeOf x₂ c env = Except.ok (ty₂, c₂)) + (∃ c₁, typeOf x₁ c env (l == .infinite) = Except.ok (.set ty₁, c₁)) ∧ + (∃ c₂, typeOf x₂ c env (l == .infinite) = Except.ok (ty₂, c₂)) := by simp [typeOf] at * - cases h₂ : typeOf x₁ c env <;> simp [h₂] at h₁ - cases h₃ : typeOf x₂ c env <;> simp [h₃] at h₁ + cases h₂ : typeOf x₁ c env (l == .infinite) <;> simp [h₂] at h₁ + cases h₃ : typeOf x₂ c env (l == .infinite) <;> simp [h₃] at h₁ simp [typeOfBinaryApp, err, ok] at h₁ split at h₁ <;> try contradiction simp [ifLubThenBool, err, ok] at h₁ @@ -296,10 +296,10 @@ theorem type_of_contains_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env · exists tc₁.snd · exists tc₂.snd -theorem type_of_contains_is_sound {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_contains_is_sound {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.binaryApp .contains x₁ x₂) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.binaryApp .contains x₁ x₂) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih₁ : TypeOfIsSound x₁) (ih₂ : TypeOfIsSound x₂) : GuardedCapabilitiesInvariant (Expr.binaryApp .contains x₁ x₂) c₂ request entities ∧ @@ -324,19 +324,19 @@ theorem type_of_contains_is_sound {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} simp [apply₂] apply bool_is_instance_of_anyBool -theorem type_of_containsA_inversion {op₂ : BinaryOp} {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} +theorem type_of_containsA_inversion {op₂ : BinaryOp} {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} (h₁ : op₂ = .containsAll ∨ op₂ = .containsAny) - (h₂ : typeOf (Expr.binaryApp op₂ x₁ x₂) c env = Except.ok (ty, c')) : + (h₂ : typeOf (Expr.binaryApp op₂ x₁ x₂) c env (l == .infinite) = Except.ok (ty, c')) : c' = ∅ ∧ ty = .bool .anyBool ∧ ∃ ty₁ ty₂, (ty₁ ⊔ ty₂).isSome ∧ - (∃ c₁, typeOf x₁ c env = Except.ok (.set ty₁, c₁)) ∧ - (∃ c₂, typeOf x₂ c env = Except.ok (.set ty₂, c₂)) + (∃ c₁, typeOf x₁ c env (l == .infinite) = Except.ok (.set ty₁, c₁)) ∧ + (∃ c₂, typeOf x₂ c env (l == .infinite) = Except.ok (.set ty₂, c₂)) := by simp [typeOf] at * - cases h₃ : typeOf x₁ c env <;> simp [h₃] at h₂ - cases h₄ : typeOf x₂ c env <;> simp [h₄] at h₂ + cases h₃ : typeOf x₁ c env (l == .infinite) <;> simp [h₃] at h₂ + cases h₄ : typeOf x₂ c env (l == .infinite) <;> simp [h₄] at h₂ rcases h₁ with h₁ | h₁ all_goals { subst h₁ @@ -354,11 +354,11 @@ theorem type_of_containsA_inversion {op₂ : BinaryOp} {x₁ x₂ : Expr} {c c' } -theorem type_of_containsA_is_sound {op₂ : BinaryOp} {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_containsA_is_sound {op₂ : BinaryOp} {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₀ : op₂ = .containsAll ∨ op₂ = .containsAny) (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.binaryApp op₂ x₁ x₂) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.binaryApp op₂ x₁ x₂) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih₁ : TypeOfIsSound x₁) (ih₂ : TypeOfIsSound x₂) : GuardedCapabilitiesInvariant (Expr.binaryApp op₂ x₁ x₂) c₂ request entities ∧ @@ -388,32 +388,146 @@ theorem type_of_containsA_is_sound {op₂ : BinaryOp} {x₁ x₂ : Expr} {c₁ c apply bool_is_instance_of_anyBool } -theorem type_of_mem_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.binaryApp .mem x₁ x₂) c env = Except.ok (ty, c')) : +theorem type_of_mem_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.binaryApp .mem x₁ x₂) c env (l == .infinite) = Except.ok (ty, c')) : c' = ∅ ∧ - ∃ (ety₁ ety₂ : EntityType), - (∃ c₁, typeOf x₁ c env = Except.ok (.entity ety₁, c₁)) ∧ + ∃ (ety₁ ety₂ : EntityType) (l₁ l₂ : Level), + (∃ c₁, typeOf x₁ c env (l == .infinite) = Except.ok (.entity ety₁ l₁, c₁)) ∧ (∃ c₂, - (typeOf x₂ c env = Except.ok (.entity ety₂, c₂) ∧ - ty = .bool (typeOfInₑ ety₁ ety₂ x₁ x₂ env)) ∨ - (typeOf x₂ c env = Except.ok (.set (.entity ety₂), c₂) ∧ - ty = .bool (typeOfInₛ ety₁ ety₂ x₁ x₂ env))) + (typeOf x₂ c env (l == .infinite) = Except.ok (.entity ety₂ l₂, c₂) ∧ + .ok (ty, ∅) = typeOfInₑ ety₁ ety₂ l₁ x₁ x₂ env) ∨ + (typeOf x₂ c env (l == .infinite) = Except.ok (.set (.entity ety₂ l₂), c₂) ∧ + .ok (ty, ∅) = typeOfInₛ ety₁ ety₂ l₁ x₁ x₂ env)) := by simp [typeOf] at h₁ - cases h₂ : typeOf x₁ c env <;> simp [h₂] at h₁ - cases h₃ : typeOf x₂ c env <;> simp [h₃] at h₁ + cases h₂ : typeOf x₁ c env (l == .infinite) <;> simp [h₂] at h₁ + cases h₃ : typeOf x₂ c env (l == .infinite) <;> simp [h₃] at h₁ simp [typeOfBinaryApp, ok] at h₁ - split at h₁ <;> try { contradiction } + split at h₁ + <;> try { contradiction } + <;> rename_i tc₁ tc₂ op ty₁ ty₂ ety₁ l₁ ety₂ l₂ _ heq₁ heq₂ all_goals { - simp only [Except.ok.injEq, Prod.mk.injEq] at h₁ - simp [h₁] - rename_i tc₁ tc₂ _ _ _ ety₁ ety₂ _ h₄ h₅ + try unfold typeOfInₑ at h₁ + try unfold typeOfInₛ at h₁ + + split at h₁ <;> try contradiction + rename_i hlt + injection h₁ + rename_i heq₃ + injection heq₃ + rename_i heq₃ heq₄ + rw [← heq₄] + constructor + simp exists ety₁ + exists ety₂ + exists l₁ + exists l₂ constructor - · exists tc₁.snd ; simp [←h₄] - · exists ety₂, tc₂.snd ; simp [←h₅, h₁] + case isTrue.right.left => + exists tc₁.snd + simp + rw [← heq₁] + case isTrue.right.right => + exists tc₂.snd + try ( + apply Or.inl ; + constructor ; + rw [← heq₂] + rw [← heq₃] + unfold typeOfInₑ + rw [if_pos] + simp [ok, Functor.map, Except.map] + apply hlt + ) + try ( + apply Or.inr + constructor + rw [← heq₂] + unfold typeOfInₛ + rw [if_pos] + simp [Prod.fst, Functor.map, Except.map, ok] + rw [heq₃] + apply hlt + ) } +theorem type_of_mem_inversion_finite {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.binaryApp .mem x₁ x₂) c env (l == .infinite) = Except.ok (ty, c')) : + c' = ∅ ∧ + ∃ (ety₁ ety₂ : EntityType) (l₁ l₂ : Level), + .finite 0 < l₁ ∧ + (∃ c₁, typeOf x₁ c env (l == .infinite) = Except.ok (.entity ety₁ l₁, c₁)) ∧ + (∃ c₂, + (typeOf x₂ c env (l == .infinite) = Except.ok (.entity ety₂ l₂, c₂) ∧ + .ok (ty, ∅) = typeOfInₑ ety₁ ety₂ l₁ x₁ x₂ env) ∨ + (typeOf x₂ c env (l == .infinite) = Except.ok (.set (.entity ety₂ l₂), c₂) ∧ + .ok (ty, ∅) = typeOfInₛ ety₁ ety₂ l₁ x₁ x₂ env)) + := by + have h := type_of_mem_inversion h₁ + replace ⟨h₂, ety₁, ety₂, l₁, l₂, h⟩ := h + constructor <;> try assumption + exists ety₁ + exists ety₂ + exists l₁ + exists l₂ + have ⟨h₃, h₄⟩ := h + clear h + replace ⟨c₁, h₃⟩ := h₃ + replace ⟨c₂, h₄⟩ := h₄ + cases h₄ + case _ h₄ => + constructor + case _ => + simp [typeOf, h₃, h₄, typeOfBinaryApp, typeOfInₑ] at h₁ + split at h₁ + case _ hlevel => + assumption + case _ hlevel => + simp [err] at h₁ + case _ => + constructor + case _ => + exists c₁ + case _ => + exists c₂ + unfold EmptyCollection.emptyCollection at h₄ + unfold List.instEmptyCollection at h₄ + have ⟨_, _⟩ := h₄ + simp only [EmptyCollection.emptyCollection, and_self, Except.ok.injEq, Prod.mk.injEq, + and_true, false_and, or_false, h₄] + case _ h₄ => + replace ⟨h₄, h₅⟩ := h₄ + constructor + case _ => + simp [typeOf, h₂, h₃, h₄, h₅, typeOfBinaryApp, typeOfInₛ] at h₁ + split at h₁ + case _ hlevel => + assumption + case _ hlevel => + simp [err] at h₁ + case _ => + constructor + case _ => + exists c₁ + case _ => + exists c₂ + simp only [Except.ok.injEq, Prod.mk.injEq, and_true, List.empty_eq, false_and, true_and, + false_or, h₄] + simp [EmptyCollection.emptyCollection, List.instEmptyCollection] at h₅ + assumption + + + + + + + + + + + + theorem entityUID?_some_implies_entity_lit {x : Expr} {euid : EntityUID} (h₁ : entityUID? x = some euid) : x = Expr.lit (.entityUID euid) @@ -498,16 +612,17 @@ theorem action_type_in_eq_action_inₑ (euid₁ euid₂ : EntityUID) {env : Envi simp [inₑ, ActionSchema.descendentOf, h₃, Entities.ancestorsOrEmpty, h₁₁] rcases h₄ : euid₁ == euid₂ <;> simp at h₄ <;> simp [h₄, h₁₂] -theorem type_of_mem_is_soundₑ {x₁ x₂ : Expr} {c₁ c₁' c₂' : Capabilities} {env : Environment} {request : Request} {entities : Entities} {ety₁ ety₂ : EntityType} +theorem type_of_mem_is_soundₑ {x₁ x₂ : Expr} {c₁ c₁' c₂' : Capabilities} {env : Environment} {request : Request} {entities : Entities} {ety₁ ety₂ : EntityType} {ty : CedarType} {l l₁ l₂ : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf x₁ c₁ env = Except.ok (CedarType.entity ety₁, c₁')) - (h₄ : typeOf x₂ c₁ env = Except.ok (CedarType.entity ety₂, c₂')) + (h₃ : typeOf x₁ c₁ env (l == .infinite) = Except.ok (CedarType.entity ety₁ l₁, c₁')) + (h₄ : typeOf x₂ c₁ env (l == .infinite) = Except.ok (CedarType.entity ety₂ l₂, c₂')) + (h₅ : .ok (ty,∅) = typeOfInₑ ety₁ ety₂ l₁ x₁ x₂ env) (ih₁ : TypeOfIsSound x₁) (ih₂ : TypeOfIsSound x₂) : ∃ v, EvaluatesTo (Expr.binaryApp BinaryOp.mem x₁ x₂) request entities v ∧ - InstanceOfType v (CedarType.bool (typeOfInₑ ety₁ ety₂ x₁ x₂ env)) + InstanceOfType v ty := by have ⟨_, v₁, hev₁, hty₁⟩ := ih₁ h₁ h₂ h₃ have ⟨_, v₂, hev₂, hty₂⟩ := ih₂ h₁ h₂ h₄ @@ -526,17 +641,32 @@ theorem type_of_mem_is_soundₑ {x₁ x₂ : Expr} {c₁ c₁' c₂' : Capabilit replace ⟨euid₂, hty₂, hty₂'⟩ := hty₂ subst hty₂ hty₂' simp [apply₂] + + unfold typeOfInₑ at h₅ + split at h₅ <;> try contradiction + rename_i hgt + simp [ok] at h₅ + subst h₅ apply InstanceOfType.instance_of_bool simp [InstanceOfBoolType] split <;> try simp only rename_i b bty h₇ h₈ h₉ simp [typeOfInₑ] at * have ⟨_, hents, hacts⟩ := h₂ ; clear h₂ + unfold typeOfInₑ.type at h₇ h₈ h₉ cases hₐ : actionUID? x₁ env.acts <;> simp [hₐ] at h₇ h₈ h₉ case none => - cases hin : EntitySchema.descendentOf env.ets euid₁.ty euid₂.ty <;> - simp [hin] at h₇ h₈ h₉ - simp [entity_type_in_false_implies_inₑ_false hents hin] at h₉ + cases hin : EntitySchema.descendentOf env.ets euid₁.ty euid₂.ty + case _ => + rw [entity_type_in_false_implies_inₑ_false] at h₉ + have h₁₀ := h₉ (by rfl) + rw [hin] at h₁₀ + contradiction + apply hents + apply hin + case _ => + rw [h₇] at hin + contradiction case some => cases he : entityUID? x₂ <;> simp [he] at h₇ h₈ h₉ case none => @@ -553,8 +683,8 @@ theorem type_of_mem_is_soundₑ {x₁ x₂ : Expr} {c₁ c₁' c₂' : Capabilit simp [h₁₀] at h₈ h₉ cases heq : ActionSchema.descendentOf env.acts auid euid <;> simp [heq] at h₈ h₉ -theorem entity_set_type_implies_set_of_entities {vs : List Value} {ety : EntityType} - (h₁ : InstanceOfType (Value.set (Set.mk vs)) (CedarType.set (CedarType.entity ety))) : +theorem entity_set_type_implies_set_of_entities {vs : List Value} {ety : EntityType} {l : Level} + (h₁ : InstanceOfType (Value.set (Set.mk vs)) (CedarType.set (CedarType.entity ety l))) : ∃ (euids : List EntityUID), vs.mapM Value.asEntityUID = Except.ok euids ∧ ∀ euid, euid ∈ euids → euid.ty = ety @@ -572,7 +702,7 @@ theorem entity_set_type_implies_set_of_entities {vs : List Value} {ety : EntityT subst h₂ rw [Value.asEntityUID] ; simp only [Except.bind_ok] rw [List.mapM'_eq_mapM] - have h₃ : InstanceOfType (Value.set (Set.mk tl)) (CedarType.set (CedarType.entity ety)) := by + have h₃ : InstanceOfType (Value.set (Set.mk tl)) (CedarType.set (CedarType.entity ety l)) := by apply InstanceOfType.instance_of_set intro v h₃ apply h₁ v @@ -734,17 +864,61 @@ theorem action_type_in_eq_action_inₛ {auid : EntityUID} {euids euids' : List E case neg => simp [inₑ, Entities.ancestorsOrEmpty, hl₁, hr₁, h₅] -theorem type_of_mem_is_soundₛ {x₁ x₂ : Expr} {c₁ c₁' c₂' : Capabilities} {env : Environment} {request : Request} {entities : Entities} {ety₁ ety₂ : EntityType} +theorem entityUIDS?_of_lits {euids : List EntityUID} : + entityUIDs? (Expr.set (List.map (Expr.lit ∘ Prim.entityUID) euids)) = some euids + := by + cases euids + case nil => + simp [entityUIDs?] + case cons head tail => + + simp [entityUIDs?, Option.bind] + split + case _ h => + simp [entityUID?] at h + case _ h => + simp + have ih := @entityUIDS?_of_lits tail + split + case _ h' => + simp [h', entityUIDs?] at ih + case _ h' => + simp + simp [entityUIDs?] at ih + simp [ih] at h' + simp [entityUID?] at h + simp [h, h'] + +theorem type_of_mem_is_soundₛ {x₁ x₂ : Expr} {c₁ c₁' c₂' : Capabilities} {env : Environment} {request : Request} {entities : Entities} {ety₁ ety₂ : EntityType} {l l₁ l₂ : Level} {ty : CedarType} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf x₁ c₁ env = Except.ok (CedarType.entity ety₁, c₁')) - (h₄ : typeOf x₂ c₁ env = Except.ok (CedarType.set (CedarType.entity ety₂), c₂')) + (h₃ : typeOf x₁ c₁ env (l == .infinite) = Except.ok (CedarType.entity ety₁ l₁, c₁')) + (h₄ : typeOf x₂ c₁ env (l == .infinite) = Except.ok (CedarType.set (CedarType.entity ety₂ l₂), c₂')) + (h₅ : .ok (ty, ∅) = typeOfInₛ ety₁ ety₂ l₁ x₁ x₂ env) (ih₁ : TypeOfIsSound x₁) (ih₂ : TypeOfIsSound x₂) : ∃ v, EvaluatesTo (Expr.binaryApp BinaryOp.mem x₁ x₂) request entities v ∧ - InstanceOfType v (CedarType.bool (typeOfInₛ ety₁ ety₂ x₁ x₂ env)) + InstanceOfType v ty := by + have hbool : ∃ b, ty = .bool b := by + unfold typeOfInₛ at h₅ + split at h₅ <;> try contradiction + simp [ok] at h₅ + cases ty <;> try contradiction + rename_i btype + unfold typeOfInₛ.type at h₅ + split at h₅ <;> split at h₅ + case _ => + exists BoolType.tt + case _ => + exists BoolType.ff + case _ => + exists BoolType.anyBool + case _ => + exists BoolType.ff + have ⟨boolType, hbool⟩ := hbool + subst hbool have ⟨_, v₁, hev₁, hty₁⟩ := ih₁ h₁ h₂ h₃ have ⟨_, v₂, hev₂, hty₂⟩ := ih₂ h₁ h₂ h₄ simp only [EvaluatesTo] at * @@ -764,30 +938,41 @@ theorem type_of_mem_is_soundₛ {x₁ x₂ : Expr} {c₁ c₁' c₂' : Capabilit simp only [Set.mapOrErr, Set.elts] have ⟨euids, h₇, hty₇⟩ := entity_set_type_implies_set_of_entities hty₂ simp only [h₇, Except.bind_ok, Except.ok.injEq, false_or, exists_eq_left'] + rename_i h₈ apply InstanceOfType.instance_of_bool simp only [InstanceOfBoolType] split <;> try simp only rename_i h₈ h₉ h₁₀ have ⟨_, hents, hacts⟩ := h₂ ; clear h₂ - simp only [typeOfInₛ, List.any_eq_true, imp_false] at * - cases ha : actionUID? x₁ env.acts <;> - simp only [ha, ite_eq_left_iff, Bool.not_eq_true, imp_false, Bool.not_eq_false, - ite_eq_right_iff] at h₈ h₉ h₁₀ + simp only [List.any_eq_true, imp_false] at * + cases ha : actionUID? x₁ env.acts case none => - cases hin : EntitySchema.descendentOf env.ets euid.ty ety₂ <;> - simp only [hin, Bool.false_eq_true, ↓reduceIte, not_false_eq_true, implies_true, imp_false, - Bool.not_eq_false, Bool.true_eq_false] at h₈ h₉ h₁₀ - simp only [entity_type_in_false_implies_inₛ_false hents hin hty₇, - Bool.false_eq_true] at h₁₀ + rename_i hok _ + cases hin : EntitySchema.descendentOf env.ets euid.ty ety₂ + <;> simp [typeOfInₛ, typeOfInₛ.type, ha, hin] at hok + <;> split at hok + <;> simp [ok, err] at hok + case false.isTrue => + apply h₁₀ + apply entity_type_in_false_implies_inₛ_false + repeat assumption + case true.isTrue => + apply h₈ + assumption case some => - cases he : entityUIDs? x₂ <;> - simp only [he, ite_eq_left_iff, not_exists, not_and, Bool.not_eq_true, imp_false, - Classical.not_forall, not_imp, Bool.not_eq_false, ite_eq_right_iff] at h₈ h₉ h₁₀ + cases he : entityUIDs? x₂ case none => - cases hin : EntitySchema.descendentOf env.ets euid.ty ety₂ <;> - simp only [hin, Bool.false_eq_true, ↓reduceIte, not_false_eq_true, implies_true, imp_false, - Bool.not_eq_false, Bool.true_eq_false] at h₈ h₉ h₁₀ - simp only [entity_type_in_false_implies_inₛ_false hents hin hty₇, Bool.false_eq_true] at h₁₀ + cases hin : EntitySchema.descendentOf env.ets euid.ty ety₂ + <;> rename_i hok _ euid + <;> simp [typeOfInₛ, typeOfInₛ.type] at hok + <;> split at hok <;> simp [ha,he,hin,ok,err] at hok + case false => + apply h₁₀ + apply entity_type_in_false_implies_inₛ_false + repeat assumption + case true => + apply h₈ + assumption case some => replace ⟨ha, hac⟩ := actionUID?_some_implies_action_lit ha subst ha @@ -803,34 +988,82 @@ theorem type_of_mem_is_soundₛ {x₁ x₂ : Expr} {c₁ c₁' c₂' : Capabilit exists_prop, false_implies, true_implies, false_iff, true_iff, not_exists, not_and, Bool.not_eq_true] at h₉ h₁₀ h₁₂ case false => - replace ⟨euid', h₁₀⟩ := h₁₀ - specialize h₁₂ euid' h₁₀.left - simp only [h₁₀.right, Bool.true_eq_false] at h₁₂ + simp [typeOfInₛ, typeOfInₛ.type] at h₈ + split at h₈ <;> simp [ok,err] at h₈ + case _ => + clear h₉ + split at h₈ <;> split at h₈ + case _ => + subst h₈ + rename_i heq₁ heq₂ h + simp [actionUID?, entityUID?] at heq₁ + replace ⟨heq₁, heq₃⟩ := heq₁ + rw [entityUIDS?_of_lits] at heq₂ + injection heq₂ + rename_i heq₂ + have ⟨x, h₁, h₂⟩ := h + subst heq₃ + subst heq₂ + rw [h₁₂] at h₂ + contradiction + repeat assumption + case _ => + apply h₁₀ + assumption + case _ => + rename_i hcontra _ _ _ _ _ _ _ + apply hcontra + assumption + case _ => + apply h₁₀ + assumption case true => replace ⟨euid', h₁₂⟩ := h₁₂ - specialize h₉ euid' h₁₂.left - simp only [h₁₂.right, Bool.true_eq_false] at h₉ + simp [typeOfInₛ, typeOfInₛ.type, err] at h₈ + split at h₈ <;> try assumption + simp [actionUID?, entityUID?, hac, entityUIDS?_of_lits] at h₈ + split at h₈ + case _ => + simp [ok] at h₈ + apply h₉ + assumption + case _ => + rename_i contra + apply contra + exists euid' -theorem type_of_mem_is_sound {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_mem_is_sound {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.binaryApp .mem x₁ x₂) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.binaryApp .mem x₁ x₂) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih₁ : TypeOfIsSound x₁) (ih₂ : TypeOfIsSound x₂) : GuardedCapabilitiesInvariant (Expr.binaryApp .mem x₁ x₂) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.binaryApp .mem x₁ x₂) request entities v ∧ InstanceOfType v ty := by - have ⟨hc, ety₁, ety₂, ⟨c₁', h₄⟩ , c₂', h₅⟩ := type_of_mem_inversion h₃ + have ⟨hc, ety₁, ety₂, l₁, l₂, ⟨c₁', h₄⟩ , c₂', h₅⟩ := type_of_mem_inversion h₃ subst hc apply And.intro empty_guarded_capabilities_invariant - rcases h₅ with ⟨h₅, h₆⟩ | ⟨h₅, h₆⟩ <;> subst h₆ - · exact type_of_mem_is_soundₑ h₁ h₂ h₄ h₅ ih₁ ih₂ - · exact type_of_mem_is_soundₛ h₁ h₂ h₄ h₅ ih₁ ih₂ + rcases h₅ with ⟨h₅, h₆⟩ | ⟨h₅, h₆⟩ --<;> subst h₆ + case _ => + have heq : .ok (ty, ∅) = typeOfInₑ ety₁ ety₂ l₁ x₁ x₂ env := by + simp [typeOfInₑ] at h₆ + split at h₆ <;> simp [ok,err] at h₆ + simp only [typeOfInₑ.type, List.empty_eq, typeOfInₑ, gt_iff_lt, ↓reduceIte, *] + rfl + exact type_of_mem_is_soundₑ h₁ h₂ h₄ h₅ heq ih₁ ih₂ + case _ => + have heq : .ok (ty, ∅) = typeOfInₛ ety₁ ety₂ l₁ x₁ x₂ env := by + simp [typeOfInₛ] at h₆ + split at h₆ <;> simp [ok,err] at h₆ + simp only [typeOfInₛ.type, List.empty_eq, typeOfInₛ, gt_iff_lt, ↓reduceIte, *] + rfl + exact type_of_mem_is_soundₛ h₁ h₂ h₄ h₅ heq ih₁ ih₂ -theorem type_of_binaryApp_is_sound {op₂ : BinaryOp} {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_binaryApp_is_sound {op₂ : BinaryOp} {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.binaryApp op₂ x₁ x₂) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.binaryApp op₂ x₁ x₂) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih₁ : TypeOfIsSound x₁) (ih₂ : TypeOfIsSound x₂) : GuardedCapabilitiesInvariant (Expr.binaryApp op₂ x₁ x₂) c₂ request entities ∧ diff --git a/cedar-lean/Cedar/Thm/Validation/Typechecker/Call.lean b/cedar-lean/Cedar/Thm/Validation/Typechecker/Call.lean index 0c98e400b..8504db488 100644 --- a/cedar-lean/Cedar/Thm/Validation/Typechecker/Call.lean +++ b/cedar-lean/Cedar/Thm/Validation/Typechecker/Call.lean @@ -26,8 +26,8 @@ open Cedar.Data open Cedar.Spec open Cedar.Validation -theorem type_of_call_decimal_inversion {xs : List Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.call .decimal xs) c env = Except.ok (ty, c')) : +theorem type_of_call_decimal_inversion {xs : List Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.call .decimal xs) c env (l == .infinite) = Except.ok (ty, c')) : ty = .ext .decimal ∧ c' = ∅ ∧ ∃ (s : String), @@ -35,7 +35,7 @@ theorem type_of_call_decimal_inversion {xs : List Expr} {c c' : Capabilities} {e (Cedar.Spec.Ext.Decimal.decimal s).isSome := by simp [typeOf] at h₁ - cases h₂ : List.mapM₁ xs fun x => justType (typeOf x.val c env) <;> + cases h₂ : List.mapM₁ xs fun x => justType (typeOf x.val c env (l == .infinite)) <;> simp [h₂] at h₁ rename_i tys simp [typeOfCall, typeOfConstructor] at h₁ @@ -46,8 +46,8 @@ theorem type_of_call_decimal_inversion {xs : List Expr} {c c' : Capabilities} {e rename_i h₃ simp [h₃] -theorem type_of_call_decimal_is_sound {xs : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} - (h₁ : typeOf (Expr.call .decimal xs) c₁ env = Except.ok (ty, c₂)) : +theorem type_of_call_decimal_is_sound {xs : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (h₁ : typeOf (Expr.call .decimal xs) c₁ env (l == .infinite) = Except.ok (ty, c₂)) : GuardedCapabilitiesInvariant (Expr.call .decimal xs) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.call .decimal xs) request entities v ∧ InstanceOfType v ty := by @@ -62,8 +62,8 @@ theorem type_of_call_decimal_is_sound {xs : List Expr} {c₁ c₂ : Capabilities · apply InstanceOfType.instance_of_ext simp [InstanceOfExtType] -theorem type_of_call_ip_inversion {xs : List Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.call .ip xs) c env = Except.ok (ty, c')) : +theorem type_of_call_ip_inversion {xs : List Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.call .ip xs) c env (l == .infinite) = Except.ok (ty, c')) : ty = .ext .ipAddr ∧ c' = ∅ ∧ ∃ (s : String), @@ -71,7 +71,7 @@ theorem type_of_call_ip_inversion {xs : List Expr} {c c' : Capabilities} {env : (Cedar.Spec.Ext.IPAddr.ip s).isSome := by simp [typeOf] at h₁ - cases h₂ : List.mapM₁ xs fun x => justType (typeOf x.val c env) <;> + cases h₂ : List.mapM₁ xs fun x => justType (typeOf x.val c env (l == .infinite)) <;> simp [h₂] at h₁ rename_i tys simp [typeOfCall, typeOfConstructor] at h₁ @@ -82,8 +82,8 @@ theorem type_of_call_ip_inversion {xs : List Expr} {c c' : Capabilities} {env : rename_i h₃ simp [h₃] -theorem type_of_call_ip_is_sound {xs : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} - (h₁ : typeOf (Expr.call .ip xs) c₁ env = Except.ok (ty, c₂)) : +theorem type_of_call_ip_is_sound {xs : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} + (h₁ : typeOf (Expr.call .ip xs) c₁ env (l == .infinite) = Except.ok (ty, c₂)) : GuardedCapabilitiesInvariant (Expr.call .ip xs) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.call .ip xs) request entities v ∧ InstanceOfType v ty := by @@ -98,12 +98,12 @@ theorem type_of_call_ip_is_sound {xs : List Expr} {c₁ c₂ : Capabilities} {en · apply InstanceOfType.instance_of_ext simp [InstanceOfExtType] -theorem typeOf_of_binary_call_inversion {xs : List Expr} {c : Capabilities} {env : Environment} {ty₁ ty₂ : CedarType} - (h₁ : (xs.mapM₁ fun x => justType (typeOf x.val c env)) = Except.ok [ty₁, ty₂]) : +theorem typeOf_of_binary_call_inversion {xs : List Expr} {c : Capabilities} {env : Environment} {ty₁ ty₂ : CedarType} {l : Level} + (h₁ : (xs.mapM₁ fun x => justType (typeOf x.val c env (l == .infinite))) = Except.ok [ty₁, ty₂]) : ∃ x₁ x₂ c₁ c₂, xs = [x₁, x₂] ∧ - typeOf x₁ c env = ok ty₁ c₁ ∧ - typeOf x₂ c env = ok ty₂ c₂ + typeOf x₁ c env (l == .infinite) = ok ty₁ c₁ ∧ + typeOf x₂ c env (l == .infinite) = ok ty₂ c₂ := by simp [List.mapM₁] at h₁ cases xs @@ -113,7 +113,7 @@ theorem typeOf_of_binary_call_inversion {xs : List Expr} {c : Capabilities} {env cases tl₁ case nil => simp [List.mapM, List.mapM.loop] at h₁ - cases h₂ : justType (typeOf hd₁ c env) <;> + cases h₂ : justType (typeOf hd₁ c env (l == .infinite)) <;> simp [h₂] at h₁ simp [pure, Except.pure] at h₁ case cons hd₂ tl₂ => @@ -133,7 +133,7 @@ theorem typeOf_of_binary_call_inversion {xs : List Expr} {c : Capabilities} {env simp [h₂, h₃] case cons hd₃ tl₃ => simp only [List.attach_def, List.pmap, List.mapM_cons, - List.mapM_pmap_subtype (fun x => justType (typeOf x c env)), bind_assoc, pure_bind] at h₁ + List.mapM_pmap_subtype (fun x => justType (typeOf x c env (l == .infinite))), bind_assoc, pure_bind] at h₁ rw [justType, Except.map] at h₁ split at h₁ <;> simp at h₁ rw [justType, Except.map] at h₁ @@ -142,7 +142,7 @@ theorem typeOf_of_binary_call_inversion {xs : List Expr} {c : Capabilities} {env split at h₁ <;> simp at h₁ rename_i res₁ _ _ res₂ _ _ res₃ _ simp only [pure, Except.pure] at h₁ - cases h₂ : List.mapM (fun x => justType (typeOf x c env)) tl₃ <;> simp [h₂] at h₁ + cases h₂ : List.mapM (fun x => justType (typeOf x c env (l == .infinite))) tl₃ <;> simp [h₂] at h₁ def IsDecimalComparator : ExtFun → Prop | .lessThan @@ -151,18 +151,18 @@ def IsDecimalComparator : ExtFun → Prop | .greaterThanOrEqual => True | _ => False -theorem type_of_call_decimal_comparator_inversion {xfn : ExtFun} {xs : List Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} +theorem type_of_call_decimal_comparator_inversion {xfn : ExtFun} {xs : List Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} (h₀ : IsDecimalComparator xfn) - (h₁ : typeOf (Expr.call xfn xs) c env = Except.ok (ty, c')) : + (h₁ : typeOf (Expr.call xfn xs) c env (l == .infinite) = Except.ok (ty, c')) : ty = .bool .anyBool ∧ c' = ∅ ∧ ∃ (x₁ x₂ : Expr) (c₁ c₂ : Capabilities), xs = [x₁, x₂] ∧ - typeOf x₁ c env = ok (.ext .decimal) c₁ ∧ - typeOf x₂ c env = ok (.ext .decimal) c₂ + typeOf x₁ c env (l == .infinite) = ok (.ext .decimal) c₁ ∧ + typeOf x₂ c env (l == .infinite) = ok (.ext .decimal) c₂ := by simp [typeOf] at h₁ - cases h₂ : List.mapM₁ xs fun x => justType (typeOf x.val c env) <;> + cases h₂ : List.mapM₁ xs fun x => justType (typeOf x.val c env (l == .infinite)) <;> simp [h₂] at h₁ rename_i tys simp [typeOfCall] at h₁ @@ -177,11 +177,11 @@ theorem type_of_call_decimal_comparator_inversion {xfn : ExtFun} {xs : List Expr } } -theorem type_of_call_decimal_comparator_is_sound {xfn : ExtFun} {xs : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_call_decimal_comparator_is_sound {xfn : ExtFun} {xs : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₀ : IsDecimalComparator xfn) (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.call xfn xs) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.call xfn xs) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih : ∀ (xᵢ : Expr), xᵢ ∈ xs → TypeOfIsSound xᵢ) : GuardedCapabilitiesInvariant (Expr.call xfn xs) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.call xfn xs) request entities v ∧ InstanceOfType v ty @@ -214,17 +214,17 @@ theorem type_of_call_decimal_comparator_is_sound {xfn : ExtFun} {xs : List Expr} } -theorem type_of_call_isInRange_inversion {xs : List Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.call .isInRange xs) c env = Except.ok (ty, c')) : +theorem type_of_call_isInRange_inversion {xs : List Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.call .isInRange xs) c env (l == .infinite) = Except.ok (ty, c')) : ty = .bool .anyBool ∧ c' = ∅ ∧ ∃ (x₁ x₂ : Expr) (c₁ c₂ : Capabilities), xs = [x₁, x₂] ∧ - typeOf x₁ c env = ok (.ext .ipAddr) c₁ ∧ - typeOf x₂ c env = ok (.ext .ipAddr) c₂ + typeOf x₁ c env (l == .infinite) = ok (.ext .ipAddr) c₁ ∧ + typeOf x₂ c env (l == .infinite) = ok (.ext .ipAddr) c₂ := by simp [typeOf] at h₁ - cases h₂ : List.mapM₁ xs fun x => justType (typeOf x.val c env) <;> + cases h₂ : List.mapM₁ xs fun x => justType (typeOf x.val c env (l == .infinite)) <;> simp [h₂] at h₁ rename_i tys simp [typeOfCall] at h₁ @@ -235,10 +235,10 @@ theorem type_of_call_isInRange_inversion {xs : List Expr} {c c' : Capabilities} apply typeOf_of_binary_call_inversion h₂ } -theorem type_of_call_isInRange_comparator_is_sound {xs : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_call_isInRange_comparator_is_sound {xs : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.call .isInRange xs) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.call .isInRange xs) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih : ∀ (xᵢ : Expr), xᵢ ∈ xs → TypeOfIsSound xᵢ) : GuardedCapabilitiesInvariant (Expr.call .isInRange xs) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.call .isInRange xs) request entities v ∧ InstanceOfType v ty @@ -274,11 +274,11 @@ def IsIpAddrRecognizer : ExtFun → Prop | _ => False -theorem typeOf_of_unary_call_inversion {xs : List Expr} {c : Capabilities} {env : Environment} {ty₁ : CedarType} - (h₁ : (xs.mapM₁ fun x => justType (typeOf x.val c env)) = Except.ok [ty₁]) : +theorem typeOf_of_unary_call_inversion {xs : List Expr} {c : Capabilities} {env : Environment} {ty₁ : CedarType} {l : Level} + (h₁ : (xs.mapM₁ fun x => justType (typeOf x.val c env (l == .infinite))) = Except.ok [ty₁]) : ∃ x₁ c₁, xs = [x₁] ∧ - typeOf x₁ c env = ok ty₁ c₁ + typeOf x₁ c env (l == .infinite) = ok ty₁ c₁ := by simp [List.mapM₁] at h₁ cases xs @@ -288,7 +288,7 @@ theorem typeOf_of_unary_call_inversion {xs : List Expr} {c : Capabilities} {env cases tl₁ case nil => simp [List.mapM, List.mapM.loop] at h₁ - cases h₂ : justType (typeOf hd₁ c env) <;> + cases h₂ : justType (typeOf hd₁ c env (l == .infinite)) <;> simp [h₂] at h₁ simp [justType, Except.map] at h₂ simp [pure, Except.pure] at h₁ @@ -299,26 +299,26 @@ theorem typeOf_of_unary_call_inversion {xs : List Expr} {c : Capabilities} {env simp [ok, h₃, ←h₂] case cons hd₂ tl₂ => simp only [List.attach_def, List.pmap, List.mapM_cons, - List.mapM_pmap_subtype (fun x => justType (typeOf x c env)), bind_assoc, pure_bind] at h₁ + List.mapM_pmap_subtype (fun x => justType (typeOf x c env (l == .infinite))), bind_assoc, pure_bind] at h₁ rw [justType, Except.map] at h₁ split at h₁ <;> simp at h₁ rw [justType, Except.map] at h₁ split at h₁ <;> simp at h₁ rename_i res₁ _ _ res₂ _ simp [pure, Except.pure] at h₁ - cases h₂ : List.mapM (fun x => justType (typeOf x c env)) tl₂ <;> simp [h₂] at h₁ + cases h₂ : List.mapM (fun x => justType (typeOf x c env (l == .infinite))) tl₂ <;> simp [h₂] at h₁ -theorem type_of_call_ipAddr_recognizer_inversion {xfn : ExtFun} {xs : List Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} +theorem type_of_call_ipAddr_recognizer_inversion {xfn : ExtFun} {xs : List Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} (h₀ : IsIpAddrRecognizer xfn) - (h₁ : typeOf (Expr.call xfn xs) c env = Except.ok (ty, c')) : + (h₁ : typeOf (Expr.call xfn xs) c env (l == .infinite) = Except.ok (ty, c')) : ty = .bool .anyBool ∧ c' = ∅ ∧ ∃ (x₁ : Expr) (c₁ : Capabilities), xs = [x₁] ∧ - typeOf x₁ c env = ok (.ext .ipAddr) c₁ + typeOf x₁ c env (l == .infinite) = ok (.ext .ipAddr) c₁ := by simp [typeOf] at h₁ - cases h₂ : List.mapM₁ xs fun x => justType (typeOf x.val c env) <;> + cases h₂ : List.mapM₁ xs fun x => justType (typeOf x.val c env (l == .infinite)) <;> simp [h₂] at h₁ rename_i tys simp [typeOfCall] at h₁ @@ -333,11 +333,11 @@ theorem type_of_call_ipAddr_recognizer_inversion {xfn : ExtFun} {xs : List Expr} } } -theorem type_of_call_ipAddr_recognizer_is_sound {xfn : ExtFun} {xs : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_call_ipAddr_recognizer_is_sound {xfn : ExtFun} {xs : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₀ : IsIpAddrRecognizer xfn) (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.call xfn xs) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.call xfn xs) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih : ∀ (xᵢ : Expr), xᵢ ∈ xs → TypeOfIsSound xᵢ) : GuardedCapabilitiesInvariant (Expr.call xfn xs) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.call xfn xs) request entities v ∧ InstanceOfType v ty @@ -363,10 +363,10 @@ theorem type_of_call_ipAddr_recognizer_is_sound {xfn : ExtFun} {xs : List Expr} apply bool_is_instance_of_anyBool } -theorem type_of_call_is_sound {xfn : ExtFun} {xs : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_call_is_sound {xfn : ExtFun} {xs : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.call xfn xs) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.call xfn xs) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih : ∀ (xᵢ : Expr), xᵢ ∈ xs → TypeOfIsSound xᵢ) : GuardedCapabilitiesInvariant (Expr.call xfn xs) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.call xfn xs) request entities v ∧ InstanceOfType v ty diff --git a/cedar-lean/Cedar/Thm/Validation/Typechecker/GetAttr.lean b/cedar-lean/Cedar/Thm/Validation/Typechecker/GetAttr.lean index 1413b74fa..54b0e585e 100644 --- a/cedar-lean/Cedar/Thm/Validation/Typechecker/GetAttr.lean +++ b/cedar-lean/Cedar/Thm/Validation/Typechecker/GetAttr.lean @@ -37,31 +37,76 @@ theorem getAttrInRecord_has_empty_capabilities {x₁ : Expr} {a : Attr} {c₁ c split at h₁ <;> simp at h₁ simp [h₁] -theorem type_of_getAttr_inversion {x₁ : Expr} {a : Attr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.getAttr x₁ a) c₁ env = Except.ok (ty, c₂)) : +theorem type_of_getAttr_inversion {x₁ : Expr} {a : Attr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {l₁ : Level} + (h₁ : typeOf (Expr.getAttr x₁ a) c₁ env (l₁ == .infinite) = Except.ok (ty, c₂)) : c₂ = ∅ ∧ ∃ c₁', - (∃ ety, typeOf x₁ c₁ env = Except.ok (.entity ety, c₁')) ∨ - (∃ rty, typeOf x₁ c₁ env = Except.ok (.record rty, c₁')) + (∃ ety l₂, typeOf x₁ c₁ env (l₁ == .infinite) = Except.ok (.entity ety l₂, c₁')) ∨ + (∃ rty, typeOf x₁ c₁ env (l₁ == .infinite) = Except.ok (.record rty, c₁')) := by simp [typeOf] at h₁ - cases h₂ : typeOf x₁ c₁ env <;> simp [h₂] at h₁ + cases h₂ : typeOf x₁ c₁ env (l₁ == .infinite) <;> simp [h₂] at h₁ case ok res => have ⟨ty₁, c₁'⟩ := res simp [typeOfGetAttr] at h₁ - split at h₁ <;> try contradiction - · simp only [List.empty_eq, Except.ok.injEq, Prod.mk.injEq, false_and, exists_const, - CedarType.record.injEq, exists_and_right, exists_eq', true_and, false_or, and_true] - apply getAttrInRecord_has_empty_capabilities h₁ - · simp only [List.empty_eq, Except.ok.injEq, Prod.mk.injEq, CedarType.entity.injEq, - exists_and_right, exists_eq', true_and, false_and, exists_const, or_false, and_true] + split at h₁ + case _ ty' rty' => -- Records + simp + apply getAttrInRecord_has_empty_capabilities + assumption + case _ ty' rty' l' => -- Entities split at h₁ <;> try simp [err] at h₁ - apply getAttrInRecord_has_empty_capabilities h₁ + rename_i hlevel + split at h₁ <;> try simp [err] at h₁ + rename_i otp rty'' _ + simp [bind, Except.bind] at h₁ + split at h₁ <;> try simp [err] at h₁ + simp + rename_i heq + have ⟨h₁₁, h₁₂⟩ := h₁ + subst h₁₂ + rename_i monad v + have ⟨result_ty, result_caps⟩ := v + simp + apply getAttrInRecord_has_empty_capabilities + apply heq + case _ => + simp [err] at h₁ -theorem type_of_getAttr_is_sound_for_records {x₁ : Expr} {a : Attr} {c₁ c₁' : Capabilities} {env : Environment} {rty : RecordType} {request : Request} {entities : Entities} {v₁ : Value} +theorem type_of_getAttr_inversion_levels {x₁ : Expr} {a : Attr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {l₁ : Level} + (h₁ : typeOf (Expr.getAttr x₁ a) c₁ env (l₁ == .infinite)= Except.ok (ty, c₂)) : + c₂ = ∅ ∧ + ∃ c₁', + (∃ ety l₂, typeOf x₁ c₁ env (l₁ == .infinite)= Except.ok (.entity ety l₂, c₁') ∧ l₂ > .finite 0) ∨ + (∃ rty, typeOf x₁ c₁ env (l₁ == .infinite) = Except.ok (.record rty, c₁')) + := by + have hinv := type_of_getAttr_inversion h₁ + replace ⟨hinv₁, c₁', hinv⟩ := hinv + constructor <;> try assumption + exists c₁' + cases hinv + case _ hinv => + apply Or.inl + replace ⟨ety, l₂, hinv⟩ := hinv + exists ety + exists l₂ + constructor <;> try assumption + simp [typeOf, hinv, typeOfGetAttr] at h₁ + split at h₁ + case _ hzero => + simp + assumption + case _ => + simp [err] at h₁ + case _ hinv => + replace ⟨rty, hinv⟩ := hinv + apply Or.inr + exists rty + +theorem type_of_getAttr_is_sound_for_records {x₁ : Expr} {a : Attr} {c₁ c₁' : Capabilities} {env : Environment} {rty : RecordType} {request : Request} {entities : Entities} {v₁ : Value} {l₁ : Level} (h₁ : CapabilitiesInvariant c₁ request entities) - (h₂ : typeOf (Expr.getAttr x₁ a) c₁ env = Except.ok (ty, ∅)) - (h₃ : typeOf x₁ c₁ env = Except.ok (CedarType.record rty, c₁')) + (h₂ : typeOf (Expr.getAttr x₁ a) c₁ env (l₁ == .infinite) = Except.ok (ty, ∅)) + (h₃ : typeOf x₁ c₁ env (l₁ == .infinite) = Except.ok (CedarType.record rty, c₁')) (h₄ : evaluate x₁ request entities = Except.ok v₁) (h₅ : InstanceOfType v₁ (CedarType.record rty)) : ∃ v, @@ -100,13 +145,183 @@ theorem type_of_getAttr_is_sound_for_records {x₁ : Expr} {a : Attr} {c₁ c₁ subst h₂ apply instance_of_attribute_type h₅ h₉ (by simp [Qualified.getType]) h₈ -theorem type_of_getAttr_is_sound_for_entities {x₁ : Expr} {a : Attr} {c₁ c₁' : Capabilities} {env : Environment} {ety : EntityType} {request : Request} {entities : Entities} {v₁ : Value} +theorem setLevel_preserves_type {v : Value} {ty : CedarType} {l : Level} + (h : InstanceOfType v ty) : + InstanceOfType v (setLevel l.sub1 ty) := by + cases ty <;> cases v <;> cases h + case _ => + simp [setLevel] + apply InstanceOfType.instance_of_bool + assumption + case _ => + simp [setLevel] + apply InstanceOfType.instance_of_int + case _ => + simp [setLevel] + apply InstanceOfType.instance_of_string + case _ => + simp [setLevel] + apply InstanceOfType.instance_of_entity + assumption + case _ => + simp [setLevel] + rename_i ty s h₁ + apply InstanceOfType.instance_of_set + intros v h₂ + apply setLevel_preserves_type + apply h₁ + apply h₂ + case _ => + rename_i rty m h₁ h₂ h₃ + simp [setLevel] + apply InstanceOfType.instance_of_record + case _ => + intros attr h₄ + apply Map.mapOnValuesAttach_preservesKeys_adapter + apply h₁ + apply h₄ + exists (Functor.map (setLevel l.sub1)) + simp [Functor.map] + funext + split <;> rename_i heq <;> rw [heq] + case _ => + intros k v qty h₄ h₅ + have h₆ : rty.contains k = true := by + apply h₁ + rw [Map.contains_iff_some_find?] + exists v + have h₇ : ∃ qty_orig, rty.find? k = .some qty_orig := by + rw [← Map.contains_iff_some_find?] + assumption + replace ⟨qty_orig, h₇⟩ := h₇ + have h₈ : some qty = some (setLevel l.sub1 <$> qty_orig) + := by + rw [← h₅] + rw [@Map.mapOnValues_maps_adapter Attr QualifiedType _ _ _ _ _ _ rty _ (Functor.map (setLevel l.sub1)) k qty_orig] + apply h₇ + simp [Functor.map] + funext + split <;> rename_i heq <;> rw [heq] + simp at h₈ + have htype : InstanceOfType v (Qualified.getType qty_orig) := by + apply h₂ + repeat assumption + rw [h₈] + simp [Functor.map] + split + <;> simp [Qualified.getType] + <;> apply setLevel_preserves_type + <;> simp [Qualified.getType] at htype + <;> assumption + case _ => + intros k qty h₄ h₅ + have h₆ : (rty.mapOnValuesAttach (λ prod => setLevel l.sub1 <$> prod.val)).contains k = true + := by + rw [Map.contains_iff_some_find?] + exists qty + simp [Functor.map] + rw [← h₄] + rw [Map.mapOnValuesAttachFunEq] + funext + split <;> rename_i heq <;> rw [heq] + have h₇ : rty.contains k = true := by + rw [Map.mapOnValuesAttach_preservesContains_adapter] + apply h₆ + exists (Functor.map (setLevel l.sub1)) + funext + rw [Map.contains_iff_some_find?] at h₇ + replace ⟨qty_orig, h₇⟩ := h₇ + apply h₃ + assumption + have heq : qty = setLevel l.sub1 <$> qty_orig := by + rw [@Map.mapOnValues_maps_adapter _ _ _ _ _ _ _ _ _ _ (Functor.map (setLevel l.sub1)) k qty_orig] at h₄ + simp at h₄ + rw [← h₄] + assumption + funext + simp [Functor.map] + split <;> rename_i heq <;> rw [heq] + cases qty_orig + case _ => + simp [heq, Functor.map, Qualified.isRequired] at h₅ + case _ => + simp [Qualified.isRequired] + case _ => + simp [setLevel] + apply InstanceOfType.instance_of_ext + assumption +termination_by sizeOf ty +decreasing_by + all_goals simp_wf + all_goals try omega + case _ => + rename_i c _ _ _ _ _ _ _ _ _ _ _ _ _ + rw [c] + simp [sizeOf, CedarType._sizeOf_1] + omega + case _ => + rename_i ty' h_orig + have hsize₁ : sizeOf ty' < sizeOf qty_orig := by + rw [h_orig] + simp + omega + rename_i rty heq_ty _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ + rw [heq_ty] + simp + rw [← h_orig] at h₇ + + have hsize₂ : sizeOf qty_orig < sizeOf rty := by + have h : (k, qty_orig) ∈ rty.kvs := by + apply Map.find_means_mem + assumption + + have h₂ : sizeOf (k, qty_orig) < sizeOf rty.kvs := by + apply List.sizeOf_lt_of_mem + assumption + + have h₃ : sizeOf rty.kvs < sizeOf rty := by + apply Map.sizeOf_lt_of_kvs + + have h₄ : sizeOf qty_orig < sizeOf (k, qty_orig) := by + simp + omega + + omega + + + omega + case _ => + rename_i h₁₀ h₁₁ hinst₁ type hinst₂ + rename_i heq₁ _ _ _ _ rty _ _ _ _ _ _ _ _ _ heq₂ _ _ _ _ _ + rw [heq₂] at h₁₀ + rw [heq₁] + rw [heq₂] + rw [heq₂] at h₇ + have hsize₁ : sizeOf type < sizeOf (Qualified.required type) := by + simp + omega + have hsize₂ : sizeOf (Qualified.required type) < sizeOf rty := by + have hmem : (k, (Qualified.required type)) ∈ rty.kvs := by + apply Map.find_means_mem + assumption + have h₁ : sizeOf (Qualified.required type) < sizeOf (k, Qualified.required type) := by + simp + omega + have h₂ : sizeOf (k, Qualified.required type) < sizeOf rty.kvs := by + apply List.sizeOf_lt_of_mem + assumption + have h₃ : sizeOf rty.kvs < sizeOf rty := by + apply Map.sizeOf_lt_of_kvs + omega + simp + omega +theorem type_of_getAttr_is_sound_for_entities {x₁ : Expr} {a : Attr} {c₁ c₁' : Capabilities} {env : Environment} {ety : EntityType} {request : Request} {entities : Entities} {v₁ : Value} {l₁ l₂ : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.getAttr x₁ a) c₁ env = Except.ok (ty, ∅)) - (h₄ : typeOf x₁ c₁ env = Except.ok (CedarType.entity ety, c₁')) + (h₃ : typeOf (Expr.getAttr x₁ a) c₁ env (l₁ == .infinite) = Except.ok (ty, ∅)) + (h₄ : typeOf x₁ c₁ env (l₁ == .infinite) = Except.ok (CedarType.entity ety l₂, c₁')) (h₅ : evaluate x₁ request entities = Except.ok v₁) - (h₆ : InstanceOfType v₁ (CedarType.entity ety)) : + (h₆ : InstanceOfType v₁ (CedarType.entity ety l₂)) : ∃ v, (getAttr v₁ a entities = Except.error Error.entityDoesNotExist ∨ getAttr v₁ a entities = Except.error Error.extensionError ∨ @@ -130,6 +345,7 @@ theorem type_of_getAttr_is_sound_for_entities {x₁ : Expr} {a : Attr} {c₁ c simp only [typeOf, h₄, typeOfGetAttr, getAttrInRecord, List.empty_eq, Except.bind_ok] at h₃ split at h₃ <;> simp [ok, err] at h₃ split at h₃ <;> try simp at h₃ + split at h₃ <;> try simp at h₃ case h_1.h_1 _ _ h₁₀ _ _ h₁₁ => subst h₃ have h₁₂ := well_typed_entity_attributes h₂ h₈ h₁₀ @@ -145,20 +361,23 @@ theorem type_of_getAttr_is_sound_for_entities {x₁ : Expr} {a : Attr} {c₁ c simp [typeOf, h₄, typeOfGetAttr, getAttrInRecord] at h₃ split at h₃ <;> simp [ok, err] at h₃ split at h₃ <;> try simp at h₃ + split at h₃ <;> try simp at h₃ case h_1.h_1 _ _ h₁₀ _ _ h₁₁ => subst h₃ + apply setLevel_preserves_type apply instance_of_attribute_type _ h₁₁ (by simp [Qualified.getType]) h₉ apply well_typed_entity_attributes h₂ h₈ h₁₀ case h_1.h_2 _ _ h₁₀ _ _ h₁₁ => split at h₃ <;> simp at h₃ subst h₃ + apply setLevel_preserves_type apply instance_of_attribute_type _ h₁₁ (by simp [Qualified.getType]) h₉ apply well_typed_entity_attributes h₂ h₈ h₁₀ -theorem type_of_getAttr_is_sound {x₁ : Expr} {a : Attr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_getAttr_is_sound {x₁ : Expr} {a : Attr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l₁ : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.getAttr x₁ a) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.getAttr x₁ a) c₁ env (l₁ == .infinite) = Except.ok (ty, c₂)) (ih : TypeOfIsSound x₁) : GuardedCapabilitiesInvariant (Expr.getAttr x₁ a) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.getAttr x₁ a) request entities v ∧ InstanceOfType v ty @@ -166,7 +385,7 @@ theorem type_of_getAttr_is_sound {x₁ : Expr} {a : Attr} {c₁ c₂ : Capabilit have ⟨h₅, c₁', h₄⟩ := type_of_getAttr_inversion h₃ subst h₅ apply And.intro empty_guarded_capabilities_invariant - rcases h₄ with ⟨ety, h₄⟩ | ⟨rty, h₄⟩ <;> + rcases h₄ with ⟨ety, l₂, h₄⟩ | ⟨rty, h₄⟩ <;> have ⟨_, v₁, h₆, h₇⟩ := ih h₁ h₂ h₄ <;> simp [EvaluatesTo] at h₆ <;> simp [EvaluatesTo, evaluate] <;> diff --git a/cedar-lean/Cedar/Thm/Validation/Typechecker/HasAttr.lean b/cedar-lean/Cedar/Thm/Validation/Typechecker/HasAttr.lean index 7fbcdf2aa..12236699e 100644 --- a/cedar-lean/Cedar/Thm/Validation/Typechecker/HasAttr.lean +++ b/cedar-lean/Cedar/Thm/Validation/Typechecker/HasAttr.lean @@ -26,15 +26,15 @@ open Cedar.Data open Cedar.Spec open Cedar.Validation -theorem type_of_hasAttr_inversion {x₁ : Expr} {a : Attr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.hasAttr x₁ a) c₁ env = Except.ok (ty, c₂)) : +theorem type_of_hasAttr_inversion {x₁ : Expr} {a : Attr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.hasAttr x₁ a) c₁ env (l == .infinite) = Except.ok (ty, c₂)) : (c₂ = ∅ ∨ c₂ = Capabilities.singleton x₁ a) ∧ ∃ c₁', - (∃ ety, typeOf x₁ c₁ env = Except.ok (.entity ety, c₁')) ∨ - (∃ rty, typeOf x₁ c₁ env = Except.ok (.record rty, c₁')) + (∃ ety l', typeOf x₁ c₁ env (l == .infinite) = Except.ok (.entity ety l', c₁')) ∨ + (∃ rty, typeOf x₁ c₁ env (l == .infinite)= Except.ok (.record rty, c₁')) := by simp [typeOf, typeOfHasAttr] at h₁ - cases h₂ : typeOf x₁ c₁ env <;> simp [h₂] at h₁ + cases h₂ : typeOf x₁ c₁ env (l == .infinite) <;> simp [h₂] at h₁ case ok res => have ⟨ty₁, c₁'⟩ := res simp at h₁ @@ -43,15 +43,141 @@ theorem type_of_hasAttr_inversion {x₁ : Expr} {a : Attr} {c₁ c₂ : Capabili <;> split at h₁ <;> try split at h₁ <;> try split at h₁ - all_goals { + <;> try (simp [ok] at h₁ ; try simp [h₁]) + case _ => simp [ok] at h₁ - try simp [h₁] - } + constructor + simp [h₁] + exists c₁' + apply Or.inr + rename_i rty _ _ _ _ + exists rty + case _ => + simp [ok] at h₁ + constructor + simp [h₁] + exists c₁' + apply Or.inr + rename_i rty _ _ _ _ + exists rty + case _ => + simp [ok] at h₁ + constructor + simp [h₁] + exists c₁' + rename_i rty _ _ + apply Or.inr + exists rty + case _ => + split at h₁ + <;> simp at h₁ + <;> constructor + <;> simp [h₁] + <;> exists c₁' + <;> rename_i ety l _ _ _ _ _ _ _ + <;> apply Or.inl + <;> exists ety + <;> exists l + case _ => + exfalso + assumption + +-- Above theorem doesn't constrain `ty` to be a bool +theorem type_of_hasAttr_inversion' {x₁ : Expr} {a : Attr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.hasAttr x₁ a) c₁ env (l == .infinite) = Except.ok (ty, c₂)) : + ∃ bty, ty = .bool bty + := by + simp [typeOf] at h₁ + cases h_expr_type : typeOf x₁ c₁ env (l == .infinite) + case _ => + simp [h_expr_type] at h₁ + case _ => + simp [h_expr_type, typeOfHasAttr] at h₁ + split at h₁ + case _ _ => + simp [hasAttrInRecord] at h₁ + split at h₁ <;> simp [ok] at h₁ + case _ _ => + split at h₁ + case _ => + simp [ok] at h₁ + exists .tt + simp [h₁] + case _ => + simp [ok] at h₁ + exists .anyBool + simp [h₁] + case _ _ => + exists .ff + simp [h₁] + case _ _ => + split at h₁ + case _ h₁ => + split at h₁ + case _ _ => + simp [hasAttrInRecord] at h₁ + split at h₁ + case _ _ => + split at h₁ <;> simp [ok] at h₁ + case _ => + exists .tt + simp [h₁] + case _ => + exists .anyBool + simp [h₁] + case _ => + simp [ok] at h₁ + exists .ff + simp [h₁] + case _ _ => + split at h₁ + case _ => + simp [ok] at h₁ + exists .ff + simp [h₁] + case _ => + simp [ok, err] at h₁ + case _ h₁ => + simp [ok, err] at h₁ + case _ h => + simp [ok, err] at h₁ + + + +theorem type_of_hasAttr_inversion_finite {x₁ : Expr} {a : Attr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.hasAttr x₁ a) c₁ env (l == .infinite) = Except.ok (ty, c₂)) : + (c₂ = ∅ ∨ c₂ = Capabilities.singleton x₁ a) ∧ + ∃ c₁', + (∃ ety l', .finite 0 < l' ∧ typeOf x₁ c₁ env (l == .infinite) = Except.ok (.entity ety l', c₁')) ∨ + (∃ rty, typeOf x₁ c₁ env (l == .infinite)= Except.ok (.record rty, c₁')) + := by + have hinv := type_of_hasAttr_inversion h₁ + replace ⟨hinv₁, c₁', hinv⟩ := hinv + constructor <;> try assumption + exists c₁' + cases hinv + case _ hinv => + apply Or.inl + replace ⟨ety, l', hinv⟩ := hinv + exists ety + exists l' + constructor <;> try assumption + simp [typeOf, hinv, typeOfHasAttr] at h₁ + split at h₁ + case _ => + assumption + case _ => + simp [err] at h₁ + case _ hinv => + apply Or.inr + assumption + + -theorem type_of_hasAttr_is_sound_for_records {x₁ : Expr} {a : Attr} {c₁ c₁' : Capabilities} {env : Environment} {rty : RecordType} {request : Request} {entities : Entities} {v₁ : Value} +theorem type_of_hasAttr_is_sound_for_records {x₁ : Expr} {a : Attr} {c₁ c₁' : Capabilities} {env : Environment} {rty : RecordType} {request : Request} {entities : Entities} {v₁ : Value} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) - (h₂ : typeOf (Expr.hasAttr x₁ a) c₁ env = Except.ok (ty, c₂)) - (h₃ : typeOf x₁ c₁ env = Except.ok (CedarType.record rty, c₁')) + (h₂ : typeOf (Expr.hasAttr x₁ a) c₁ env (l == .infinite) = Except.ok (ty, c₂)) + (h₃ : typeOf x₁ c₁ env (l == .infinite) = Except.ok (CedarType.record rty, c₁')) (h₄ : evaluate x₁ request entities = Except.ok v₁) (h₅ : InstanceOfType v₁ (CedarType.record rty)) : ∃ v, @@ -95,13 +221,13 @@ theorem type_of_hasAttr_is_sound_for_records {x₁ : Expr} {a : Attr} {c₁ c₁ simp [Map.contains_iff_some_find?, h₇] at h₆ -theorem type_of_hasAttr_is_sound_for_entities {x₁ : Expr} {a : Attr} {c₁ c₁' : Capabilities} {env : Environment} {ety : EntityType} {request : Request} {entities : Entities} {v₁ : Value} +theorem type_of_hasAttr_is_sound_for_entities {x₁ : Expr} {a : Attr} {c₁ c₁' : Capabilities} {env : Environment} {ety : EntityType} {request : Request} {entities : Entities} {v₁ : Value} {l l₁ : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.hasAttr x₁ a) c₁ env = Except.ok (ty, c₂)) - (h₄ : typeOf x₁ c₁ env = Except.ok (CedarType.entity ety, c₁')) + (h₃ : typeOf (Expr.hasAttr x₁ a) c₁ env (l == .infinite) = Except.ok (ty, c₂)) + (h₄ : typeOf x₁ c₁ env (l == .infinite) = Except.ok (CedarType.entity ety l₁, c₁')) (h₅ : evaluate x₁ request entities = Except.ok v₁) - (h₆ : InstanceOfType v₁ (CedarType.entity ety)) : + (h₆ : InstanceOfType v₁ (CedarType.entity ety l₁)) : ∃ v, (hasAttr v₁ a entities = Except.error Error.entityDoesNotExist ∨ hasAttr v₁ a entities = Except.error Error.extensionError ∨ @@ -113,7 +239,7 @@ theorem type_of_hasAttr_is_sound_for_entities {x₁ : Expr} {a : Attr} {c₁ c subst h₆ h₇ simp [hasAttr, attrsOf] simp [typeOf, h₄, typeOfHasAttr] at h₃ - split at h₃ <;> try simp [err, hasAttrInRecord] at h₃ + split at h₃ <;> try simp [err, hasAttrInRecord] at h₃ <;> split at h₃ rename_i _ rty h₇ split at h₃ case h_1.h_1 => @@ -159,10 +285,10 @@ theorem type_of_hasAttr_is_sound_for_entities {x₁ : Expr} {a : Attr} {c₁ c contradiction -theorem type_of_hasAttr_is_sound {x₁ : Expr} {a : Attr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_hasAttr_is_sound {x₁ : Expr} {a : Attr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.hasAttr x₁ a) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.hasAttr x₁ a) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih : TypeOfIsSound x₁) : GuardedCapabilitiesInvariant (Expr.hasAttr x₁ a) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.hasAttr x₁ a) request entities v ∧ InstanceOfType v ty @@ -177,7 +303,7 @@ theorem type_of_hasAttr_is_sound {x₁ : Expr} {a : Attr} {c₁ c₂ : Capabilit subst h₇; subst h₈ simp [EvaluatesTo, h₆] case right => - rcases h₄ with ⟨ety, h₄⟩ | ⟨rty, h₄⟩ <;> + rcases h₄ with ⟨ety, l, h₄⟩ | ⟨rty, h₄⟩ <;> have ⟨_, v₁, h₆, h₇⟩ := ih h₁ h₂ h₄ <;> simp [EvaluatesTo] at h₆ <;> simp [EvaluatesTo, evaluate] <;> diff --git a/cedar-lean/Cedar/Thm/Validation/Typechecker/IfThenElse.lean b/cedar-lean/Cedar/Thm/Validation/Typechecker/IfThenElse.lean index 908ac0476..0392d6ce0 100644 --- a/cedar-lean/Cedar/Thm/Validation/Typechecker/IfThenElse.lean +++ b/cedar-lean/Cedar/Thm/Validation/Typechecker/IfThenElse.lean @@ -25,29 +25,29 @@ namespace Cedar.Thm open Cedar.Spec open Cedar.Validation -theorem type_of_ite_inversion {x₁ x₂ x₃ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.ite x₁ x₂ x₃) c env = Except.ok (ty, c')) : +theorem type_of_ite_inversion {x₁ x₂ x₃ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.ite x₁ x₂ x₃) c env (l == .infinite) = Except.ok (ty, c')) : ∃ bty₁ c₁ ty₂ c₂ ty₃ c₃, - typeOf x₁ c env = .ok (.bool bty₁, c₁) ∧ + typeOf x₁ c env (l == .infinite) = .ok (.bool bty₁, c₁) ∧ match bty₁ with | .ff => - typeOf x₃ c env = .ok (ty₃, c₃) ∧ ty = ty₃ ∧ c' = c₃ + typeOf x₃ c env (l == .infinite) = .ok (ty₃, c₃) ∧ ty = ty₃ ∧ c' = c₃ | .tt => - typeOf x₂ (c ∪ c₁) env = .ok (ty₂, c₂) ∧ + typeOf x₂ (c ∪ c₁) env (l == .infinite) = .ok (ty₂, c₂) ∧ ty = ty₂ ∧ c' = c₁ ∪ c₂ | .anyBool => - typeOf x₂ (c ∪ c₁) env = .ok (ty₂, c₂) ∧ - typeOf x₃ c env = .ok (ty₃, c₃) ∧ + typeOf x₂ (c ∪ c₁) env (l == .infinite) = .ok (ty₂, c₂) ∧ + typeOf x₃ c env (l == .infinite) = .ok (ty₃, c₃) ∧ (ty₂ ⊔ ty₃) = (.some ty) ∧ c' = (c₁ ∪ c₂) ∩ c₃ := by simp [typeOf] at h₁ - cases h₂ : typeOf x₁ c env <;> simp [h₂, typeOfIf] at * + cases h₂ : typeOf x₁ c env (l == .infinite) <;> simp [h₂, typeOfIf] at * rename_i res₁ split at h₁ <;> try { simp [ok, err] at h₁ } <;> rename_i c₁ hr₁ <;> simp at hr₁ <;> have ⟨ht₁, hc₁⟩ := hr₁ case ok.h_1 => exists BoolType.tt, res₁.snd ; simp [←ht₁] - cases h₃ : typeOf x₂ (c ∪ res₁.snd) env <;> simp [h₃] at h₁ + cases h₃ : typeOf x₂ (c ∪ res₁.snd) env (l == .infinite) <;> simp [h₃] at h₁ rename_i res₂ ; simp [ok] at h₁ have ⟨ht₂, hc₂⟩ := h₁ exists res₂.fst, res₂.snd @@ -58,8 +58,8 @@ theorem type_of_ite_inversion {x₁ x₂ x₃ : Expr} {c c' : Capabilities} {env exists ty case ok.h_3 => exists BoolType.anyBool, res₁.snd ; simp [←ht₁] - cases h₃ : typeOf x₂ (c ∪ res₁.snd) env <;> simp [h₃] at h₁ - cases h₄ : typeOf x₃ c env <;> simp [h₄] at h₁ + cases h₃ : typeOf x₂ (c ∪ res₁.snd) env (l == .infinite) <;> simp [h₃] at h₁ + cases h₄ : typeOf x₃ c env (l == .infinite)<;> simp [h₄] at h₁ split at h₁ <;> simp [ok, err] at h₁ rename_i ty' res₂ res₃ _ ty' hty have ⟨ht, hc⟩ := h₁ @@ -68,10 +68,10 @@ theorem type_of_ite_inversion {x₁ x₂ x₃ : Expr} {c c' : Capabilities} {env simp only [Except.ok.injEq, true_and] exists res₃.fst, res₃.snd -theorem type_of_ite_is_sound {x₁ x₂ x₃ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_ite_is_sound {x₁ x₂ x₃ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.ite x₁ x₂ x₃) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.ite x₁ x₂ x₃) c₁ env (l == .infinite)= Except.ok (ty, c₂)) (ih₁ : TypeOfIsSound x₁) (ih₂ : TypeOfIsSound x₂) (ih₃ : TypeOfIsSound x₃) : diff --git a/cedar-lean/Cedar/Thm/Validation/Typechecker/LUB.lean b/cedar-lean/Cedar/Thm/Validation/Typechecker/LUB.lean index eb9f357a5..a0bd6b3a6 100644 --- a/cedar-lean/Cedar/Thm/Validation/Typechecker/LUB.lean +++ b/cedar-lean/Cedar/Thm/Validation/Typechecker/LUB.lean @@ -175,6 +175,82 @@ theorem lubRecord_comm {rty₁ rty₂ : List (Attr × Qualified CedarType)} : by_contra apply h₂ a₁ hd₂ tl₂ a₁ hd₁ tl₁ <;> rfl +theorem Level.min_comm {l₁ l₂ : Level} : + min l₁ l₂ = min l₂ l₁ + := by + cases l₁ <;> cases l₂ <;> simp + case _ n₁ n₂ => + apply Nat.min_comm + + + + + +theorem Level.min_one_of (l₁ l₂ l : Level) : + min l₁ l₂ = l → + (l = l₁) ∨ (l = l₂) + := by + intros h + cases l₁ <;> cases l₂ + any_goals try { simp at h ; rw [h] ; simp } + rename_i n₁ n₂ + cases l <;> simp at h + rename_i n + simp [min,instMinNat,minOfLe] at h + split at h <;> rw[h] <;> simp + +theorem Level.min_same (l : Level) : min l l = l + := by + cases l <;> simp + + + +theorem Level.min_left (l₁ l₂ l' l'' : Level) : + min l₁ l₂ = l' → + min l₁ l' = l'' → + l' = l'' + := by + intros h₁ h₂ + have h₃ := Level.min_one_of l₁ l₂ l' h₁ + cases h₃ + case inl h' => + rw [h'] at h₂ + rw [Level.min_same] at h₂ + simp [*] + case inr h' => + rw [h'] at h₂ + rw [h₂] at h₁ + simp [*] + +theorem Level.lub_none_assoc (ety : EntityType) (ty : CedarType) (l₁ l₂ : Level) : + (ty ⊔ .entity ety l₁) = .none → + (ty ⊔ .entity ety l₂) = .none + := by + intros h + cases ty <;> simp [lub?] at h <;> simp [lub?] + apply h + +theorem Level.min_lemma (l₁ l₂ l₃ l₄ l₅ : Level) : + min l₁ l₂ = l₄ → + min l₂ l₃ = l₅ → + (min l₄ l₃) = (min l₁ l₅) + := by + intros h₁ h₂ + cases l₁ <;> cases l₂ <;> cases l₃ <;> cases l₄ <;> cases l₅ <;> simp <;> simp at h₁ <;> simp at h₂ + all_goals { omega } + + + + + + + + +theorem EntityType.eq_comm {t₁ t₂ : EntityType} : t₁ = t₂ → t₂ = t₁ := by + intros h + rw [h] + + theorem lub_comm {ty₁ ty₂ : CedarType} : (ty₁ ⊔ ty₂) = (ty₂ ⊔ ty₁) := by @@ -186,22 +262,41 @@ theorem lub_comm {ty₁ ty₂ : CedarType} : have h := @lub_comm s₁ s₂ simp [h] case h_3 => + rename_i ety₁ l₁ ety₂ l₂ + simp + rw [Level.min_comm] + cases h : (decide (ety₁ = ety₂)) + case false => + simp at h + rw [if_neg] + rw [if_neg] + intros h' + rw [h'] at h + apply h + rfl + apply h + case true => + simp at h + rw [h] + case h_4 => rename_i rty₁ rty₂ have h := @lubRecord_comm rty₁ rty₂ simp [h] - case h_4 => - rename_i h₁ h₂ h₃ - split <;> split <;> rename_i h₄ - case isTrue.h_4 | isFalse.h_4 => - rename_i _ _ h₅ _ _ _ _ - rw [eq_comm] at h₅ - simp [h₅] + case h_5 => + rename_i h₁ h₂ h₃ h₄ + split <;> split + case isTrue.h_5 | isFalse.h_5 => + rename_i _ _ h _ _ _ _ _ _ + rw [eq_comm] at h + simp [h] + all_goals { - rename_i v₁ v₂ + rename_i h₄ v₁ v₂ _ by_contra try { apply h₁ v₂ v₁ <;> rfl } try { apply h₂ v₂ v₁ <;> rfl } - try { apply h₃ v₂ v₁ <;> rfl } + try { apply h₃ v₁ v₂ <;> rfl } + try { apply h₄ v₂ v₁ <;> rfl } } end @@ -215,7 +310,9 @@ theorem lub_refl (ty : CedarType) : case h_2 eltTy => have h₁ := lub_refl eltTy simp [h₁] - case h_3 rty => + case h_3 ety level => + cases level <;> simp [min, Option.min] + case h_4 rty => have h₁ := lubRecordType_refl rty simp [h₁] @@ -267,6 +364,15 @@ theorem lubQualified_is_lub_of_getType {qty qty₁ qty₂: Qualified CedarType} simp only [Qualified.getType, ← h₁] } +theorem Level.min_trans {l₁ l₂ l₃ : Level} : + min l₁ l₂ = l₂ → + min l₂ l₃ = l₃ → + min l₁ l₃ = l₃ + := by + intros h₁ h₂ + cases l₁ <;> cases l₂ <;> cases l₃ <;> simp <;> simp at h₁ <;> simp at h₂ + omega + mutual theorem lub_trans {ty₁ ty₂ ty₃ : CedarType} : @@ -294,7 +400,26 @@ theorem lub_trans {ty₁ ty₂ ty₃ : CedarType} : rw [eq_comm] at h₁ h₂ ; subst h₁ h₂ have h₅ := lub_trans h₃ h₄ simp [h₅] - case h_3 rty₁ rty₃ => + case h_3 ety₁ l₁ ety₂ l₂ => + unfold lub? at h₁ h₂ + cases ty₂ <;> simp at h₁ h₂ + rename_i ty' l' + have ⟨heq_ety₁, heq_l₁⟩ := h₁ + have ⟨heq_ety₂, heq_l₂⟩ := h₂ + cases heq : (decide (ety₁ = ety₂)) <;> simp at heq + case false => + simp + exfalso + apply heq + rw [← heq_ety₁] at heq_ety₂ + apply heq_ety₂ + case true => + rw [heq] + simp + apply Level.min_trans + apply heq_l₁ + apply heq_l₂ + case h_4 rty₁ rty₃ => unfold lub? at h₁ h₂ cases ty₂ <;> simp at h₁ h₂ rename_i mty₂ ; cases mty₂ ; rename_i rty₂ @@ -304,10 +429,11 @@ theorem lub_trans {ty₁ ty₂ ty₃ : CedarType} : rw [eq_comm] at h₁ h₂ ; subst h₁ h₂ have h₅ := lubRecordType_trans h₃ h₄ simp [h₅] - case h_4 => + case h_5 => + split case isTrue h₃ => simp [h₃] - case isFalse h₃ h₄ h₅ h₆ => + case isFalse h₃ h₄ h₅ h₆ h₇ => unfold lub? at h₁ h₂ cases ty₁ <;> cases ty₂ <;> simp at h₁ <;> cases ty₃ <;> simp at h₂ <;> simp at h₆ @@ -318,10 +444,18 @@ theorem lub_trans {ty₁ ty₂ ty₃ : CedarType} : case record rty₁ _ rty₃ => cases rty₁ ; cases rty₃ rename_i rty₁' rty₃' - apply h₅ rty₁' rty₃' <;> rfl + apply h₆ rty₁' rty₃' <;> rfl + case entity => + rename_i cty₁ cty₂ ety₁ l₁ ety₂ l₂ ety' l' + apply h₅ ety₁ l₁ <;> rfl + case ext => + apply h₇ + rw [← h₂] + rw [h₁] all_goals { - rename_i ety₁ ety₂ ety₃ - rw [h₁] at h₆ ; contradiction + -- rename_i ety₁ ety₂ ety₃ + apply h₇ + rfl } theorem lubRecordType_trans {rty₁ rty₂ rty₃ : List (Attr × QualifiedType)} : @@ -412,7 +546,45 @@ theorem lub_left_subty {ty₁ ty₂ ty₃ : CedarType} : have h₄ := lub_left_subty h₂ <;> simp [subty, h₃] at h₄ assumption - case h_3 rty₁ rty₂ => + case h_3 ety₁ l₁ ety₂ l₂ => + cases heq₁ : decide (ety₁ = ety₂) <;> split <;> simp at heq₁ + any_goals try { + rw [if_neg] at h₁ + contradiction + simp + apply heq₁ + } + case true.h_1 o t heq₂ => + rw [if_pos] at h₁ + injection h₁ + rename_i heq₃ + cases ty₃ <;> simp at heq₃ + rename_i ety' l' + unfold lub? at heq₂ + cases t <;> simp at heq₂ + rename_i ety'' l'' + have ⟨_, heq₄⟩ := heq₃ + have ⟨heq₅, heq₆ ,heq₇⟩ := heq₂ + simp + rw [← heq₆] + rw [heq₅] + apply And.intro + rfl + rw [eq_comm] + apply Level.min_left + apply heq₄ + apply heq₇ + rw [heq₁] + simp + case true.h_2 o t heq₂ => + exfalso + rw [heq₁] at h₁ + simp at h₁ + cases ty₃ <;> simp at h₁ + rename_i ety' l' + unfold lub? at heq₂ + simp [*] at heq₂ + case h_4 rty₁ rty₂ => cases h₂ : lubRecordType rty₁ rty₂ <;> simp [h₂] at h₁ rename_i rty₃ subst h₁ @@ -422,7 +594,7 @@ theorem lub_left_subty {ty₁ ty₂ ty₃ : CedarType} : have h₄ := lubRecordType_left_subty h₂ <;> simp [h₃] at h₄ assumption - case h_4 => + case h_5 => split at h₁ <;> try contradiction rename_i h₂ subst h₂ @@ -492,9 +664,9 @@ theorem lubBool_assoc_none_some {ty₁ ty₂ : CedarType} {bty₁ bty₂ : BoolT simp at h₂ unfold lub? at h₁ split at h₁ <;> try contradiction - rename_i ty₁' ty₂' ty₃' h₃ h₄ h₅ + rename_i ty₁' ty₂' ty₃' h₃ h₄ h₅ h₆ subst h₂ - cases ty₁' <;> simp [lub?] + cases ty₁' <;> try simp [lub?] rename_i bty₃ split at h₁ <;> try contradiction apply h₃ bty₃ bty₁ <;> rfl @@ -520,7 +692,29 @@ theorem lub_assoc_none_some {ty₁ ty₂ ty₃ ty₄ : CedarType} cases h₄ : ty₁' ⊔ sty₂ <;> simp [h₄] at h₁ have h₅ := lub_assoc_none_some h₄ h₃ simp [h₅] - case h_3 rty₂ rty₃ => + case h_3 ety₁ l₁ ety₂ l₂ => + cases heq : decide (ety₁ = ety₂) <;> simp at heq + case false => + rw [if_neg] at h₂ + contradiction + cases heq₂ : (ety₁ == ety₂) + simp + simp at heq₂ + rw [heq₂] at heq + contradiction + case true => + rw [heq] at h₂ + rw [if_pos] at h₂ + simp at h₂ + cases ty₄ <;> simp at h₂ + rename_i ety' l' + have ⟨h₂, _⟩ := h₂ + rw [← h₂] + rw [← heq] + apply Level.lub_none_assoc + apply h₁ + simp + case h_4 rty₂ rty₃ => cases h₃ : lubRecordType rty₂ rty₃ <;> simp [h₃] at h₂ subst h₂ unfold lub? at h₁ ; unfold lub? @@ -530,7 +724,7 @@ theorem lub_assoc_none_some {ty₁ ty₂ ty₃ ty₄ : CedarType} cases h₄ : lubRecordType rty₁ rty₂ <;> simp [h₄] at h₁ have h₅ := lubRecordType_assoc_none_some h₄ h₃ simp [h₅] - case h_4 => + case h_5 => split at h₂ <;> try contradiction rename_i h₃ ; simp at h₂ subst h₂ h₃ @@ -621,7 +815,40 @@ theorem lub_assoc_some_some {ty₁ ty₂ ty₃ ty₄ ty₅ : CedarType} case int | string => subst h₁ h₂ simp [lub?] - case entity | ext => + case entity ety₁ l₁ ety₂ l₂ ety₃ l₃ => + have ⟨h₁, h₃⟩ := h₁ + have ⟨h₄, h₅⟩ := h₂ + cases ty₄ <;> cases ty₅ <;> simp [lub?] <;> simp at h₃ <;> simp at h₅ + rename_i h₆ ety₄ l₄ ety₅ l₅ + have ⟨h₇, h₈⟩ := h₃ + have ⟨h₉, h₁₀⟩ := h₅ + cases heq₁ : decide (ety₄ = ety₃) <;> cases heq₂ : decide (ety₁ = ety₅) <;> simp at heq₁ <;> simp at heq₂ + case entity.entity.false.false => + rw [if_neg] + rw [if_neg] + apply heq₂ + apply heq₁ + case entity.entity.false.true => + exfalso + apply heq₁ + subst h₇ h₆ h₄ + rfl + case entity.entity.true.false => + exfalso + apply heq₂ + subst h₆ h₉ + rfl + case entity.entity.true.true => + rw [heq₁] + rw [heq₂] + simp + apply And.intro + rw [← h₉] + rw [h₄] + apply Level.min_lemma + apply h₈ + apply h₁₀ + case ext => have ⟨hl₁, hr₁⟩ := h₁ have ⟨hl₂, hr₂⟩ := h₂ subst hl₁ hr₁ hl₂ hr₂ diff --git a/cedar-lean/Cedar/Thm/Validation/Typechecker/LitVar.lean b/cedar-lean/Cedar/Thm/Validation/Typechecker/LitVar.lean index cd60be75b..935b19431 100644 --- a/cedar-lean/Cedar/Thm/Validation/Typechecker/LitVar.lean +++ b/cedar-lean/Cedar/Thm/Validation/Typechecker/LitVar.lean @@ -26,8 +26,8 @@ namespace Cedar.Thm open Cedar.Spec open Cedar.Validation -theorem type_of_lit_is_sound {l : Prim} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} - (h₃ : typeOf (Expr.lit l) c₁ env = Except.ok (ty, c₂)) : +theorem type_of_lit_is_sound {l : Prim} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {level : Level} + (h₃ : typeOf (Expr.lit l) c₁ env (level == .infinite) = Except.ok (ty, c₂)) : GuardedCapabilitiesInvariant (Expr.lit l) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.lit l) request entities v ∧ InstanceOfType v ty := by @@ -49,9 +49,9 @@ theorem type_of_lit_is_sound {l : Prim} {c₁ c₂ : Capabilities} {env : Enviro apply InstanceOfType.instance_of_entity; simp [InstanceOfEntityType] } -theorem type_of_var_is_sound {var : Var} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_var_is_sound {var : Var} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.var var) c₁ env = Except.ok (ty, c₂)) : + (h₃ : typeOf (Expr.var var) c₁ env (l == .infinite) = Except.ok (ty, c₂)) : GuardedCapabilitiesInvariant (Expr.var var) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.var var) request entities v ∧ InstanceOfType v ty := by diff --git a/cedar-lean/Cedar/Thm/Validation/Typechecker/Or.lean b/cedar-lean/Cedar/Thm/Validation/Typechecker/Or.lean index 96b8cb00a..6421b3c08 100644 --- a/cedar-lean/Cedar/Thm/Validation/Typechecker/Or.lean +++ b/cedar-lean/Cedar/Thm/Validation/Typechecker/Or.lean @@ -26,14 +26,14 @@ open Cedar.Data open Cedar.Spec open Cedar.Validation -theorem type_of_or_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.or x₁ x₂) c env = Except.ok (ty, c')) : +theorem type_of_or_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.or x₁ x₂) c env (l == .infinite) = Except.ok (ty, c')) : ∃ bty₁ c₁, - typeOf x₁ c env = .ok (.bool bty₁, c₁) ∧ + typeOf x₁ c env (l == .infinite) = .ok (.bool bty₁, c₁) ∧ if bty₁ = BoolType.tt then ty = .bool BoolType.tt ∧ c' = ∅ else ∃ bty₂ c₂, - typeOf x₂ c env = .ok (.bool bty₂, c₂) ∧ + typeOf x₂ c env (l == .infinite) = .ok (.bool bty₂, c₂) ∧ if bty₁ = BoolType.ff then ty = .bool bty₂ ∧ c' = c₂ else if bty₂ = BoolType.tt @@ -43,7 +43,7 @@ theorem type_of_or_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Env else ty = .bool BoolType.anyBool ∧ c' = c₁ ∩ c₂ := by simp [typeOf] at h₁ - cases h₂ : typeOf x₁ c env <;> simp [h₂] at * + cases h₂ : typeOf x₁ c env (l == .infinite) <;> simp [h₂] at * rename_i res₁ simp [typeOfOr] at h₁ split at h₁ <;> simp [ok, err] at h₁ <;> @@ -54,7 +54,7 @@ theorem type_of_or_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Env have ⟨ht, hc⟩ := h₁ simp [←ht₁, ←hc₁, hc, ←ht] case ok.h_2 c₁ => - cases h₃ : typeOf x₂ c env <;> simp [h₃] at * + cases h₃ : typeOf x₂ c env (l == .infinite) <;> simp [h₃] at * rename_i res₂ split at h₁ <;> simp [ok, err] at h₁ rename_i bty₂ hr₂ @@ -68,7 +68,7 @@ theorem type_of_or_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Env cases bty₁ <;> simp at hneq₁ hneq₂ exists BoolType.anyBool, c₁ simp [←ht₁, ←hc₁] - cases h₃ : typeOf x₂ c env <;> simp [h₃] at * + cases h₃ : typeOf x₂ c env (l == .infinite) <;> simp [h₃] at * rename_i res₂ split at h₁ <;> simp [ok, err] at h₁ <;> rename_i hr₂ <;> @@ -84,10 +84,10 @@ theorem type_of_or_inversion {x₁ x₂ : Expr} {c c' : Capabilities} {env : Env simp [←hr₂, ←ht₁, ←hc₁] cases bty₂ <;> simp at * -theorem type_of_or_is_sound {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_or_is_sound {x₁ x₂ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.or x₁ x₂) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.or x₁ x₂) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih₁ : TypeOfIsSound x₁) (ih₂ : TypeOfIsSound x₂) : GuardedCapabilitiesInvariant (Expr.or x₁ x₂) c₂ request entities ∧ diff --git a/cedar-lean/Cedar/Thm/Validation/Typechecker/Record.lean b/cedar-lean/Cedar/Thm/Validation/Typechecker/Record.lean index 87a976b8f..5334db522 100644 --- a/cedar-lean/Cedar/Thm/Validation/Typechecker/Record.lean +++ b/cedar-lean/Cedar/Thm/Validation/Typechecker/Record.lean @@ -30,14 +30,14 @@ open Cedar.Validation def AttrValueHasAttrType (av : Attr × Value) (aqty : Attr × QualifiedType) : Prop := av.fst = aqty.fst ∧ InstanceOfType av.snd (Qualified.getType aqty.snd) -def AttrExprHasAttrType (c : Capabilities) (env : Environment) (ax : Attr × Expr) (aqty : Attr × QualifiedType) : Prop := +def AttrExprHasAttrType (c : Capabilities) (env : Environment) (l : Level) (ax : Attr × Expr) (aqty : Attr × QualifiedType) : Prop := ∃ ty' c', aqty = (ax.fst, .required ty') ∧ - typeOf ax.snd c env = Except.ok (ty', c') + typeOf ax.snd c env (l == .infinite) = Except.ok (ty', c') -theorem type_of_record_inversion_forall {axs : List (Attr × Expr)} {c : Capabilities} {env : Environment} {rty : List (Attr × QualifiedType)} - (h₁ : List.mapM (fun x => requiredAttr x.fst (typeOf x.snd c env)) axs = Except.ok rty) : - List.Forall₂ (AttrExprHasAttrType c env) axs rty +theorem type_of_record_inversion_forall {axs : List (Attr × Expr)} {c : Capabilities} {env : Environment} {rty : List (Attr × QualifiedType)} {l : Level} + (h₁ : List.mapM (fun x => requiredAttr x.fst (typeOf x.snd c env (l == .infinite))) axs = Except.ok rty) : + List.Forall₂ (AttrExprHasAttrType c env l) axs rty := by cases axs case nil => @@ -46,15 +46,15 @@ theorem type_of_record_inversion_forall {axs : List (Attr × Expr)} {c : Capabil cases rty case nil => simp [←List.mapM'_eq_mapM] at h₁ - cases h₂ : requiredAttr hd.fst (typeOf hd.snd c env) <;> simp [h₂] at h₁ - cases h₃ : List.mapM' (fun x => requiredAttr x.fst (typeOf x.snd c env)) tl <;> simp [h₃] at h₁ + cases h₂ : requiredAttr hd.fst (typeOf hd.snd c env (l == .infinite)) <;> simp [h₂] at h₁ + cases h₃ : List.mapM' (fun x => requiredAttr x.fst (typeOf x.snd c env (l == .infinite))) tl <;> simp [h₃] at h₁ simp [pure, Except.pure] at h₁ case cons hd' tl' => apply List.Forall₂.cons { simp [←List.mapM'_eq_mapM] at h₁ - cases h₂ : requiredAttr hd.fst (typeOf hd.snd c env) <;> simp [h₂] at h₁ - cases h₃ : List.mapM' (fun x => requiredAttr x.fst (typeOf x.snd c env)) tl <;> simp [h₃] at h₁ + cases h₂ : requiredAttr hd.fst (typeOf hd.snd c env (l == .infinite)) <;> simp [h₂] at h₁ + cases h₃ : List.mapM' (fun x => requiredAttr x.fst (typeOf x.snd c env (l == .infinite))) tl <;> simp [h₃] at h₁ simp [pure, Except.pure] at h₁ have ⟨hl₁, hr₁⟩ := h₁ rw [eq_comm] at hl₁ hr₁ ; subst hl₁ hr₁ @@ -68,26 +68,26 @@ theorem type_of_record_inversion_forall {axs : List (Attr × Expr)} {c : Capabil { apply type_of_record_inversion_forall simp [←List.mapM'_eq_mapM] at * - cases h₂ : requiredAttr hd.fst (typeOf hd.snd c env) <;> simp [h₂] at h₁ - cases h₃ : List.mapM' (fun x => requiredAttr x.fst (typeOf x.snd c env)) tl <;> simp [h₃] at h₁ + cases h₂ : requiredAttr hd.fst (typeOf hd.snd c env (l == .infinite)) <;> simp [h₂] at h₁ + cases h₃ : List.mapM' (fun x => requiredAttr x.fst (typeOf x.snd c env (l == .infinite))) tl <;> simp [h₃] at h₁ simp [pure, Except.pure] at h₁ simp [h₁] } -theorem type_of_record_inversion {axs : List (Attr × Expr)} {c c' : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.record axs) c env = Except.ok (ty, c')) : +theorem type_of_record_inversion {axs : List (Attr × Expr)} {c c' : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.record axs) c env (l == .infinite) = Except.ok (ty, c')) : c' = ∅ ∧ ∃ (rty : List (Attr × QualifiedType)), ty = .record (Map.make rty) ∧ - List.Forall₂ (AttrExprHasAttrType c env) axs rty + List.Forall₂ (AttrExprHasAttrType c env l) axs rty := by simp [typeOf, ok] at h₁ - cases h₂ : List.mapM₂ axs fun x => requiredAttr x.1.fst (typeOf x.1.snd c env) <;> simp [h₂] at h₁ + cases h₂ : List.mapM₂ axs fun x => requiredAttr x.1.fst (typeOf x.1.snd c env (l == .infinite)) <;> simp [h₂] at h₁ rename_i rty simp [h₁] exists rty ; simp [h₁] simp [List.mapM₂, List.attach₂] at h₂ - simp [List.mapM_pmap_subtype (fun (x : Attr × Expr) => requiredAttr x.fst (typeOf x.snd c env))] at h₂ + simp [List.mapM_pmap_subtype (fun (x : Attr × Expr) => requiredAttr x.fst (typeOf x.snd c env (l == .infinite)))] at h₂ exact type_of_record_inversion_forall h₂ theorem mk_vals_instance_of_mk_types₁ {a : Attr} {avs : List (Attr × Value)} {rty : List (Attr × QualifiedType)} @@ -163,11 +163,11 @@ theorem mk_vals_instance_of_mk_types {avs : List (Attr × Value)} {rty : List (A intro a qty h₂ _ exact mk_vals_instance_of_mk_types₃ h₁ h₂ -theorem head_of_vals_instance_of_head_of_types {xhd : Attr × Expr} {c₁ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {rhd : Attr × QualifiedType} {vhd : Attr × Value} +theorem head_of_vals_instance_of_head_of_types {xhd : Attr × Expr} {c₁ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {rhd : Attr × QualifiedType} {vhd : Attr × Value} {l : Level} (h₁ : TypeOfIsSound xhd.snd) (h₂ : CapabilitiesInvariant c₁ request entities) (h₃ : RequestAndEntitiesMatchEnvironment env request entities) - (h₄ : AttrExprHasAttrType c₁ env xhd rhd) + (h₄ : AttrExprHasAttrType c₁ env l xhd rhd) (h₅ : bindAttr xhd.fst (evaluate xhd.snd request entities) = Except.ok vhd) : AttrValueHasAttrType vhd rhd := by @@ -180,11 +180,11 @@ theorem head_of_vals_instance_of_head_of_types {xhd : Attr × Expr} {c₁ : Capa simp [EvaluatesTo, h₈] at h₁ ; subst h₁ h₅ simp [AttrValueHasAttrType, Qualified.getType, h₇] -theorem vals_instance_of_types {axs : List (Attr × Expr)} {c₁ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {rty : List (Attr × QualifiedType)} {avs : List (Attr × Value)} +theorem vals_instance_of_types {axs : List (Attr × Expr)} {c₁ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {rty : List (Attr × QualifiedType)} {avs : List (Attr × Value)} {l : Level} (ih : ∀ (axᵢ : Attr × Expr), axᵢ ∈ axs → TypeOfIsSound axᵢ.snd) (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : List.Forall₂ (AttrExprHasAttrType c₁ env) axs rty) + (h₃ : List.Forall₂ (AttrExprHasAttrType c₁ env l) axs rty) (h₄ : (axs.mapM fun x => bindAttr x.fst (evaluate x.snd request entities)) = Except.ok avs) : List.Forall₂ AttrValueHasAttrType avs rty := by @@ -205,15 +205,15 @@ theorem vals_instance_of_types {axs : List (Attr × Expr)} {c₁ : Capabilities} apply ih ; simp only [List.mem_cons, true_or] } { - apply @vals_instance_of_types xtl c₁ env request entities rtl vtl _ h₁ h₂ h₆ h₈ + apply @vals_instance_of_types xtl c₁ env request entities rtl vtl l _ h₁ h₂ h₆ h₈ intro ax h₉ ; apply ih ; simp [h₉] } -theorem type_of_record_is_sound_ok {axs : List (Attr × Expr)} {c₁ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {rty : List (Attr × QualifiedType)} {avs : List (Attr × Value)} +theorem type_of_record_is_sound_ok {axs : List (Attr × Expr)} {c₁ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {rty : List (Attr × QualifiedType)} {avs : List (Attr × Value)} {l : Level} (ih : ∀ (axᵢ : Attr × Expr), axᵢ ∈ axs → TypeOfIsSound axᵢ.snd) (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : List.Forall₂ (AttrExprHasAttrType c₁ env) axs rty) + (h₃ : List.Forall₂ (AttrExprHasAttrType c₁ env l) axs rty) (h₄ : (axs.mapM fun x => bindAttr x.fst (evaluate x.snd request entities)) = Except.ok avs) : InstanceOfType (Value.record (Map.make avs)) (CedarType.record (Map.make rty)) := by @@ -224,11 +224,11 @@ theorem type_of_record_is_sound_ok {axs : List (Attr × Expr)} {c₁ : Capabilit simp [List.Forallᵥ] at h₆ exact h₆ h₅ -theorem type_of_record_is_sound_err {axs : List (Attr × Expr)} {c₁ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {rty : List (Attr × QualifiedType)} {err : Error} +theorem type_of_record_is_sound_err {axs : List (Attr × Expr)} {c₁ : Capabilities} {env : Environment} {request : Request} {entities : Entities} {rty : List (Attr × QualifiedType)} {err : Error} {l : Level} (ih : ∀ (axᵢ : Attr × Expr), axᵢ ∈ axs → TypeOfIsSound axᵢ.snd) (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : List.Forall₂ (AttrExprHasAttrType c₁ env) axs rty) + (h₃ : List.Forall₂ (AttrExprHasAttrType c₁ env l) axs rty) (h₄ : (axs.mapM fun x => bindAttr x.fst (evaluate x.snd request entities)) = Except.error err) : err = Error.entityDoesNotExist ∨ err = Error.extensionError ∨ err = Error.arithBoundsError := by @@ -255,16 +255,16 @@ theorem type_of_record_is_sound_err {axs : List (Attr × Expr)} {c₁ : Capabili cases h₅ : tl.mapM f <;> simp [h₅, pure, Except.pure] at h₄ rw [eq_comm] at h₄ ; subst h₄ apply @type_of_record_is_sound_err - tl c₁ env request entities tl' err - (by intro axᵢ h ; apply ih ; simp [h]) + tl c₁ env request entities tl' err l + (by intro axᵢ h ; apply ih ; simp [h]) h₁ h₂ ht₃ (by simp [List.mapM₂, List.attach₂, List.mapM_pmap_subtype f, h₅]) -theorem type_of_record_is_sound {axs : List (Attr × Expr)} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_record_is_sound {axs : List (Attr × Expr)} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.record axs) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.record axs) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih : ∀ (axᵢ : Attr × Expr), axᵢ ∈ axs → TypeOfIsSound axᵢ.snd) : GuardedCapabilitiesInvariant (Expr.record axs) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.record axs) request entities v ∧ InstanceOfType v ty diff --git a/cedar-lean/Cedar/Thm/Validation/Typechecker/Set.lean b/cedar-lean/Cedar/Thm/Validation/Typechecker/Set.lean index d6ebb0f71..ff64eecb4 100644 --- a/cedar-lean/Cedar/Thm/Validation/Typechecker/Set.lean +++ b/cedar-lean/Cedar/Thm/Validation/Typechecker/Set.lean @@ -67,11 +67,11 @@ theorem foldlM_of_lub_assoc (ty₁ ty₂ : CedarType) (tys : List CedarType) : theorem type_of_set_tail {x xhd : Expr } {xtl : List Expr} {c : Capabilities} - {env : Environment} {ty hd : CedarType } {tl : List CedarType} - (h₁ : (List.mapM₁ (xhd :: xtl) fun x => justType (typeOf x.val c env)) = Except.ok (hd :: tl)) + {env : Environment} {ty hd : CedarType } {tl : List CedarType} {l : Level} + (h₁ : (List.mapM₁ (xhd :: xtl) fun x => justType (typeOf x.val c env (l == .infinite))) = Except.ok (hd :: tl)) (h₂ : List.foldlM lub? hd tl = some ty) (h₃ : List.Mem x xtl) : - ∃ ty', typeOf (Expr.set xtl) c env = Except.ok (.set ty', []) ∧ + ∃ ty', typeOf (Expr.set xtl) c env (l == .infinite) = Except.ok (.set ty', []) ∧ (ty' ⊔ ty) = some ty := by cases xtl @@ -84,14 +84,14 @@ theorem type_of_set_tail case nil => simp only [List.mapM₁, List.attach_def, List.pmap, List.mapM_cons, bind_assoc, pure_bind] at h₁ - simp only [List.mapM_pmap_subtype (fun x => justType (typeOf x c env))] at h₁ - cases h₄ : justType (typeOf xhd c env) <;> simp [h₄] at h₁ - cases h₅ : justType (typeOf xhd' c env) <;> simp [h₅] at h₁ - cases h₆ : List.mapM (fun x => justType (typeOf x c env)) xtl' <;> simp [h₆] at h₁ + simp only [List.mapM_pmap_subtype (fun x => justType (typeOf x c env (l == .infinite)))] at h₁ + cases h₄ : justType (typeOf xhd c env (l == .infinite)) <;> simp [h₄] at h₁ + cases h₅ : justType (typeOf xhd' c env (l == .infinite)) <;> simp [h₅] at h₁ + cases h₆ : List.mapM (fun x => justType (typeOf x c env (l == .infinite))) xtl' <;> simp [h₆] at h₁ simp [pure, Except.pure] at h₁ case cons hd' tl' => simp only [List.mapM₁, List.attach_def] at h₁ - rw [List.mapM_pmap_subtype (fun x => justType (typeOf x c env))] at h₁ + rw [List.mapM_pmap_subtype (fun x => justType (typeOf x c env (l == .infinite)))] at h₁ have h₁ := List.mapM_head_tail h₁ rw [←List.mapM₁_eq_mapM] at h₁ simp only [h₁, typeOfSet, List.empty_eq, Except.bind_ok] @@ -112,18 +112,18 @@ theorem type_of_set_tail subst h₆ assumption -theorem type_of_set_inversion {xs : List Expr} {c c' : Capabilities} {env : Environment} {sty : CedarType} - (h₁ : typeOf (Expr.set xs) c env = Except.ok (sty, c')) : +theorem type_of_set_inversion {xs : List Expr} {c c' : Capabilities} {env : Environment} {sty : CedarType} {l : Level} + (h₁ : typeOf (Expr.set xs) c env (l == .infinite) = Except.ok (sty, c')) : c' = ∅ ∧ ∃ ty, sty = .set ty ∧ ∀ xᵢ, xᵢ ∈ xs → ∃ tyᵢ cᵢ, - typeOf xᵢ c env = Except.ok (tyᵢ, cᵢ) ∧ + typeOf xᵢ c env (l == .infinite) = Except.ok (tyᵢ, cᵢ) ∧ (tyᵢ ⊔ ty) = .some ty := by simp [typeOf] at h₁ - cases h₂ : List.mapM₁ xs fun x => justType (typeOf x.val c env) <;> + cases h₂ : List.mapM₁ xs fun x => justType (typeOf x.val c env (l == .infinite)) <;> simp [h₂] at h₁ simp [typeOfSet] at h₁ split at h₁ <;> simp [ok, err] at h₁ @@ -137,10 +137,10 @@ theorem type_of_set_inversion {xs : List Expr} {c c' : Capabilities} {env : Envi cases h₄ case head xtl => simp only [List.mapM₁, List.attach_def, List.pmap, List.mapM_cons] at h₂ - rcases h₅ : justType (typeOf x c env) <;> + rcases h₅ : justType (typeOf x c env (l == .infinite)) <;> simp only [h₅, Except.bind_err] at h₂ - simp only [List.mapM_pmap_subtype (fun x => justType (typeOf x c env)), Except.bind_ok] at h₂ - rcases h₆ : List.mapM (fun x => justType (typeOf x c env)) xtl <;> + simp only [List.mapM_pmap_subtype (fun x => justType (typeOf x c env (l == .infinite))), Except.bind_ok] at h₂ + rcases h₆ : List.mapM (fun x => justType (typeOf x c env (l == .infinite))) xtl <;> simp only [h₆, pure, Except.pure, Except.bind_err, Except.bind_ok, Except.ok.injEq, List.cons.injEq] at h₂ have ⟨hl₂, hr₂⟩ := h₂ ; subst hl₂ hr₂ rename_i hdty tlty @@ -154,7 +154,7 @@ theorem type_of_set_inversion {xs : List Expr} {c c' : Capabilities} {env : Envi · exact foldlM_of_lub_is_LUB h₃ case tail xhd xtl h₄ => have ⟨ty', h₅, h₆⟩ := type_of_set_tail h₂ h₃ h₄ - have h₇ := @type_of_set_inversion xtl c ∅ env (.set ty') + have h₇ := @type_of_set_inversion xtl c ∅ env (.set ty') l simp only [h₅, List.empty_eq, CedarType.set.injEq, exists_and_right, exists_eq_left', true_and, true_implies] at h₇ specialize h₇ x h₄ @@ -173,19 +173,19 @@ theorem list_is_sound_implies_tail_is_sound {hd : Expr} {tl : List Expr} apply h₁ simp [h₂] -theorem list_is_typed_implies_tail_is_typed {hd : Expr} {tl : List Expr} {c₁ : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : ∀ (xᵢ : Expr), xᵢ ∈ hd :: tl → ∃ tyᵢ cᵢ, typeOf xᵢ c₁ env = Except.ok (tyᵢ, cᵢ) ∧ (tyᵢ ⊔ ty) = some ty) : - ∀ (xᵢ : Expr), xᵢ ∈ tl → ∃ tyᵢ cᵢ, typeOf xᵢ c₁ env = Except.ok (tyᵢ, cᵢ) ∧ (tyᵢ ⊔ ty) = some ty +theorem list_is_typed_implies_tail_is_typed {hd : Expr} {tl : List Expr} {c₁ : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : ∀ (xᵢ : Expr), xᵢ ∈ hd :: tl → ∃ tyᵢ cᵢ, typeOf xᵢ c₁ env (l == .infinite) = Except.ok (tyᵢ, cᵢ) ∧ (tyᵢ ⊔ ty) = some ty) : + ∀ (xᵢ : Expr), xᵢ ∈ tl → ∃ tyᵢ cᵢ, typeOf xᵢ c₁ env (l == .infinite) = Except.ok (tyᵢ, cᵢ) ∧ (tyᵢ ⊔ ty) = some ty := by intro xᵢ h₂ apply h₁ simp [h₂] -theorem type_of_set_is_sound_err {xs : List Expr} {c₁ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {err : Error} +theorem type_of_set_is_sound_err {xs : List Expr} {c₁ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {err : Error} {l : Level} (ih : ∀ (xᵢ : Expr), xᵢ ∈ xs → TypeOfIsSound xᵢ) (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₄ : ∀ (xᵢ : Expr), xᵢ ∈ xs → ∃ tyᵢ cᵢ, typeOf xᵢ c₁ env = Except.ok (tyᵢ, cᵢ) ∧ (tyᵢ ⊔ ty) = some ty) + (h₄ : ∀ (xᵢ : Expr), xᵢ ∈ xs → ∃ tyᵢ cᵢ, typeOf xᵢ c₁ env (l == .infinite) = Except.ok (tyᵢ, cᵢ) ∧ (tyᵢ ⊔ ty) = some ty) (h₅ : (xs.mapM fun x => evaluate x request entities) = Except.error err) : err = Error.entityDoesNotExist ∨ err = Error.extensionError ∨ @@ -212,17 +212,17 @@ theorem type_of_set_is_sound_err {xs : List Expr} {c₁ : Capabilities} {env : E simp [h₁₀, pure, Except.pure] at h₅ ; rw [eq_comm] at h₅ subst h₅ apply @type_of_set_is_sound_err - tl c₁ env ty request entities err + tl c₁ env ty request entities err l (list_is_sound_implies_tail_is_sound ih) h₁ h₂ (list_is_typed_implies_tail_is_typed h₄) exact h₁₀ -theorem type_of_set_is_sound_ok { xs : List Expr } { c₁ : Capabilities } { env : Environment } { request : Request } { entities : Entities } { ty : CedarType } { v : Value } { vs : List Value } +theorem type_of_set_is_sound_ok { xs : List Expr } { c₁ : Capabilities } { env : Environment } { request : Request } { entities : Entities } { ty : CedarType } { v : Value } { vs : List Value } {l : Level} (ih : ∀ (xᵢ : Expr), xᵢ ∈ xs → TypeOfIsSound xᵢ) (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : ∀ (xᵢ : Expr), xᵢ ∈ xs → ∃ tyᵢ cᵢ, typeOf xᵢ c₁ env = Except.ok (tyᵢ, cᵢ) ∧ (tyᵢ ⊔ ty) = some ty) + (h₃ : ∀ (xᵢ : Expr), xᵢ ∈ xs → ∃ tyᵢ cᵢ, typeOf xᵢ c₁ env (l == .infinite) = Except.ok (tyᵢ, cᵢ) ∧ (tyᵢ ⊔ ty) = some ty) (h₄ : (xs.mapM fun x => evaluate x request entities) = Except.ok vs) (h₅ : v ∈ vs): InstanceOfType v ty @@ -257,17 +257,17 @@ theorem type_of_set_is_sound_ok { xs : List Expr } { c₁ : Capabilities } { env exact instance_of_lub_left h₆ ihr case tail h₉ => apply @type_of_set_is_sound_ok - tl c₁ env request entities ty v vtl + tl c₁ env request entities ty v vtl l (list_is_sound_implies_tail_is_sound ih) h₁ h₂ (list_is_typed_implies_tail_is_typed h₃) _ h₉ simp [List.mapM₁, List.attach, List.mapM_pmap_subtype (evaluate · request entities), h₈] -theorem type_of_set_is_sound {xs : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {sty : CedarType} {request : Request} {entities : Entities} +theorem type_of_set_is_sound {xs : List Expr} {c₁ c₂ : Capabilities} {env : Environment} {sty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.set xs) c₁ env = Except.ok (sty, c₂)) + (h₃ : typeOf (Expr.set xs) c₁ env (l == .infinite) = Except.ok (sty, c₂)) (ih : ∀ (xᵢ : Expr), xᵢ ∈ xs → TypeOfIsSound xᵢ) : GuardedCapabilitiesInvariant (Expr.set xs) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.set xs) request entities v ∧ InstanceOfType v sty diff --git a/cedar-lean/Cedar/Thm/Validation/Typechecker/Types.lean b/cedar-lean/Cedar/Thm/Validation/Typechecker/Types.lean index 00f9bf4d2..be8fcde33 100644 --- a/cedar-lean/Cedar/Thm/Validation/Typechecker/Types.lean +++ b/cedar-lean/Cedar/Thm/Validation/Typechecker/Types.lean @@ -54,9 +54,9 @@ inductive InstanceOfType : Value → CedarType → Prop := InstanceOfType (.prim (.int _)) .int | instance_of_string : InstanceOfType (.prim (.string _)) .string - | instance_of_entity (e : EntityUID) (ety: EntityType) + | instance_of_entity (e : EntityUID) (ety: EntityType) (l : Level) (h₁ : InstanceOfEntityType e ety) : - InstanceOfType (.prim (.entityUID e)) (.entity ety) + InstanceOfType (.prim (.entityUID e)) (.entity ety l) | instance_of_set (s : Set Value) (ty : CedarType) (h₁ : forall v, v ∈ s → InstanceOfType v ty) : InstanceOfType (.set s) (.set ty) @@ -74,11 +74,141 @@ inductive InstanceOfType : Value → CedarType → Prop := InstanceOfType (.ext x) (.ext xty) def InstanceOfRequestType (request : Request) (reqty : RequestType) : Prop := - InstanceOfEntityType request.principal reqty.principal ∧ - request.action = reqty.action ∧ - InstanceOfEntityType request.resource reqty.resource ∧ + InstanceOfEntityType request.principal reqty.principal.fst ∧ + request.action = reqty.action.fst ∧ + InstanceOfEntityType request.resource reqty.resource.fst ∧ InstanceOfType request.context (.record reqty.context) +def InstanceOfRequestTypeLevel (request : Request) (reqty : RequestType) (l : Level) : Prop := + InstanceOfEntityType request.principal reqty.principal.fst ∧ + reqty.principal.snd = l ∧ + reqty.action.snd = l ∧ + reqty.resource.snd = l ∧ + request.action = reqty.action.fst ∧ + InstanceOfEntityType request.resource reqty.resource.fst ∧ + InstanceOfType request.context (.record reqty.context) + + +section +set_option hygiene false + +notation:10 μ " ⊢ " e " : " τ => WellFormed μ e τ + +inductive WellFormed : Entities → Value → CedarType → Prop := + | bool (μ : Entities) (b : Bool) (bty : BoolType) + (h₁ : InstanceOfBoolType b bty) : + μ ⊢ .prim (.bool b) : .bool bty + | int (μ : Entities) : + μ ⊢ .prim (.int _) : .int + | string (μ : Entities) : + μ ⊢ .prim (.string _) : .string + | set (μ : Entities) (s : Cedar.Data.Set Value) (ty : CedarType) + (h₁ : forall v, v ∈ s → (μ ⊢ v : ty)) : + μ ⊢ .set s : .set ty + | record (μ : Entities) (r : Cedar.Data.Map Attr Value) (rty : RecordType) + -- if an attribute is present in the record, then it is present in the type + (h₁ : ∀ (k : Attr), r.contains k → rty.contains k) + -- if an attribute is present, then it has the expected type + (h₂ : ∀ (k : Attr) (v : Value) (qty : QualifiedType), + r.find? k = some v → rty.find? k = some qty → (μ ⊢ v : qty.getType)) + -- required attributes are present + (h₃ : ∀ (k : Attr) (qty : QualifiedType), rty.find? k = some qty → qty.isRequired → r.contains k) : + μ ⊢ .record r : .record rty + | ext (μ : Entities) (x : Ext) (xty : ExtType) + (h₁ : InstanceOfExtType x xty) : + μ ⊢ .ext x : .ext xty + -- Now for the wacky cases: Entities + -- Any entity is well formed at level 0 + | entity₀ (μ : Entities) (e : EntityUID) (ety: EntityType) + (h₁ : InstanceOfEntityType e ety) : + μ ⊢ .prim (.entityUID e) : .entity ety (.finite 0) + -- Entities can be given `n` as long as all entities in the attributes can be given a level bounded from below by `n - 1` + | entityₙ (μ : Entities) (e : EntityUID) (ety : EntityType) (attrs : Cedar.Data.Map Attr Value) {l : Level} + (h₁ : InstanceOfEntityType e ety) + (h₂ : μ.attrs e = .ok attrs) + -- All attributes must be well formed + (h₃ : ∀ k v t', + (k,v) ∈ attrs.kvs → + (μ ⊢ v : t') + ) + -- All attributes must be bounded by `l - 1` + (h₄ : ∀ k v t', + (k,v) ∈ attrs.kvs → + level t' ≥ l.sub1 + ) : + μ ⊢ .prim (.entityUID e) : .entity ety l + +end + + +theorem WellFormed_is_instanceOf (μ : Entities) (v : Value) (t : CedarType) : + (μ ⊢ v : t) → + InstanceOfType v t + := by + intros h + cases v <;> cases h + case _ => + apply InstanceOfType.instance_of_bool + assumption + case _ => + apply InstanceOfType.instance_of_int + case _ => + apply InstanceOfType.instance_of_string + case _ => + apply InstanceOfType.instance_of_entity + assumption + case _ => + apply InstanceOfType.instance_of_entity + assumption + case _ => + apply InstanceOfType.instance_of_set + rename_i s ty h₁ + intros v h₂ + have h_v_wf := h₁ v h₂ + apply WellFormed_is_instanceOf + apply h_v_wf + case _ => + rename_i attrs rty h₁ h₂ h₃ + apply InstanceOfType.instance_of_record <;> try assumption + intros k v qty h₄ h₅ + have h_v_wf := h₃ k v qty h₄ h₅ + apply WellFormed_is_instanceOf + assumption + case _ => + apply InstanceOfType.instance_of_ext + assumption +termination_by sizeOf v +decreasing_by + all_goals simp_wf + all_goals try omega + case _ => + rename_i heq _ _ _ _ _ _ _ _ _ _ _ _ + subst heq + rename_i s _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ + simp + have hsize : sizeOf v < sizeOf s := by + apply Set.sizeOf_lt_of_mem + assumption + omega + case _ => + rename_i m heq _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ + rw [heq] + simp + have hsize : sizeOf v < sizeOf m := by + have h₁ : sizeOf (k,v) < sizeOf m.kvs := by + apply List.sizeOf_lt_of_mem + apply Map.find_means_mem + assumption + have h₂ : sizeOf v < sizeOf (k,v) := by + simp + omega + have h₃ : sizeOf m.kvs < sizeOf m := by + apply Map.sizeOf_lt_of_kvs + omega + omega + + + /-- For every entity in the store, 1. The entity's type is defined in the type store. @@ -108,6 +238,15 @@ def RequestAndEntitiesMatchEnvironment (env : Environment) (request : Request) ( InstanceOfEntitySchema entities env.ets ∧ InstanceOfActionSchema entities env.acts +def RequestAndEntitiesMatchEnvironmentLeveled (env : Environment) (request : Request) (entities : Entities) (l : Level) : Prop := + InstanceOfRequestTypeLevel request env.reqty l ∧ + InstanceOfEntitySchema entities env.ets ∧ + InstanceOfActionSchema entities env.acts ∧ + (entities ⊢ request.principal : .entity env.reqty.principal.fst env.reqty.principal.snd) ∧ + (entities ⊢ request.resource : .entity env.reqty.resource.fst env.reqty.resource.snd) ∧ + (entities ⊢ request.action : .entity env.reqty.action.fst.ty env.reqty.action.snd) ∧ + (entities ⊢ request.context : .record env.reqty.context) + ----- Theorems ----- theorem false_is_instance_of_ff : @@ -182,7 +321,7 @@ theorem instance_of_string_is_string {v₁ : Value} : exists y theorem instance_of_entity_type_is_entity {ety : EntityType} : - InstanceOfType v₁ (.entity ety) → + InstanceOfType v₁ (.entity ety l) → ∃ euid, euid.ty = ety ∧ v₁ = .prim (.entityUID euid) := by intro h₁ @@ -385,10 +524,10 @@ theorem type_is_inhabited (ty : CedarType) : | .string => exists (.prim (.string default)) apply InstanceOfType.instance_of_string - | .entity ety => + | .entity ety l => have ⟨euid, h₁⟩ := entity_type_is_inhabited ety exists (.prim (.entityUID euid)) - apply InstanceOfType.instance_of_entity _ _ h₁ + apply InstanceOfType.instance_of_entity _ _ _ h₁ | .set ty₁ => exists (.set Set.empty) apply InstanceOfType.instance_of_set @@ -496,7 +635,30 @@ theorem instance_of_lub_left {v : Value} {ty ty₁ ty₂ : CedarType} intro w h₅ specialize h₄ w h₅ apply instance_of_lub_left h₃ (by simp [h₄]) - case h_3 _ _ rty₁ rty₂ => + case h_3 ety₁ l₁ ety₂ l₂ => + cases heq : decide (ety₁ = ety₂) <;> simp at heq + case false => + rw [if_neg] at h₁ + contradiction + simp + apply heq + case true => + rw [if_pos] at h₁ + simp at h₁ + cases ty₁' <;> simp at hty₁ + rename_i ety' l' + have ⟨h₃, _h₄⟩ := hty₁ + cases ty <;> simp at h₁ + cases h₂ + apply InstanceOfType.instance_of_entity + have ⟨h₁, _h₂ ⟩ := h₁ + rename_i h + rw [← h₁] + rw [h₃] + apply h + rw [heq] + simp + case h_4 _ _ rty₁ rty₂ => cases h₃ : lubRecordType rty₁ rty₂ <;> simp [h₃] at h₁ rename_i rty have hl := lubRecordType_is_lub_of_record_types h₃ @@ -520,7 +682,7 @@ theorem instance_of_lub_left {v : Value} {ty ty₁ ty₂ : CedarType} have ⟨qty₁, h₉, h₁₀⟩ := lubRecord_find_implies_find_left hl h₇ apply h₆ a qty₁ h₉ simp [h₁₀, h₈] - case h_4 => + case h_5 => split at h₁ <;> simp at h₁ rename_i h₃ subst h₁ h₃ hty₁ hty₂ @@ -537,5 +699,27 @@ theorem instance_of_lub {v : Value} {ty ty₁ ty₂ : CedarType} · rw [lub_comm] at h₁ exact instance_of_lub_left h₁ h₃ +theorem request_leveled_instance_implies_instance {request : Request} {reqty : RequestType} {l : Level} : + InstanceOfRequestTypeLevel request reqty l → + InstanceOfRequestType request reqty + := by + intros h + unfold InstanceOfRequestType + unfold InstanceOfRequestTypeLevel at h + simp [h] + +theorem request_and_entity_match_level_implies_regular {env : Environment} {request : Request} {entities : Entities} {l : Level} : + RequestAndEntitiesMatchEnvironmentLeveled env request entities l → + RequestAndEntitiesMatchEnvironment env request entities + := by + intros h + unfold RequestAndEntitiesMatchEnvironment + unfold RequestAndEntitiesMatchEnvironmentLeveled at h + simp [h] + apply request_leveled_instance_implies_instance + have ⟨h₁, _⟩ := h + apply h₁ + + end Cedar.Thm diff --git a/cedar-lean/Cedar/Thm/Validation/Typechecker/UnaryApp.lean b/cedar-lean/Cedar/Thm/Validation/Typechecker/UnaryApp.lean index 6fe16b124..e28cdaa4d 100644 --- a/cedar-lean/Cedar/Thm/Validation/Typechecker/UnaryApp.lean +++ b/cedar-lean/Cedar/Thm/Validation/Typechecker/UnaryApp.lean @@ -26,15 +26,15 @@ open Cedar.Data open Cedar.Spec open Cedar.Validation -theorem type_of_not_inversion {x₁ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.unaryApp .not x₁) c₁ env = Except.ok (ty, c₂)) : +theorem type_of_not_inversion {x₁ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.unaryApp .not x₁) c₁ env (l == .infinite) = Except.ok (ty, c₂)) : c₂ = ∅ ∧ ∃ bty c₁', ty = .bool bty.not ∧ - typeOf x₁ c₁ env = Except.ok (.bool bty, c₁') + typeOf x₁ c₁ env (l == .infinite) = Except.ok (.bool bty, c₁') := by simp [typeOf] at h₁ - cases h₂ : typeOf x₁ c₁ env <;> simp [h₂] at h₁ + cases h₂ : typeOf x₁ c₁ env (l == .infinite) <;> simp [h₂] at h₁ case ok res => have ⟨ty₁, c₁'⟩ := res simp [typeOfUnaryApp] at h₁ @@ -46,10 +46,10 @@ theorem type_of_not_inversion {x₁ : Expr} {c₁ c₂ : Capabilities} {env : En · exists bty, c₁' simp only [and_true, h₁] -theorem type_of_not_is_sound {x₁ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_not_is_sound {x₁ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.unaryApp .not x₁) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.unaryApp .not x₁) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih : TypeOfIsSound x₁) : GuardedCapabilitiesInvariant (Expr.unaryApp .not x₁) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.unaryApp .not x₁) request entities v ∧ InstanceOfType v ty @@ -83,14 +83,14 @@ theorem type_of_not_is_sound {x₁ : Expr} {c₁ c₂ : Capabilities} {env : Env exact type_is_inhabited (CedarType.bool (BoolType.not bty)) } -theorem type_of_neg_inversion {x₁ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.unaryApp .neg x₁) c₁ env = Except.ok (ty, c₂)) : +theorem type_of_neg_inversion {x₁ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.unaryApp .neg x₁) c₁ env (l == .infinite) = Except.ok (ty, c₂)) : c₂ = ∅ ∧ ty = .int ∧ - ∃ c₁', typeOf x₁ c₁ env = Except.ok (.int, c₁') + ∃ c₁', typeOf x₁ c₁ env (l == .infinite) = Except.ok (.int, c₁') := by simp [typeOf] at h₁ - cases h₂ : typeOf x₁ c₁ env <;> simp [h₂] at h₁ + cases h₂ : typeOf x₁ c₁ env (l == .infinite) <;> simp [h₂] at h₁ case ok res => have ⟨ty₁, c₁'⟩ := res simp [typeOfUnaryApp] at h₁ @@ -98,10 +98,10 @@ theorem type_of_neg_inversion {x₁ : Expr} {c₁ c₂ : Capabilities} {env : En simp [ok] at h₁ simp [h₁] -theorem type_of_neg_is_sound {x₁ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_neg_is_sound {x₁ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.unaryApp .neg x₁) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.unaryApp .neg x₁) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih : TypeOfIsSound x₁) : GuardedCapabilitiesInvariant (Expr.unaryApp .neg x₁) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.unaryApp .neg x₁) request entities v ∧ InstanceOfType v ty @@ -128,14 +128,14 @@ theorem type_of_neg_is_sound {x₁ : Expr} {c₁ c₂ : Capabilities} {env : Env exact type_is_inhabited CedarType.int } -theorem type_of_like_inversion {x₁ : Expr} {p : Pattern} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.unaryApp (.like p) x₁) c₁ env = Except.ok (ty, c₂)) : +theorem type_of_like_inversion {x₁ : Expr} {p : Pattern} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.unaryApp (.like p) x₁) c₁ env (l == .infinite) = Except.ok (ty, c₂)) : c₂ = ∅ ∧ ty = .bool .anyBool ∧ - ∃ c₁', typeOf x₁ c₁ env = Except.ok (.string, c₁') + ∃ c₁', typeOf x₁ c₁ env (l == .infinite) = Except.ok (.string, c₁') := by simp [typeOf] at h₁ - cases h₂ : typeOf x₁ c₁ env <;> simp [h₂] at h₁ + cases h₂ : typeOf x₁ c₁ env (l == .infinite) <;> simp [h₂] at h₁ case ok res => have ⟨ty₁, c₁'⟩ := res simp [typeOfUnaryApp] at h₁ @@ -143,10 +143,10 @@ theorem type_of_like_inversion {x₁ : Expr} {p : Pattern} {c₁ c₂ : Capabili simp [ok] at h₁ simp [h₁] -theorem type_of_like_is_sound {x₁ : Expr} {p : Pattern} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_like_is_sound {x₁ : Expr} {p : Pattern} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.unaryApp (.like p) x₁) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.unaryApp (.like p) x₁) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih : TypeOfIsSound x₁) : GuardedCapabilitiesInvariant (Expr.unaryApp (.like p) x₁) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.unaryApp (.like p) x₁) request entities v ∧ InstanceOfType v ty @@ -167,37 +167,37 @@ theorem type_of_like_is_sound {x₁ : Expr} {p : Pattern} {c₁ c₂ : Capabilit exact type_is_inhabited (.bool .anyBool) } -theorem type_of_is_inversion {x₁ : Expr} {ety : EntityType} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} - (h₁ : typeOf (Expr.unaryApp (.is ety) x₁) c₁ env = Except.ok (ty, c₂)) : +theorem type_of_is_inversion {x₁ : Expr} {ety : EntityType} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {l : Level} + (h₁ : typeOf (Expr.unaryApp (.is ety) x₁) c₁ env (l == .infinite) = Except.ok (ty, c₂)) : c₂ = ∅ ∧ - ∃ ety' c₁', + ∃ ety' l' c₁', ty = (.bool (if ety = ety' then .tt else .ff)) ∧ - typeOf x₁ c₁ env = Except.ok (.entity ety', c₁') + typeOf x₁ c₁ env (l == .infinite) = Except.ok (.entity ety' l', c₁') := by simp [typeOf] at h₁ - cases h₂ : typeOf x₁ c₁ env <;> simp [h₂] at h₁ + cases h₂ : typeOf x₁ c₁ env (l == .infinite) <;> simp [h₂] at h₁ case ok res => have ⟨ty₁, c₁'⟩ := res simp [typeOfUnaryApp] at h₁ split at h₁ <;> try contradiction - case h_4 _ _ ety' h₃ => + case h_4 _ _ ety' l h₃ => simp only [UnaryOp.is.injEq] at h₃ subst h₃ simp [ok] at h₁ - apply And.intro + constructor · simp [h₁] - · exists ety', c₁' + · exists ety', l, c₁' simp only [h₁, and_self] -theorem type_of_is_is_sound {x₁ : Expr} {ety : EntityType} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_is_is_sound {x₁ : Expr} {ety : EntityType} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.unaryApp (.is ety) x₁) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.unaryApp (.is ety) x₁) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih : TypeOfIsSound x₁) : GuardedCapabilitiesInvariant (Expr.unaryApp (.is ety) x₁) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.unaryApp (.is ety) x₁) request entities v ∧ InstanceOfType v ty := by - have ⟨h₅, ety', c₁', h₆, h₄⟩ := type_of_is_inversion h₃ + have ⟨h₅, ety', l', c₁', h₆, h₄⟩ := type_of_is_inversion h₃ subst h₅; subst h₆ apply And.intro empty_guarded_capabilities_invariant have ⟨_, v₁, h₆, h₇⟩ := ih h₁ h₂ h₄ -- IH @@ -216,10 +216,10 @@ theorem type_of_is_is_sound {x₁ : Expr} {ety : EntityType} {c₁ c₂ : Capabi apply type_is_inhabited } -theorem type_of_unaryApp_is_sound {op₁ : UnaryOp} {x₁ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} +theorem type_of_unaryApp_is_sound {op₁ : UnaryOp} {x₁ : Expr} {c₁ c₂ : Capabilities} {env : Environment} {ty : CedarType} {request : Request} {entities : Entities} {l : Level} (h₁ : CapabilitiesInvariant c₁ request entities) (h₂ : RequestAndEntitiesMatchEnvironment env request entities) - (h₃ : typeOf (Expr.unaryApp op₁ x₁) c₁ env = Except.ok (ty, c₂)) + (h₃ : typeOf (Expr.unaryApp op₁ x₁) c₁ env (l == .infinite) = Except.ok (ty, c₂)) (ih : TypeOfIsSound x₁) : GuardedCapabilitiesInvariant (Expr.unaryApp op₁ x₁) c₂ request entities ∧ ∃ v, EvaluatesTo (Expr.unaryApp op₁ x₁) request entities v ∧ InstanceOfType v ty diff --git a/cedar-lean/Cedar/Thm/Validation/Validator.lean b/cedar-lean/Cedar/Thm/Validation/Validator.lean index 8d6333f89..6cded185e 100644 --- a/cedar-lean/Cedar/Thm/Validation/Validator.lean +++ b/cedar-lean/Cedar/Thm/Validation/Validator.lean @@ -35,27 +35,27 @@ def EvaluatesToBool (expr : Expr) (request : Request) (entities : Entities) : Pr def AllEvaluateToBool (policies : Policies) (request : Request) (entities : Entities) : Prop := ∀ policy ∈ policies, EvaluatesToBool policy.toExpr request entities -def RequestAndEntitiesMatchSchema (schema : Schema) (request : Request) (entities : Entities) :Prop := - ∃ env ∈ schema.toEnvironments, +def RequestAndEntitiesMatchSchema (schema : Schema) (request : Request) (entities : Entities) (l : Level) :Prop := + ∃ env ∈ schema.toEnvironments l, RequestAndEntitiesMatchEnvironment env request entities theorem action_matches_env (env : Environment) (request : Request) (entities : Entities) : RequestAndEntitiesMatchEnvironment env request entities → - request.action = env.reqty.action + request.action = env.reqty.action.fst := by intro h₀ simp only [RequestAndEntitiesMatchEnvironment, InstanceOfRequestType] at h₀ obtain ⟨ ⟨ _, h₁, _, _ ⟩ , _ , _⟩ := h₀ exact h₁ -theorem typecheck_policy_is_sound (policy : Policy) (env : Environment) (ty : CedarType) (request : Request) (entities : Entities) : +theorem typecheck_policy_is_sound (policy : Policy) (env : Environment) (ty : CedarType) (request : Request) (entities : Entities) (l : Level) : RequestAndEntitiesMatchEnvironment env request entities → - typecheckPolicy policy env = .ok ty → + typecheckPolicy policy l env = .ok ty → ∃ b : Bool, EvaluatesTo policy.toExpr request entities b := by intro h₁ h₂ simp only [typecheckPolicy] at h₂ - cases h₃ : typeOf (substituteAction env.reqty.action policy.toExpr) [] env <;> + cases h₃ : typeOf (substituteAction env.reqty.action.fst policy.toExpr) [] env (l == .infinite) <;> simp only [List.empty_eq, h₃] at h₂ split at h₂ <;> simp only [Except.ok.injEq] at h₂ rename_i cp ht @@ -94,14 +94,14 @@ theorem typecheck_policy_is_sound (policy : Policy) (env : Environment) (ty : Ce rw [action_matches_env] repeat assumption -theorem typecheck_policy_with_environments_is_sound (policy : Policy) (envs : List Environment) (request : Request) (entities : Entities) : +theorem typecheck_policy_with_environments_is_sound (policy : Policy) (envs : List Environment) (request : Request) (entities : Entities) (l : Level) : (∃ env ∈ envs, RequestAndEntitiesMatchEnvironment env request entities) → - typecheckPolicyWithEnvironments policy envs = .ok () → + typecheckPolicyWithEnvironments l policy envs = .ok () → ∃ b : Bool, EvaluatesTo policy.toExpr request entities b := by intro h₀ h₂ simp only [typecheckPolicyWithEnvironments] at h₂ - cases h₃ : List.mapM (typecheckPolicy policy) envs with + cases h₃ : List.mapM (typecheckPolicy policy l) envs with | error => simp only [h₃, Except.bind_err] at h₂ | ok ts => simp only [h₃, Except.bind_ok, ite_eq_right_iff, imp_false, Bool.not_eq_true, allFalse] at h₂ @@ -109,7 +109,7 @@ theorem typecheck_policy_with_environments_is_sound (policy : Policy) (envs : Li rw [List.mapM_ok_iff_forall₂] at h₃ have h₄ := List.forall₂_implies_all_left h₃ specialize h₄ env h₀ - obtain ⟨ty, ⟨h₄, h₅⟩⟩ := h₄ - exact typecheck_policy_is_sound policy env ty request entities h₁ h₅ - + obtain ⟨ty, ⟨_, h₅⟩⟩ := h₄ + exact typecheck_policy_is_sound policy env ty request entities l h₁ h₅ + end Cedar.Thm diff --git a/cedar-lean/Cedar/Thm/Validation/WellFormed.lean b/cedar-lean/Cedar/Thm/Validation/WellFormed.lean new file mode 100644 index 000000000..d27a2ebc4 --- /dev/null +++ b/cedar-lean/Cedar/Thm/Validation/WellFormed.lean @@ -0,0 +1,52 @@ +import Cedar.Spec.Entities +import Cedar.Spec +import Cedar.Validation +import Cedar.Thm.Data + +namespace Cedar.Validation +open Cedar.Spec + +def lowerBound (ty : CedarType) : Level := + match ty with + | .bool _ + | .int + | .string + | .ext _ + | .set _ + => .infinite + | .entity _ l => l + | .record members => + let levels := (List.map₁ members.kvs (λ prod => lowerBound prod.val.snd.getType) ) + List.foldr min .infinite levels +termination_by sizeOf ty +decreasing_by + all_goals simp_wf + all_goals try omega + have h₁ : sizeOf prod.val < sizeOf members.kvs := by + apply List.sizeOf_lt_of_mem + exact prod.property + have h₂ : sizeOf members.kvs < sizeOf members := by + apply Cedar.Data.Map.sizeOf_lt_of_kvs + have h₃ : sizeOf prod.val.snd.getType < sizeOf prod.val := by + cases prod.val + rename_i a b + simp [Prod.snd, Qualified.getType] + split <;> simp <;> omega + omega + +def InstanceOfBoolType : Bool → BoolType → Prop + | true, .tt => True + | false, .ff => True + | _, .anyBool => True + | _, _ => False + +def InstanceOfEntityType (e : EntityUID) (ety: EntityType) : Prop := + ety = e.ty + +def InstanceOfExtType : Ext → ExtType → Prop + | .decimal _, .decimal => True + | .ipaddr _, .ipAddr => True + | _, _ => False + + +end Cedar.Validation diff --git a/cedar-lean/Cedar/Validation/RequestEntityValidator.lean b/cedar-lean/Cedar/Validation/RequestEntityValidator.lean index 5185cdb6c..dc1df108f 100644 --- a/cedar-lean/Cedar/Validation/RequestEntityValidator.lean +++ b/cedar-lean/Cedar/Validation/RequestEntityValidator.lean @@ -57,7 +57,7 @@ def instanceOfType (v : Value) (ty : CedarType) : Bool := | .prim (.bool b), .bool bty => instanceOfBoolType b bty | .prim (.int _), .int => true | .prim (.string _), .string => true - | .prim (.entityUID e), .entity ety => instanceOfEntityType e ety + | .prim (.entityUID e), .entity ety _ => instanceOfEntityType e ety | .set s, .set ty => s.elts.attach.all (λ ⟨v, _⟩ => instanceOfType v ty) | .record r, .record rty => r.kvs.all (λ (k, _) => rty.contains k) && @@ -80,9 +80,9 @@ def instanceOfType (v : Value) (ty : CedarType) : Bool := omega def instanceOfRequestType (request : Request) (reqty : RequestType) : Bool := - instanceOfEntityType request.principal reqty.principal && - request.action == reqty.action && - instanceOfEntityType request.resource reqty.resource && + instanceOfEntityType request.principal reqty.principal.fst && + request.action == reqty.action.fst && + instanceOfEntityType request.resource reqty.resource.fst && instanceOfType request.context (.record reqty.context) /-- @@ -119,8 +119,8 @@ where def requestMatchesEnvironment (env : Environment) (request : Request) : Bool := instanceOfRequestType request env.reqty -def validateRequest (schema : Schema) (request : Request) : RequestValidationResult := - if ((schema.toEnvironments.any (requestMatchesEnvironment · request))) +def validateRequest (schema : Schema) (request : Request) (l : Level) : RequestValidationResult := + if (((schema.toEnvironments l).any (requestMatchesEnvironment · request))) then .ok () else .error (.typeError "request could not be validated in any environment") @@ -155,8 +155,8 @@ def updateSchema (schema : Schema) (actionSchemaEntities : Entities) : Schema := } (ty, ese) -def validateEntities (schema : Schema) (entities : Entities) : EntityValidationResult := - schema.toEnvironments.forM (entitiesMatchEnvironment · entities) +def validateEntities (schema : Schema) (entities : Entities) (l : Level) : EntityValidationResult := + (schema.toEnvironments l).forM (entitiesMatchEnvironment · entities) -- json diff --git a/cedar-lean/Cedar/Validation/Subtyping.lean b/cedar-lean/Cedar/Validation/Subtyping.lean index 05a6ab596..99360d845 100644 --- a/cedar-lean/Cedar/Validation/Subtyping.lean +++ b/cedar-lean/Cedar/Validation/Subtyping.lean @@ -51,6 +51,11 @@ mutual | .set s₁, .set s₂ => do let lub ← lub? s₁ s₂ .some (.set lub) + | .entity t₁ l₁, .entity t₂ l₂ => + if t₁ == t₂ then + .some (.entity t₁ (min l₁ l₂)) + else + .none | .record (.mk r₁), .record (.mk r₂) => do let lub ← lubRecordType r₁ r₂ .some (.record (Map.mk lub)) diff --git a/cedar-lean/Cedar/Validation/TypeSystem.lean b/cedar-lean/Cedar/Validation/TypeSystem.lean new file mode 100644 index 000000000..846775a94 --- /dev/null +++ b/cedar-lean/Cedar/Validation/TypeSystem.lean @@ -0,0 +1,19 @@ +/- + Copyright Cedar Contributors + + Licensed under the Apache License, Version 2.0 (the "License"); + you may not use this file except in compliance with the License. + You may obtain a copy of the License at + + https://www.apache.org/licenses/LICENSE-2.0 + + Unless required by applicable law or agreed to in writing, software + distributed under the License is distributed on an "AS IS" BASIS, + WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. + See the License for the specific language governing permissions and + limitations under the License. +-/ + +namespace Cedar.Validation + +end Cedar.Validation diff --git a/cedar-lean/Cedar/Validation/Typechecker.lean b/cedar-lean/Cedar/Validation/Typechecker.lean index 1524f98ab..0a791ee2c 100644 --- a/cedar-lean/Cedar/Validation/Typechecker.lean +++ b/cedar-lean/Cedar/Validation/Typechecker.lean @@ -23,6 +23,7 @@ open Cedar.Data open Cedar.Spec inductive TypeError where + | levelError (ety : EntityType) | lubErr (ty₁ : CedarType) (ty₂ : CedarType) | unexpectedType (ty : CedarType) | attrNotFound (ty : CedarType) (attr : Attr) @@ -41,7 +42,7 @@ abbrev ResultType := Except TypeError (CedarType × Capabilities) def ok (ty : CedarType) (c : Capabilities := ∅) : ResultType := .ok (ty, c) def err (e : TypeError) : ResultType := .error e -def typeOfLit (p : Prim) (env : Environment) : ResultType := +def typeOfLit (p : Prim) (env : Environment) (inf : Bool) : ResultType := match p with | .bool true => ok (.bool .tt) | .bool false => ok (.bool .ff) @@ -49,14 +50,18 @@ def typeOfLit (p : Prim) (env : Environment) : ResultType := | .string _ => ok .string | .entityUID uid => if env.ets.contains uid.ty || env.acts.contains uid - then ok (.entity uid.ty) + -- Type the entity at level 0 if we're checking a finite leve schema + -- This will forbid derefernces of entity literals + -- othewise give it level infinity + then ok (.entity uid.ty (if inf then Level.infinite else Level.finite 0 )) else err (.unknownEntity uid.ty) def typeOfVar (v : Var) (env : Environment) : ResultType := match v with - | .principal => ok (.entity env.reqty.principal) - | .action => ok (.entity env.reqty.action.ty) - | .resource => ok (.entity env.reqty.resource) + | .principal => ok (.entity env.reqty.principal.fst env.reqty.principal.snd) + | .action => + ok (.entity env.reqty.action.fst.ty env.reqty.action.snd) + | .resource => ok (.entity env.reqty.resource.fst env.reqty.resource.snd) | .context => ok (.record env.reqty.context) def typeOfIf (r₁ : CedarType × Capabilities) (r₂ r₃ : ResultType) : ResultType := @@ -79,11 +84,11 @@ def typeOfAnd (r₁ : CedarType × Capabilities) (r₂ : ResultType) : ResultTyp | (.bool ty₁, c₁) => do let (ty₂, c₂) ← r₂ match ty₂ with - | .bool .ff => ok (.bool .ff) - | .bool .tt => ok (.bool ty₁) (c₁ ∪ c₂) - | .bool _ => ok (.bool .anyBool) (c₁ ∪ c₂) - | _ => err (.unexpectedType ty₂) - | (ty₁, _) => err (.unexpectedType ty₁) + | .bool .ff => ok (.bool .ff) + | .bool .tt => ok (.bool ty₁) (c₁ ∪ c₂) + | .bool .anyBool => ok (.bool .anyBool) (c₁ ∪ c₂) + | _ => err (.unexpectedType ty₂) + | (ty₁, _) => err (.unexpectedType ty₁) def typeOfOr (r₁ : CedarType × Capabilities) (r₂ : ResultType) : ResultType := match r₁ with @@ -107,7 +112,7 @@ def typeOfUnaryApp (op : UnaryOp) (ty : CedarType) : ResultType := | .not, .bool x => ok (.bool x.not) | .neg, .int => ok .int | .like _, .string => ok (.bool .anyBool) - | .is ety₁, .entity ety₂ => ok (.bool (if ety₁ = ety₂ then .tt else .ff)) + | .is ety₁, .entity ety₂ _ => ok (.bool (if ety₁ = ety₂ then .tt else .ff)) | _, _ => err (.unexpectedType ty) def typeOfEq (ty₁ ty₂ : CedarType) (x₁ x₂ : Expr) : ResultType := @@ -118,7 +123,7 @@ def typeOfEq (ty₁ ty₂ : CedarType) (x₁ x₂ : Expr) : ResultType := | .some _ => ok (.bool .anyBool) | .none => match ty₁, ty₂ with - | .entity _, .entity _ => ok (.bool .ff) + | .entity _ _, .entity _ _ => ok (.bool .ff) | _, _ => err (.lubErr ty₁ ty₂) def entityUID? : Expr → Option EntityUID @@ -133,29 +138,29 @@ def actionUID? (x : Expr) (acts: ActionSchema) : Option EntityUID := do let uid ← entityUID? x if acts.contains uid then .some uid else .none --- x₁ in x₂ where x₁ has type ety₁ and x₂ has type ety₂ -def typeOfInₑ (ety₁ ety₂ : EntityType) (x₁ x₂ : Expr) (env : Environment) : BoolType := - match actionUID? x₁ env.acts, entityUID? x₂ with - | .some uid₁, .some uid₂ => - if env.acts.descendentOf uid₁ uid₂ - then .tt - else .ff - | _, _ => - if env.ets.descendentOf ety₁ ety₂ - then .anyBool - else .ff - --- x₁ in x₂ where x₁ has type ety₁ and x₂ has type (.set ety₂) -def typeOfInₛ (ety₁ ety₂ : EntityType) (x₁ x₂ : Expr) (env : Environment) : BoolType := - match actionUID? x₁ env.acts, entityUIDs? x₂ with - | .some uid₁, .some uids => - if uids.any (env.acts.descendentOf uid₁ ·) - then .tt - else .ff - | _, _ => - if env.ets.descendentOf ety₁ ety₂ - then .anyBool - else .ff +-- x₁ in x₂ where x₁ has type ety₁ at level l₁ and x₂ has type ety₂ +def typeOfInₑ (ety₁ ety₂ : EntityType) (l₁ : Level ) (x₁ x₂ : Expr) (env : Environment) : ResultType := + if l₁ > Level.zero then ok (.bool type) else err (.levelError ety₁) + where + type := match actionUID? x₁ env.acts, entityUID? x₂ with + | .some uid₁, .some uid₂ => + if env.acts.descendentOf uid₁ uid₂ + then .tt else .ff + | _, _ => + if env.ets.descendentOf ety₁ ety₂ + then .anyBool else .ff + +-- x₁ in x₂ where x₁ has type ety₁ at level l₁ and x₂ has type (.set ety₂) +def typeOfInₛ (ety₁ ety₂ : EntityType) (l₁ : Level) (x₁ x₂ : Expr) (env : Environment) : ResultType := + if l₁ > Level.zero then ok (.bool type) else err (.levelError ety₁) + where + type := match actionUID? x₁ env.acts, entityUIDs? x₂ with + | .some uid₁, .some uids => + if uids.any (env.acts.descendentOf uid₁ ·) + then .tt else .ff + | _, _ => + if env.ets.descendentOf ety₁ ety₂ + then .anyBool else .ff def ifLubThenBool (ty₁ ty₂ : CedarType) : ResultType := match ty₁ ⊔ ty₂ with @@ -165,8 +170,8 @@ def ifLubThenBool (ty₁ ty₂ : CedarType) : ResultType := def typeOfBinaryApp (op₂ : BinaryOp) (ty₁ ty₂ : CedarType) (x₁ x₂ : Expr) (env : Environment) : ResultType := match op₂, ty₁, ty₂ with | .eq, _, _ => typeOfEq ty₁ ty₂ x₁ x₂ - | .mem, .entity ety₁, .entity ety₂ => ok (.bool (typeOfInₑ ety₁ ety₂ x₁ x₂ env)) - | .mem, .entity ety₁, .set (.entity ety₂) => ok (.bool (typeOfInₛ ety₁ ety₂ x₁ x₂ env)) + | .mem, .entity ety₁ l₁, .entity ety₂ _ => (typeOfInₑ ety₁ ety₂ l₁ x₁ x₂ env) + | .mem, .entity ety₁ l₁, .set (.entity ety₂ _) => (typeOfInₛ ety₁ ety₂ l₁ x₁ x₂ env) | .less, .int, .int => ok (.bool .anyBool) | .lessEq, .int, .int => ok (.bool .anyBool) | .add, .int, .int => ok .int @@ -191,15 +196,63 @@ def actionType? (ety : EntityType) (acts: ActionSchema) : Bool := def typeOfHasAttr (ty : CedarType) (x : Expr) (a : Attr) (c : Capabilities) (env : Environment) : ResultType := match ty with | .record rty => hasAttrInRecord rty x a c true - | .entity ety => - match env.ets.attrs? ety with - | .some rty => hasAttrInRecord rty x a c false - | .none => - if actionType? ety env.acts - then ok (.bool .ff) -- action attributes not allowed - else err (.unknownEntity ety) + | .entity ety level => + let hasType := + match env.ets.attrs? ety with + | .some rty => hasAttrInRecord rty x a c false + | .none => + if actionType? ety env.acts + then ok (.bool .ff) -- action attributes not allowed + else err (.unknownEntity ety) + if level > Level.zero then hasType else err (.levelError ety) | _ => err (.unexpectedType ty) + +def setLevel (l: Level) (ty : CedarType) : CedarType := + match ty with + | .bool b => .bool b + | .int => .int + | .string => .string + | .entity ety l => .entity ety l + | .set ty => .set (setLevel l ty) + | .record rty => + .record (rty.mapOnValuesAttach (λ v => + match _h : v.val with + | .required ty => .required (setLevel l ty) + | .optional ty => .optional (setLevel l ty) + )) + -- .record (Map.make ((List.attach (rty.kvs)).map (λ pair => + -- let attrName := pair.val.fst + -- let ty := match _h : pair.val.snd with + -- | .required ty => .required (setLevel l ty) + -- | .optional ty => .optional (setLevel l ty) + -- (attrName, ty) + -- ))) + | .ext ty => .ext ty +termination_by sizeOf ty +decreasing_by + all_goals simp_wf + all_goals try omega + all_goals { + have h₁ := v.property + cases h₁ + rename_i k h₁ + have h₂ : sizeOf ty < sizeOf (k, v.val) := by + rw [_h] + simp + omega + have h₃ : sizeOf (k, v.val) < sizeOf rty.kvs := by + apply List.sizeOf_lt_of_mem + assumption + have h₄ : sizeOf rty.kvs < sizeOf rty := by + apply Cedar.Data.Map.sizeOf_lt_of_kvs + + omega + } + + + + def getAttrInRecord (ty : CedarType) (rty : RecordType) (x : Expr) (a : Attr) (c : Capabilities) : ResultType := match rty.find? a with | .some (.required aty) => ok aty @@ -209,10 +262,14 @@ def getAttrInRecord (ty : CedarType) (rty : RecordType) (x : Expr) (a : Attr) (c def typeOfGetAttr (ty : CedarType) (x : Expr) (a : Attr) (c : Capabilities) (env : Environment) : ResultType := match ty with | .record rty => getAttrInRecord ty rty x a c - | .entity ety => - match env.ets.attrs? ety with - | .some rty => getAttrInRecord ty rty x a c - | .none => err (.unknownEntity ety) + | .entity ety level => + if level > Level.zero then + match env.ets.attrs? ety with + | .some rty => do + let (type, caps) ← getAttrInRecord ty rty x a c + .ok (setLevel level.sub1 type, caps) + | .none => err (.unknownEntity ety) + else err $ .levelError ety | _ => err (.unexpectedType ty) def typeOfSet (tys : List CedarType) : ResultType := @@ -256,40 +313,40 @@ def typeOfCall (xfn : ExtFun) (tys : List CedarType) (xs : List Expr) : ResultTy -- Note: if x types as .tt or .ff, it is okay to replace x with the literal -- expression true or false if x can never throw an extension error at runtime. -- This is true for the current version of Cedar. -def typeOf (x : Expr) (c : Capabilities) (env : Environment) : ResultType := +def typeOf (x : Expr) (c : Capabilities) (env : Environment) (inf: Bool) : ResultType := match x with - | .lit p => typeOfLit p env + | .lit p => typeOfLit p env inf | .var v => typeOfVar v env | .ite x₁ x₂ x₃ => do - let (ty₁, c₁) ← typeOf x₁ c env - typeOfIf (ty₁, c₁) (typeOf x₂ (c ∪ c₁) env) (typeOf x₃ c env) + let (ty₁, c₁) ← typeOf x₁ c env inf + typeOfIf (ty₁, c₁) (typeOf x₂ (c ∪ c₁) env inf) (typeOf x₃ c env inf) | .and x₁ x₂ => do - let (ty₁, c₁) ← typeOf x₁ c env - typeOfAnd (ty₁, c₁) (typeOf x₂ (c ∪ c₁) env) + let (ty₁, c₁) ← typeOf x₁ c env inf + typeOfAnd (ty₁, c₁) (typeOf x₂ (c ∪ c₁) env inf) | .or x₁ x₂ => do - let (ty₁, c₁) ← typeOf x₁ c env - typeOfOr (ty₁, c₁) (typeOf x₂ c env) + let (ty₁, c₁) ← typeOf x₁ c env inf + typeOfOr (ty₁, c₁) (typeOf x₂ c env inf) | .unaryApp op₁ x₁ => do - let (ty₁, _) ← typeOf x₁ c env + let (ty₁, _) ← typeOf x₁ c env inf typeOfUnaryApp op₁ ty₁ | .binaryApp op₂ x₁ x₂ => do - let (ty₁, _) ← typeOf x₁ c env - let (ty₂, _) ← typeOf x₂ c env + let (ty₁, _) ← typeOf x₁ c env inf + let (ty₂, _) ← typeOf x₂ c env inf typeOfBinaryApp op₂ ty₁ ty₂ x₁ x₂ env | .hasAttr x₁ a => do - let (ty₁, _) ← typeOf x₁ c env + let (ty₁, _) ← typeOf x₁ c env inf typeOfHasAttr ty₁ x₁ a c env | .getAttr x₁ a => do - let (ty₁, _) ← typeOf x₁ c env + let (ty₁, _) ← typeOf x₁ c env inf typeOfGetAttr ty₁ x₁ a c env | .set xs => do - let tys ← xs.mapM₁ (λ ⟨x₁, _⟩ => justType (typeOf x₁ c env)) + let tys ← xs.mapM₁ (λ ⟨x₁, _⟩ => justType (typeOf x₁ c env inf)) typeOfSet tys | .record axs => do - let atys ← axs.mapM₂ (λ ⟨(a₁, x₁), _⟩ => requiredAttr a₁ (typeOf x₁ c env)) + let atys ← axs.mapM₂ (λ ⟨(a₁, x₁), _⟩ => requiredAttr a₁ (typeOf x₁ c env inf)) ok (.record (Map.make atys)) | .call xfn xs => do - let tys ← xs.mapM₁ (λ ⟨x₁, _⟩ => justType (typeOf x₁ c env)) + let tys ← xs.mapM₁ (λ ⟨x₁, _⟩ => justType (typeOf x₁ c env inf)) typeOfCall xfn tys xs ---- Derivations ----- diff --git a/cedar-lean/Cedar/Validation/Types.lean b/cedar-lean/Cedar/Validation/Types.lean index 40950b4bf..14d63e011 100644 --- a/cedar-lean/Cedar/Validation/Types.lean +++ b/cedar-lean/Cedar/Validation/Types.lean @@ -15,6 +15,7 @@ -/ import Cedar.Spec +import Cedar.Data.SizeOf namespace Cedar.Validation open Cedar.Data @@ -40,6 +41,13 @@ inductive Qualified (α : Type u) where | optional (a : α) | required (a : α) +instance : Functor Qualified where + map f q := + match q with + | .optional x => .optional (f x) + | .required x => .required (f x) + + def Qualified.getType {α} : Qualified α → α | optional a => a | required a => a @@ -48,15 +56,104 @@ def Qualified.isRequired {α} : Qualified α → Bool | optional _ => false | required _ => true +abbrev Level := Option Nat + +def Level.finite (n : Nat) : Level := .some n + +def Level.zero : Level := Level.finite 0 + +def Level.infinite : Level := .none + +def Level.sub1 (l : Level) : Level := + match l with + | .some n => .some (n - 1) + | .none => .none + +def Level.isInfinite (l : Level) : Bool := l.isNone + + +inductive LevelLT : Level → Level → Prop where + | finite₁ : ∀ n₁ n₂, + n₁ < n₂ → + LevelLT (.some n₁) (.some n₂) + | finite₂ : ∀ n₁, + LevelLT (.some n₁) .none + +-- Is there a way to get an anymous inductive? +instance : LT Level where + lt := LevelLT + +instance (l₁ l₂ : Level) : Decidable (l₁ < l₂) := by + cases l₁ <;> cases l₂ + case some.some n₁ n₂ => + cases hlt : decide (n₁ < n₂) + case false => + apply isFalse + simp at hlt + intros h + cases h + case _ h₁ => + omega + case true => + apply isTrue + apply LevelLT.finite₁ + apply of_decide_eq_true + apply hlt + case some.none n => + apply isTrue + apply LevelLT.finite₂ + case none.some n => + apply isFalse + intros h + cases h + case none.none => + apply isFalse + intros h + cases h + +deriving instance DecidableEq for Level + +instance : LE Level where + le l₁ l₂ := if l₁ == l₂ then true else if l₁ < l₂ then true else false + + inductive CedarType where | bool (bty : BoolType) | int | string - | entity (ety : EntityType) + | entity (ty : EntityType) (l : Level) | set (ty : CedarType) | record (rty : Map Attr (Qualified CedarType)) | ext (xty : ExtType) + + +def level (ty : CedarType) : Level := + match ty with + | .bool _ + | .int + | .string + | .set _ + | .ext _ => .infinite + | .record fields => + let levels : List Level := fields.kvs.map₁ (λ kv => level kv.val.snd.getType) + levels.foldl min .infinite + | .entity _ level => level +termination_by sizeOf ty +decreasing_by + simp_wf + have h₁ : sizeOf kv.val < sizeOf fields.kvs := by + apply List.sizeOf_lt_of_mem + apply kv.property + have h₂ : sizeOf kv.val.snd.getType < sizeOf kv.val := by + cases kv.val + case mk fst snd => + cases snd <;> simp [Qualified.getType] <;> omega + have h₃ : sizeOf fields.kvs < sizeOf fields := by + apply Cedar.Data.Map.sizeOf_lt_of_kvs + omega + + abbrev QualifiedType := Qualified CedarType abbrev RecordType := Map Attr QualifiedType @@ -103,9 +200,9 @@ structure Schema where acts : ActionSchema structure RequestType where - principal : EntityType - action : EntityUID - resource : EntityType + principal : EntityType × Level + action : (EntityUID × Level) + resource : EntityType × Level context : RecordType structure Environment where @@ -135,8 +232,10 @@ def decCedarType (a b : CedarType) : Decidable (a = b) := by case set.set t1 t2 => exact match decCedarType t1 t2 with | isTrue h => isTrue (by rw [h]) | isFalse _ => isFalse (by intro h; injection h; contradiction) - case entity.entity lub1 lub2 => exact match decEq lub1 lub2 with - | isTrue h => isTrue (by rw [h]) + case entity.entity lub1 l1 lub2 l2 => exact match decEq lub1 lub2 with + | isTrue h => match decEq l1 l2 with + | isTrue h' => isTrue (by rw [h]; rw[h']) + | isFalse _ => isFalse (by intro h; injection h; contradiction) | isFalse _ => isFalse (by intro h; injection h; contradiction) case record.record r1 r2 => exact match decAttrQualifiedCedarTypeMap r1 r2 with | isTrue h => isTrue (by rw [h]) diff --git a/cedar-lean/Cedar/Validation/Validator.lean b/cedar-lean/Cedar/Validation/Validator.lean index c56d4c02d..080a2e0d3 100644 --- a/cedar-lean/Cedar/Validation/Validator.lean +++ b/cedar-lean/Cedar/Validation/Validator.lean @@ -29,22 +29,22 @@ open Cedar.Data For a given action, compute the cross-product of the applicable principal and resource types. -/ -def ActionSchemaEntry.toRequestTypes (action : EntityUID) (entry : ActionSchemaEntry) : List RequestType := +def ActionSchemaEntry.toRequestTypes (level : Level) (action : EntityUID) (entry : ActionSchemaEntry) : List RequestType := entry.appliesToPrincipal.toList.foldl (fun acc principal => let reqtys : List RequestType := entry.appliesToResource.toList.map (fun resource => { - principal := principal, - action := action, - resource := resource, + principal := (principal, level), + action := (action, level), + resource := (resource, level), context := entry.context }) reqtys ++ acc) ∅ /-- Return every schema-defined environment. -/ -def Schema.toEnvironments (schema : Schema) : List Environment := +def Schema.toEnvironments (schema : Schema) (l : Level) : List Environment := let requestTypes : List RequestType := - schema.acts.toList.foldl (fun acc (action,entry) => entry.toRequestTypes action ++ acc) ∅ + schema.acts.toList.foldl (fun acc (action,entry) => entry.toRequestTypes l action ++ acc) ∅ requestTypes.map ({ ets := schema.ets, acts := schema.acts, @@ -105,9 +105,9 @@ def substituteAction (uid : EntityUID) (expr : Expr) : Expr := mapOnVars f expr /-- Check that a policy is Boolean-typed. -/ -def typecheckPolicy (policy : Policy) (env : Environment) : Except ValidationError CedarType := - let expr := substituteAction env.reqty.action policy.toExpr - match typeOf expr ∅ env with +def typecheckPolicy (policy : Policy) (l : Level) (env : Environment) : Except ValidationError CedarType := do + let expr := substituteAction env.reqty.action.fst policy.toExpr + match typeOf expr ∅ env (l == Level.infinite) with | .ok (ty, _) => if ty ⊑ .bool .anyBool then .ok ty @@ -118,17 +118,17 @@ def allFalse (tys : List CedarType) : Bool := tys.all (· == .bool .ff) /-- Check a policy under multiple environments. -/ -def typecheckPolicyWithEnvironments (policy : Policy) (envs : List Environment) : ValidationResult := do - let policyTypes ← envs.mapM (typecheckPolicy policy) +def typecheckPolicyWithEnvironments (l : Level) (policy : Policy) (envs : List Environment) : ValidationResult := do + let policyTypes ← envs.mapM (typecheckPolicy policy l) if allFalse policyTypes then .error (.impossiblePolicy policy.id) else .ok () /-- Analyze a set of policies to checks that all are boolean-typed, and that none are guaranteed to be false under all possible environments. -/ -def validate (policies : Policies) (schema : Schema) : ValidationResult := - let envs := schema.toEnvironments - policies.forM (typecheckPolicyWithEnvironments · envs) +def validate (policies : Policies) (schema : Schema) (l : Level) : ValidationResult := + let envs := schema.toEnvironments l + policies.forM ((typecheckPolicyWithEnvironments l) · envs) ----- Derivations ----- @@ -139,6 +139,7 @@ Lossy serialization of errors to Json. This serialization provides some extra information to DRT without having to derive `Lean.ToJson` for `Expr` and `CedarType`. -/ def validationErrorToJson : ValidationError → Lean.Json + | .typeError _ (.levelError _) => "Level Error" | .typeError _ (.lubErr _ _) => "lubErr" | .typeError _ (.unexpectedType _) => "unexpectedType" | .typeError _ (.attrNotFound _ _) => "attrNotFound"