Theory ResourcedDenotational

theory ResourcedDenotational
imports "Abstract-Denotational-Props" "CValue-Nominal" "C-restr"
begin

type_synonym CEnv = "var  (C  CValue)"

interpretation semantic_domain
  "Λ f . Λ r. CFn(Λ v. (f(v))|⇘r)"
  "Λ x y. (Λ r. (xr ↓CFn y|⇘r)r)"
  "Λ b r. CBb"
  "Λ scrut v1 v2 r. CB_project(scrutr)(v1r)(v2r)"
  "C_case".

notation ESem_syn ("𝒩⟦ _ ⟧⇘_"  [60,60] 60)
notation EvalHeapSem_syn  ("𝒩 _ _"  [0,0] 110)
notation HSem_syn ("𝒩⦃__"  [60,60] 60)
notation AHSem_bot ("𝒩⦃_"  [60] 60)

text ‹
Here we re-state the simplification rules, cleaned up by beta-reducing the locale parameters.
›

lemma CESem_simps:
  "𝒩⟦ Lam [x]. e ⟧⇘ρ= (Λ (Cr). CFn(Λ v. (𝒩⟦ e ⟧⇘ρ(x := v))|⇘r))"
  "𝒩⟦ App e x ⟧⇘ρ= (Λ (Cr). ((𝒩⟦ e ⟧⇘ρ)r ↓CFn ρ x|⇘r)r)"
  "𝒩⟦ Var x ⟧⇘ρ= (Λ (Cr). (ρ  x)  r)"
  "𝒩⟦ Bool b ⟧⇘ρ= (Λ (Cr). CB(Discr b))"
  "𝒩⟦ (scrut ? e1 : e2) ⟧⇘ρ= (Λ (Cr). CB_project((𝒩⟦ scrut ⟧⇘ρ)r)((𝒩⟦ e1 ⟧⇘ρ)r)((𝒩⟦ e2 ⟧⇘ρ)r))"
  "𝒩⟦ Let as body ⟧⇘ρ= (Λ (C  r). (𝒩⟦body⟧⇘𝒩⦃asρ)  r)"
  by (auto simp add: eta_cfun)

lemma CESem_bot[simp]:"(𝒩⟦ e ⟧⇘σ) = "
  by (nominal_induct e arbitrary: σ rule: exp_strong_induct) auto

lemma CHSem_bot[simp]:"((𝒩⦃ Γ ) x) = "
  by (cases "x  domA Γ") (auto simp add: lookup_HSem_heap lookup_HSem_other)

text ‹
Sometimes we do not care much about the resource usage and just want a simpler formula.
›

lemma CESem_simps_no_tick:
  "(𝒩⟦ Lam [x]. e ⟧⇘ρ)r  CFn(Λ v. (𝒩⟦ e ⟧⇘ρ(x := v))|⇘r)"
  "(𝒩⟦ App e x ⟧⇘ρ)r     ((𝒩⟦ e ⟧⇘ρ)r ↓CFn ρ x|⇘r)r"
  "𝒩⟦ Var x ⟧⇘ρ ρ x"
  "(𝒩⟦ (scrut ? e1 : e2) ⟧⇘ρ)r  CB_project((𝒩⟦ scrut ⟧⇘ρ)r)((𝒩⟦ e1 ⟧⇘ρ)r)((𝒩⟦ e2 ⟧⇘ρ)r)"
  "𝒩⟦ Let as body ⟧⇘ρ  𝒩⟦body⟧⇘𝒩⦃asρ⇙"
  apply -
  apply (rule below_trans[OF monofun_cfun_arg[OF below_C]], simp)
  apply (rule below_trans[OF monofun_cfun_arg[OF below_C]], simp)
  apply (rule cfun_belowI, rule below_trans[OF monofun_cfun_arg[OF below_C]], simp)
  apply (rule below_trans[OF monofun_cfun_arg[OF below_C]], simp)
  apply (rule cfun_belowI, rule below_trans[OF monofun_cfun_arg[OF below_C]], simp)
  done

lemma CELam_no_restr: "(𝒩⟦ Lam [x]. e ⟧⇘ρ)r  CFn(Λ v. (𝒩⟦ e ⟧⇘ρ(x := v)))"
proof-
  have "(𝒩⟦ Lam [x]. e ⟧⇘ρ)r  CFn(Λ v. (𝒩⟦ e ⟧⇘ρ(x := v))|⇘r)" by (rule CESem_simps_no_tick)
  also have "  CFn(Λ v. (𝒩⟦ e ⟧⇘ρ(x := v)))"
    by (intro cont2cont monofun_LAM below_trans[OF C_restr_below] monofun_cfun_arg below_refl fun_upd_mono)
  finally show ?thesis by this (intro cont2cont)
qed

lemma CEApp_no_restr: "(𝒩⟦ App e x ⟧⇘ρ)r     ((𝒩⟦ e ⟧⇘ρ)r ↓CFn ρ x)r"
proof-
  have "(𝒩⟦ App e x ⟧⇘ρ)r   ((𝒩⟦ e ⟧⇘ρ)r ↓CFn ρ x|⇘r)r" by (rule CESem_simps_no_tick)
  also have "ρ x|⇘r ρ x" by (rule C_restr_below)
  finally show ?thesis by this (intro cont2cont)
qed

end