Theory Semantics
section ‹Semantics of $\lambda\bullet$›
theory Semantics
imports FMap_Lemmas Syntax
begin
text ‹Avoid clash with substitution notation.›
no_notation inverse_divide (infixl ‹'/› 70)
text ‹Help automated provers with smallsteps.›
declare One_nat_def[simp del]
subsection ‹Equivariant Hash Function›
consts hash_real :: "term ⇒ hash"
nominal_function map_fixed :: "var ⇒ var list ⇒ term ⇒ term" where
"map_fixed fp l Unit = Unit" |
"map_fixed fp l (Var y) = (if y ∈ set l then (Var y) else (Var fp))" |
"atom y ♯ (fp, l) ⟹ map_fixed fp l (Lam y t) = (Lam y ((map_fixed fp (y # l) t)))" |
"atom y ♯ (fp, l) ⟹ map_fixed fp l (Rec y t) = (Rec y ((map_fixed fp (y # l) t)))" |
"map_fixed fp l (Inj1 t) = (Inj1 ((map_fixed fp l t)))" |
"map_fixed fp l (Inj2 t) = (Inj2 ((map_fixed fp l t)))" |
"map_fixed fp l (Pair t1 t2) = (Pair ((map_fixed fp l t1)) ((map_fixed fp l t2)))" |
"map_fixed fp l (Roll t) = (Roll ((map_fixed fp l t)))" |
"atom y ♯ (fp, l) ⟹ map_fixed fp l (Let t1 y t2) = (Let ((map_fixed fp l t1)) y ((map_fixed fp (y # l) t2)))" |
"map_fixed fp l (App t1 t2) = (App ((map_fixed fp l t1)) ((map_fixed fp l t2)))" |
"map_fixed fp l (Case t1 t2 t3) = (Case ((map_fixed fp l t1)) ((map_fixed fp l t2)) ((map_fixed fp l t3)))" |
"map_fixed fp l (Prj1 t) = (Prj1 ((map_fixed fp l t)))" |
"map_fixed fp l (Prj2 t) = (Prj2 ((map_fixed fp l t)))" |
"map_fixed fp l (Unroll t) = (Unroll ((map_fixed fp l t)))" |
"map_fixed fp l (Auth t) = (Auth ((map_fixed fp l t)))" |
"map_fixed fp l (Unauth t) = (Unauth ((map_fixed fp l t)))" |
"map_fixed fp l (Hash h) = (Hash h)" |
"map_fixed fp l (Hashed h t) = (Hashed h ((map_fixed fp l t)))"
using [[simproc del: alpha_lst defined_all]]
subgoal by (simp add: eqvt_def map_fixed_graph_aux_def)
subgoal by (erule map_fixed_graph.induct) (auto simp: fresh_star_def fresh_at_base)
apply clarify
subgoal for P fp l t
by (rule term.strong_exhaust[of t P "(fp, l)"]) (auto simp: fresh_star_def fresh_Pair)
apply (simp_all add: fresh_star_def fresh_at_base)
subgoal for y fp l t ya fpa la ta
apply (erule conjE)+
apply (erule Abs_lst1_fcb2'[where c = "(fp, l)"])
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff fresh_Pair)
done
subgoal for y fp l t ya fpa la ta
apply (erule conjE)+
apply (erule Abs_lst1_fcb2'[where c = "(fp, l)"])
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff fresh_Pair)
done
subgoal for y fp l t ya fpa la ta
apply (erule conjE)+
apply (erule Abs_lst1_fcb2'[where c = "(fp, l)"])
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff fresh_Pair)
done
done
nominal_termination (eqvt)
by lexicographic_order
definition hash where
"hash t = hash_real (map_fixed undefined [] t)"
lemma permute_map_list: "p ∙ l = map (λx. p ∙ x) l"
by (induct l) auto
lemma map_fixed_eqvt: "p ∙ l = l ⟹ map_fixed v l (p ∙ t) = map_fixed v l t"
proof (nominal_induct t avoiding: v l p rule: term.strong_induct)
case (Var x)
then show ?case
by (auto simp: term.supp supp_at_base permute_map_list list_eq_iff_nth_eq in_set_conv_nth)
next
case (Lam y e)
from Lam(1,2,3,5) Lam(4)[of p "y # l" v] show ?case
by (auto simp: fresh_perm)
next
case (Rec y e)
from Rec(1,2,3,5) Rec(4)[of p "y # l" v] show ?case
by (auto simp: fresh_perm)
next
case (Let e' y e)
from Let(1,2,3,6) Let(4)[of p l v] Let(5)[of p "y # l" v] show ?case
by (auto simp: fresh_perm)
qed (auto simp: permute_pure)
lemma hash_eqvt[eqvt]: "p ∙ hash t = hash (p ∙ t)"
unfolding permute_pure hash_def by (auto simp: map_fixed_eqvt)
lemma map_fixed_idle: "{x. ¬ atom x ♯ t} ⊆ set l ⟹ map_fixed v l t = t"
proof (nominal_induct t avoiding: v l rule: term.strong_induct)
case (Var x)
then show ?case
by (auto simp: subset_iff fresh_at_base)
next
case (Lam y e)
from Lam(1,2,4) Lam(3)[of "y # l" v] show ?case
by (auto simp: fresh_Pair Abs1_eq)
next
case (Rec y e)
from Rec(1,2,4) Rec(3)[of "y # l" v] show ?case
by (auto simp: fresh_Pair Abs1_eq)
next
case (Let e' y e)
from Let(1,2,5) Let(3)[of l v] Let(4)[of "y # l" v] show ?case
by (auto simp: fresh_Pair Abs1_eq)
qed (auto simp: subset_iff)
lemma map_fixed_idle_closed:
"closed t ⟹ map_fixed undefined [] t = t"
by (rule map_fixed_idle) auto
lemma map_fixed_inj_closed:
"closed t ⟹ closed u ⟹ map_fixed undefined [] t = map_fixed undefined [] u ⟹ t = u"
by (rule box_equals[OF _ map_fixed_idle_closed map_fixed_idle_closed])
lemma hash_eq_hash_real_closed:
assumes "closed t"
shows "hash t = hash_real t"
unfolding hash_def map_fixed_idle_closed[OF assms] ..
subsection ‹Substitution›
nominal_function subst_term :: "term ⇒ term ⇒ var ⇒ term" (‹_[_ '/ _]› [250, 200, 200] 250) where
"Unit[t' / x] = Unit" |
"(Var y)[t' / x] = (if x = y then t' else Var y)" |
"atom y ♯ (x, t') ⟹ (Lam y t)[t' / x] = Lam y (t[t' / x])" |
"atom y ♯ (x, t') ⟹ (Rec y t)[t' / x] = Rec y (t[t' / x])" |
"(Inj1 t)[t' / x] = Inj1 (t[t' / x])" |
"(Inj2 t)[t' / x] = Inj2 (t[t' / x])" |
"(Pair t1 t2)[t' / x] = Pair (t1[t' / x]) (t2[t' / x]) " |
"(Roll t)[t' / x] = Roll (t[t' / x])" |
"atom y ♯ (x, t') ⟹ (Let t1 y t2)[t' / x] = Let (t1[t' / x]) y (t2[t' / x])" |
"(App t1 t2)[t' / x] = App (t1[t' / x]) (t2[t' / x])" |
"(Case t1 t2 t3)[t' / x] = Case (t1[t' / x]) (t2[t' / x]) (t3[t' / x])" |
"(Prj1 t)[t' / x] = Prj1 (t[t' / x])" |
"(Prj2 t)[t' / x] = Prj2 (t[t' / x])" |
"(Unroll t)[t' / x] = Unroll (t[t' / x])" |
"(Auth t)[t' / x] = Auth (t[t' / x])" |
"(Unauth t)[t' / x] = Unauth (t[t' / x])" |
"(Hash h)[t' / x] = Hash h" |
"(Hashed h t)[t' / x] = Hashed h (t[t' / x])"
using [[simproc del: alpha_lst defined_all]]
subgoal by (simp add: eqvt_def subst_term_graph_aux_def)
subgoal by (erule subst_term_graph.induct) (auto simp: fresh_star_def fresh_at_base)
apply clarify
subgoal for P a b t
by (rule term.strong_exhaust[of a P "(b, t)"]) (auto simp: fresh_star_def fresh_Pair)
apply (simp_all add: fresh_star_def fresh_at_base)
subgoal
apply (erule conjE)
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff fresh_Pair)
done
subgoal
apply (erule conjE)
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff fresh_Pair)
done
subgoal
apply (erule conjE)
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff fresh_Pair)
done
done
nominal_termination (eqvt)
by lexicographic_order
type_synonym tenv = "(var, term) fmap"
nominal_function psubst_term :: "term ⇒ tenv ⇒ term" where
"psubst_term Unit f = Unit" |
"psubst_term (Var y) f = (case f $$ y of Some t ⇒ t | None ⇒ Var y)" |
"atom y ♯ f ⟹ psubst_term (Lam y t) f = Lam y (psubst_term t f)" |
"atom y ♯ f ⟹ psubst_term (Rec y t) f = Rec y (psubst_term t f)" |
"psubst_term (Inj1 t) f = Inj1 (psubst_term t f)" |
"psubst_term (Inj2 t) f = Inj2 (psubst_term t f)" |
"psubst_term (Pair t1 t2) f = Pair (psubst_term t1 f) (psubst_term t2 f) " |
"psubst_term (Roll t) f = Roll (psubst_term t f)" |
"atom y ♯ f ⟹ psubst_term (Let t1 y t2) f = Let (psubst_term t1 f) y (psubst_term t2 f)" |
"psubst_term (App t1 t2) f = App (psubst_term t1 f) (psubst_term t2 f)" |
"psubst_term (Case t1 t2 t3) f = Case (psubst_term t1 f) (psubst_term t2 f) (psubst_term t3 f)" |
"psubst_term (Prj1 t) f = Prj1 (psubst_term t f)" |
"psubst_term (Prj2 t) f = Prj2 (psubst_term t f)" |
"psubst_term (Unroll t) f = Unroll (psubst_term t f)" |
"psubst_term (Auth t) f = Auth (psubst_term t f)" |
"psubst_term (Unauth t) f = Unauth (psubst_term t f)" |
"psubst_term (Hash h) f = Hash h" |
"psubst_term (Hashed h t) f = Hashed h (psubst_term t f)"
using [[simproc del: alpha_lst defined_all]]
subgoal by (simp add: eqvt_def psubst_term_graph_aux_def)
subgoal by (erule psubst_term_graph.induct) (auto simp: fresh_star_def fresh_at_base)
apply clarify
subgoal for P a b
by (rule term.strong_exhaust[of a P b]) (auto simp: fresh_star_def fresh_Pair)
apply (simp_all add: fresh_star_def fresh_at_base)
subgoal by clarify
subgoal
apply (erule conjE)
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff)
done
subgoal
apply (erule conjE)
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff)
done
subgoal
apply (erule conjE)
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff)
done
done
nominal_termination (eqvt)
by lexicographic_order
nominal_function subst_type :: "ty ⇒ ty ⇒ tvar ⇒ ty" where
"subst_type One t' x = One" |
"subst_type (Fun t1 t2) t' x = Fun (subst_type t1 t' x) (subst_type t2 t' x)" |
"subst_type (Sum t1 t2) t' x = Sum (subst_type t1 t' x) (subst_type t2 t' x)" |
"subst_type (Prod t1 t2) t' x = Prod (subst_type t1 t' x) (subst_type t2 t' x)" |
"atom y ♯ (t', x) ⟹ subst_type (Mu y t) t' x = Mu y (subst_type t t' x)" |
"subst_type (Alpha y) t' x = (if y = x then t' else Alpha y)" |
"subst_type (AuthT t) t' x = AuthT (subst_type t t' x)"
using [[simproc del: alpha_lst defined_all]]
subgoal by (simp add: eqvt_def subst_type_graph_aux_def)
subgoal by (erule subst_type_graph.induct) (auto simp: fresh_star_def fresh_at_base)
apply clarify
subgoal for P a
by (rule ty.strong_exhaust[of a P]) (auto simp: fresh_star_def)
apply (simp_all add: fresh_star_def fresh_at_base)
subgoal
apply (erule conjE)
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff fresh_Pair)
done
done
nominal_termination (eqvt)
by lexicographic_order
lemma fresh_subst_term: "atom x ♯ t[t' / x'] ⟷ (x = x' ∨ atom x ♯ t) ∧ (atom x' ♯ t ∨ atom x ♯ t')"
by (nominal_induct t avoiding: t' x x' rule: term.strong_induct) (auto simp add: fresh_at_base)
lemma term_fresh_subst[simp]: "atom x ♯ t ⟹ atom x ♯ s ⟹ (atom (x::var)) ♯ t[s / y]"
by (nominal_induct t avoiding: s y rule: term.strong_induct) (auto)
lemma term_subst_idle[simp]: "atom y ♯ t ⟹ t[s / y] = t"
by (nominal_induct t avoiding: s y rule: term.strong_induct) (auto simp: fresh_Pair fresh_at_base)
lemma term_subst_subst: "atom y1 ≠ atom y2 ⟹ atom y1 ♯ s2 ⟹ t[s1 / y1][s2 / y2] = t[s2 / y2][s1[s2 / y2] / y1]"
by (nominal_induct t avoiding: y1 y2 s1 s2 rule: term.strong_induct) auto
lemma fresh_psubst:
fixes x :: var
assumes "atom x ♯ e" "atom x ♯ vs"
shows "atom x ♯ psubst_term e vs"
using assms
by (induct e vs rule: psubst_term.induct)
(auto simp: fresh_at_base elim: fresh_fmap_fresh_Some split: option.splits)
lemma fresh_subst_type:
"atom α ♯ subst_type τ τ' α' ⟷ ((α = α' ∨ atom α ♯ τ) ∧ (atom α' ♯ τ ∨ atom α ♯ τ'))"
by (nominal_induct τ avoiding: α α' τ' rule: ty.strong_induct) (auto simp add: fresh_at_base)
lemma type_fresh_subst[simp]: "atom x ♯ t ⟹ atom x ♯ s ⟹ (atom (x::tvar)) ♯ subst_type t s y"
by (nominal_induct t avoiding: s y rule: ty.strong_induct) (auto)
lemma type_subst_idle[simp]: "atom y ♯ t ⟹ subst_type t s y = t"
by (nominal_induct t avoiding: s y rule: ty.strong_induct) (auto simp: fresh_Pair fresh_at_base)
lemma type_subst_subst: "atom y1 ≠ atom y2 ⟹ atom y1 ♯ s2 ⟹
subst_type (subst_type t s1 y1) s2 y2 = subst_type (subst_type t s2 y2) (subst_type s1 s2 y2) y1"
by (nominal_induct t avoiding: y1 y2 s1 s2 rule: ty.strong_induct) auto
subsection ‹Weak Typing Judgement›
type_synonym tyenv = "(var, ty) fmap"
inductive judge_weak :: "tyenv ⇒ term ⇒ ty ⇒ bool" (‹_ ⊢⇩W _ : _› [150,0,150] 149) where
jw_Unit: "Γ ⊢⇩W Unit : One" |
jw_Var: "⟦ Γ $$ x = Some τ ⟧
⟹ Γ ⊢⇩W Var x : τ" |
jw_Lam: "⟦ atom x ♯ Γ; Γ(x $$:= τ⇩1) ⊢⇩W e : τ⇩2 ⟧
⟹ Γ ⊢⇩W Lam x e : Fun τ⇩1 τ⇩2" |
jw_App: "⟦ Γ ⊢⇩W e : Fun τ⇩1 τ⇩2; Γ ⊢⇩W e' : τ⇩1 ⟧
⟹ Γ ⊢⇩W App e e' : τ⇩2" |
jw_Let: "⟦ atom x ♯ (Γ, e⇩1); Γ ⊢⇩W e⇩1 : τ⇩1; Γ(x $$:= τ⇩1) ⊢⇩W e⇩2 : τ⇩2 ⟧
⟹ Γ ⊢⇩W Let e⇩1 x e⇩2 : τ⇩2" |
jw_Rec: "⟦ atom x ♯ Γ; atom y ♯ (Γ,x); Γ(x $$:= Fun τ⇩1 τ⇩2) ⊢⇩W Lam y e : Fun τ⇩1 τ⇩2 ⟧
⟹ Γ ⊢⇩W Rec x (Lam y e) : Fun τ⇩1 τ⇩2" |
jw_Inj1: "⟦ Γ ⊢⇩W e : τ⇩1 ⟧
⟹ Γ ⊢⇩W Inj1 e : Sum τ⇩1 τ⇩2" |
jw_Inj2: "⟦ Γ ⊢⇩W e : τ⇩2 ⟧
⟹ Γ ⊢⇩W Inj2 e : Sum τ⇩1 τ⇩2" |
jw_Case: "⟦ Γ ⊢⇩W e : Sum τ⇩1 τ⇩2; Γ ⊢⇩W e⇩1 : Fun τ⇩1 τ; Γ ⊢⇩W e⇩2 : Fun τ⇩2 τ ⟧
⟹ Γ ⊢⇩W Case e e⇩1 e⇩2 : τ" |
jw_Pair: "⟦ Γ ⊢⇩W e⇩1 : τ⇩1; Γ ⊢⇩W e⇩2 : τ⇩2 ⟧
⟹ Γ ⊢⇩W Pair e⇩1 e⇩2 : Prod τ⇩1 τ⇩2" |
jw_Prj1: "⟦ Γ ⊢⇩W e : Prod τ⇩1 τ⇩2 ⟧
⟹ Γ ⊢⇩W Prj1 e : τ⇩1" |
jw_Prj2: "⟦ Γ ⊢⇩W e : Prod τ⇩1 τ⇩2 ⟧
⟹ Γ ⊢⇩W Prj2 e : τ⇩2" |
jw_Roll: "⟦ atom α ♯ Γ; Γ ⊢⇩W e : subst_type τ (Mu α τ) α ⟧
⟹ Γ ⊢⇩W Roll e : Mu α τ" |
jw_Unroll: "⟦ atom α ♯ Γ; Γ ⊢⇩W e : Mu α τ ⟧
⟹ Γ ⊢⇩W Unroll e : subst_type τ (Mu α τ) α" |
jw_Auth: "⟦ Γ ⊢⇩W e : τ ⟧
⟹ Γ ⊢⇩W Auth e : τ" |
jw_Unauth: "⟦ Γ ⊢⇩W e : τ ⟧
⟹ Γ ⊢⇩W Unauth e : τ"
declare judge_weak.intros[simp]
declare judge_weak.intros[intro]
equivariance judge_weak
nominal_inductive judge_weak
avoids jw_Lam: x
| jw_Rec: x and y
| jw_Let: x
| jw_Roll: α
| jw_Unroll: α
by (auto simp: fresh_subst_type fresh_Pair)
text ‹Inversion rules for typing judgment.›
inductive_cases jw_Unit_inv[elim]: "Γ ⊢⇩W Unit : τ"
inductive_cases jw_Var_inv[elim]: "Γ ⊢⇩W Var x : τ"
lemma jw_Lam_inv[elim]:
assumes "Γ ⊢⇩W Lam x e : τ"
and "atom x ♯ Γ"
obtains τ⇩1 τ⇩2 where "τ = Fun τ⇩1 τ⇩2" "(Γ(x $$:= τ⇩1)) ⊢⇩W e : τ⇩2"
using assms proof (atomize_elim, nominal_induct Γ "Lam x e" τ avoiding: x e rule: judge_weak.strong_induct)
case (jw_Lam x Γ τ⇩1 t τ⇩2 y u)
then show ?case
by (auto simp: perm_supp_eq elim!:
iffD1[OF Abs_lst1_fcb2'[where f = "λx t (Γ, τ⇩1, τ⇩2). (Γ(x $$:= τ⇩1)) ⊢⇩W t : τ⇩2"
and c = "(Γ, τ⇩1, τ⇩2)" and a = x and b = y and x = t and y = u, unfolded prod.case],
rotated -1])
qed
lemma swap_permute_swap: "atom x ♯ π ⟹ atom y ♯ π ⟹ (x ↔ y) ∙ π ∙ (x ↔ y) ∙ t = π ∙ t"
by (subst permute_eqvt) (auto simp: flip_fresh_fresh)
lemma jw_Rec_inv[elim]:
assumes "Γ ⊢⇩W Rec x t : τ"
and "atom x ♯ Γ"
obtains y e τ⇩1 τ⇩2 where "atom y ♯ (Γ,x)" "t = Lam y e" "τ = Fun τ⇩1 τ⇩2" "Γ(x $$:= Fun τ⇩1 τ⇩2) ⊢⇩W Lam y e : Fun τ⇩1 τ⇩2"
using [[simproc del: alpha_lst]] assms
proof (atomize_elim, nominal_induct Γ "Rec x t" τ avoiding: x t rule: judge_weak.strong_induct)
case (jw_Rec x Γ y τ⇩1 τ⇩2 e z t)
then show ?case
proof (nominal_induct t avoiding: y x z rule: term.strong_induct)
case (Lam y' e')
show ?case
proof (intro exI conjI)
from Lam.prems show "atom y ♯ (Γ, z)" by simp
from Lam.hyps(1-3) Lam.prems show "Lam y' e' = Lam y ((y' ↔ y) ∙ e')"
by (subst term.eq_iff(3), intro Abs_lst_eq_flipI) (simp add: fresh_at_base)
from Lam.hyps(1-3) Lam.prems show "Γ(z $$:= Fun τ⇩1 τ⇩2) ⊢⇩W Lam y ((y' ↔ y) ∙ e') : Fun τ⇩1 τ⇩2"
by (elim judge_weak.eqvt[of "Γ(x $$:= Fun τ⇩1 τ⇩2)" "Lam y e" "Fun τ⇩1 τ⇩2" "(x ↔ z)", elim_format])
(simp add: perm_supp_eq Abs1_eq_iff fresh_at_base swap_permute_swap fresh_perm flip_commute)
qed simp
qed (simp_all add: Abs1_eq_iff)
qed
inductive_cases jw_Inj1_inv[elim]: "Γ ⊢⇩W Inj1 e : τ"
inductive_cases jw_Inj2_inv[elim]: "Γ ⊢⇩W Inj2 e : τ"
inductive_cases jw_Pair_inv[elim]: "Γ ⊢⇩W Pair e⇩1 e⇩2 : τ"
lemma jw_Let_inv[elim]:
assumes "Γ ⊢⇩W Let e⇩1 x e⇩2 : τ⇩2"
and "atom x ♯ (e⇩1, Γ)"
obtains τ⇩1 where "Γ ⊢⇩W e⇩1 : τ⇩1" "Γ(x $$:= τ⇩1) ⊢⇩W e⇩2 : τ⇩2"
using assms proof (atomize_elim, nominal_induct Γ "Let e⇩1 x e⇩2" τ⇩2 avoiding: e⇩1 x e⇩2 rule: judge_weak.strong_induct)
case (jw_Let x Γ e⇩1 τ⇩1 e⇩2 τ⇩2 x' e⇩2')
then show ?case
by (auto simp: fresh_Pair perm_supp_eq elim!:
iffD1[OF Abs_lst1_fcb2'[where f = "λx t (Γ, τ⇩1, τ⇩2). Γ(x $$:= τ⇩1) ⊢⇩W t : τ⇩2"
and c = "(Γ, τ⇩1, τ⇩2)" and a = x and b = x' and x = e⇩2 and y = e⇩2', unfolded prod.case],
rotated -1])
qed
inductive_cases jw_Prj1_inv[elim]: "Γ ⊢⇩W Prj1 e : τ⇩1"
inductive_cases jw_Prj2_inv[elim]: "Γ ⊢⇩W Prj2 e : τ⇩2"
inductive_cases jw_App_inv[elim]: "Γ ⊢⇩W App e e' : τ⇩2"
inductive_cases jw_Case_inv[elim]: "Γ ⊢⇩W Case e e⇩1 e⇩2 : τ"
inductive_cases jw_Auth_inv[elim]: "Γ ⊢⇩W Auth e : τ"
inductive_cases jw_Unauth_inv[elim]: "Γ ⊢⇩W Unauth e : τ"
lemma subst_type_perm_eq:
assumes "atom b ♯ t"
shows "subst_type t (Mu a t) a = subst_type ((a ↔ b) ∙ t) (Mu b ((a ↔ b) ∙ t)) b"
using assms proof -
have f: "atom a ♯ subst_type t (Mu a t) a" by (rule iffD2[OF fresh_subst_type]) simp
have "atom b ♯ subst_type t (Mu a t) a" by (auto simp: assms)
with f have "subst_type t (Mu a t) a = (a ↔ b) ∙ subst_type t (Mu a t) a"
by (simp add: flip_fresh_fresh)
then show "subst_type t (Mu a t) a = subst_type ((a ↔ b) ∙ t) (Mu b ((a ↔ b) ∙ t)) b"
by simp
qed
lemma jw_Roll_inv[elim]:
assumes "Γ ⊢⇩W Roll e : τ"
and "atom α ♯ (Γ, τ)"
obtains τ' where "τ = Mu α τ'" "Γ ⊢⇩W e : subst_type τ' (Mu α τ') α"
using assms [[simproc del: alpha_lst]]
proof (atomize_elim, nominal_induct Γ "Roll e" τ avoiding: e α rule: judge_weak.strong_induct)
case (jw_Roll α Γ e τ α')
then show ?case
by (auto simp: perm_supp_eq fresh_Pair fresh_at_base subst_type.eqvt
intro!: exI[of _ "(α ↔ α') ∙ τ"] Abs_lst_eq_flipI dest: judge_weak.eqvt[of _ _ _ "(α ↔ α')"])
qed
lemma jw_Unroll_inv[elim]:
assumes "Γ ⊢⇩W Unroll e : τ"
and "atom α ♯ (Γ, τ)"
obtains τ' where "τ = subst_type τ' (Mu α τ') α" "Γ ⊢⇩W e : Mu α τ'"
using assms proof (atomize_elim, nominal_induct Γ "Unroll e" τ avoiding: e α rule: judge_weak.strong_induct)
case (jw_Unroll α Γ e τ α')
then show ?case
by (auto simp: perm_supp_eq fresh_Pair subst_type_perm_eq fresh_subst_type
intro!: exI[of _ "(α ↔ α') ∙ τ"] dest: judge_weak.eqvt[of _ _ _ "(α ↔ α')"])
qed
text ‹Additional inversion rules based on type rather than term.›
inductive_cases jw_Prod_inv[elim]: "{$$} ⊢⇩W e : Prod τ⇩1 τ⇩2"
inductive_cases jw_Sum_inv[elim]: "{$$} ⊢⇩W e : Sum τ⇩1 τ⇩2"
lemma jw_Fun_inv[elim]:
assumes "{$$} ⊢⇩W v : Fun τ⇩1 τ⇩2" "value v"
obtains e x where "v = Lam x e ∨ v = Rec x e" "atom x ♯ (c::term)"
using assms [[simproc del: alpha_lst]]
proof (atomize_elim, nominal_induct "{$$} :: tyenv" v "Fun τ⇩1 τ⇩2" avoiding: τ⇩1 τ⇩2 rule: judge_weak.strong_induct)
case (jw_Lam x τ⇩1 e τ⇩2)
then obtain x' where "atom (x'::var) ♯ (c, e)" using finite_supp obtain_fresh' by blast
then have "[[atom x]]lst. e = [[atom x']]lst. (x ↔ x') ∙ e ∧ atom x' ♯ c"
by (simp add: Abs_lst_eq_flipI fresh_Pair)
then show ?case
by auto
next
case (jw_Rec x y τ⇩1 τ⇩2 e')
obtain x' where "atom (x'::var) ♯ (c, Lam y e')" using finite_supp obtain_fresh by blast
then have "[[atom x]]lst. Lam y e' = [[atom x']]lst. (x ↔ x') ∙ (Lam y e') ∧ atom x' ♯ c"
using Abs_lst_eq_flipI fresh_Pair by blast
then show ?case
by auto
qed simp_all
lemma jw_Mu_inv[elim]:
assumes "{$$} ⊢⇩W v : Mu α τ" "value v"
obtains v' where "v = Roll v'"
using assms by (atomize_elim, nominal_induct "{$$} :: tyenv" v "Mu α τ" rule: judge_weak.strong_induct) simp_all
subsection ‹Erasure of Authenticated Types›
nominal_function erase :: "ty ⇒ ty" where
"erase One = One" |
"erase (Fun τ⇩1 τ⇩2) = Fun (erase τ⇩1) (erase τ⇩2)" |
"erase (Sum τ⇩1 τ⇩2) = Sum (erase τ⇩1) (erase τ⇩2)" |
"erase (Prod τ⇩1 τ⇩2) = Prod (erase τ⇩1) (erase τ⇩2)" |
"erase (Mu α τ) = Mu α (erase τ)" |
"erase (Alpha α) = Alpha α" |
"erase (AuthT τ) = erase τ"
using [[simproc del: alpha_lst]]
subgoal by (simp add: eqvt_def erase_graph_aux_def)
subgoal by (erule erase_graph.induct) (auto simp: fresh_star_def fresh_at_base)
subgoal for P x
by (rule ty.strong_exhaust[of x P x]) (auto simp: fresh_star_def)
apply (simp_all add: fresh_star_def fresh_at_base)
subgoal
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff)
done
done
nominal_termination (eqvt)
by lexicographic_order
lemma fresh_erase_fresh:
assumes "atom x ♯ τ"
shows "atom x ♯ erase τ"
using assms by (induct τ rule: ty.induct) auto
lemma fresh_fmmap_erase_fresh:
assumes "atom x ♯ Γ"
shows "atom x ♯ fmmap erase Γ"
using assms by transfer simp
lemma erase_subst_type_shift[simp]:
"erase (subst_type τ τ' α) = subst_type (erase τ) (erase τ') α"
by (induct τ τ' α rule: subst_type.induct) (auto simp: fresh_Pair fresh_erase_fresh)
definition erase_env :: "tyenv ⇒ tyenv" where
"erase_env = fmmap erase"
subsection ‹Strong Typing Judgement›
inductive judge :: "tyenv ⇒ term ⇒ ty ⇒ bool" (‹_ ⊢ _ : _› [150,0,150] 149) where
j_Unit: "Γ ⊢ Unit : One" |
j_Var: "⟦ Γ $$ x = Some τ ⟧
⟹ Γ ⊢ Var x : τ" |
j_Lam: "⟦ atom x ♯ Γ; Γ(x $$:= τ⇩1) ⊢ e : τ⇩2 ⟧
⟹ Γ ⊢ Lam x e : Fun τ⇩1 τ⇩2" |
j_App: "⟦ Γ ⊢ e : Fun τ⇩1 τ⇩2; Γ ⊢ e' : τ⇩1 ⟧
⟹ Γ ⊢ App e e' : τ⇩2" |
j_Let: "⟦ atom x ♯ (Γ, e⇩1); Γ ⊢ e⇩1 : τ⇩1; Γ(x $$:= τ⇩1) ⊢ e⇩2 : τ⇩2 ⟧
⟹ Γ ⊢ Let e⇩1 x e⇩2 : τ⇩2" |
j_Rec: "⟦ atom x ♯ Γ; atom y ♯ (Γ,x); Γ(x $$:= Fun τ⇩1 τ⇩2) ⊢ Lam y e' : Fun τ⇩1 τ⇩2 ⟧
⟹ Γ ⊢ Rec x (Lam y e') : Fun τ⇩1 τ⇩2" |
j_Inj1: "⟦ Γ ⊢ e : τ⇩1 ⟧
⟹ Γ ⊢ Inj1 e : Sum τ⇩1 τ⇩2" |
j_Inj2: "⟦ Γ ⊢ e : τ⇩2 ⟧
⟹ Γ ⊢ Inj2 e : Sum τ⇩1 τ⇩2" |
j_Case: "⟦ Γ ⊢ e : Sum τ⇩1 τ⇩2; Γ ⊢ e⇩1 : Fun τ⇩1 τ; Γ ⊢ e⇩2 : Fun τ⇩2 τ ⟧
⟹ Γ ⊢ Case e e⇩1 e⇩2 : τ" |
j_Pair: "⟦ Γ ⊢ e⇩1 : τ⇩1; Γ ⊢ e⇩2 : τ⇩2 ⟧
⟹ Γ ⊢ Pair e⇩1 e⇩2 : Prod τ⇩1 τ⇩2" |
j_Prj1: "⟦ Γ ⊢ e : Prod τ⇩1 τ⇩2 ⟧
⟹ Γ ⊢ Prj1 e : τ⇩1" |
j_Prj2: "⟦ Γ ⊢ e : Prod τ⇩1 τ⇩2 ⟧
⟹ Γ ⊢ Prj2 e : τ⇩2" |
j_Roll: "⟦ atom α ♯ Γ; Γ ⊢ e : subst_type τ (Mu α τ) α ⟧
⟹ Γ ⊢ Roll e : Mu α τ" |
j_Unroll: "⟦ atom α ♯ Γ; Γ ⊢ e : Mu α τ ⟧
⟹ Γ ⊢ Unroll e : subst_type τ (Mu α τ) α" |
j_Auth: "⟦ Γ ⊢ e : τ ⟧
⟹ Γ ⊢ Auth e : AuthT τ" |
j_Unauth: "⟦ Γ ⊢ e : AuthT τ ⟧
⟹ Γ ⊢ Unauth e : τ"
declare judge.intros[intro]
equivariance judge
nominal_inductive judge
avoids j_Lam: x
| j_Rec: x and y
| j_Let: x
| j_Roll: α
| j_Unroll: α
by (auto simp: fresh_subst_type fresh_Pair)
lemma judge_imp_judge_weak:
assumes "Γ ⊢ e : τ"
shows "erase_env Γ ⊢⇩W e : erase τ"
using assms unfolding erase_env_def
by (induct Γ e τ rule: judge.induct) (simp_all add: fresh_Pair fresh_fmmap_erase_fresh fmmap_fmupd)
subsection ‹Shallow Projection›
nominal_function shallow :: "term ⇒ term" (‹⦇_⦈›) where
"⦇Unit⦈ = Unit" |
"⦇Var v⦈ = Var v" |
"⦇Lam x e⦈ = Lam x ⦇e⦈" |
"⦇Rec x e⦈ = Rec x ⦇e⦈" |
"⦇Inj1 e⦈ = Inj1 ⦇e⦈" |
"⦇Inj2 e⦈ = Inj2 ⦇e⦈" |
"⦇Pair e⇩1 e⇩2⦈ = Pair ⦇e⇩1⦈ ⦇e⇩2⦈" |
"⦇Roll e⦈ = Roll ⦇e⦈" |
"⦇Let e⇩1 x e⇩2⦈ = Let ⦇e⇩1⦈ x ⦇e⇩2⦈" |
"⦇App e⇩1 e⇩2⦈ = App ⦇e⇩1⦈ ⦇e⇩2⦈" |
"⦇Case e e⇩1 e⇩2⦈ = Case ⦇e⦈ ⦇e⇩1⦈ ⦇e⇩2⦈" |
"⦇Prj1 e⦈ = Prj1 ⦇e⦈" |
"⦇Prj2 e⦈ = Prj2 ⦇e⦈" |
"⦇Unroll e⦈ = Unroll ⦇e⦈" |
"⦇Auth e⦈ = Auth ⦇e⦈" |
"⦇Unauth e⦈ = Unauth ⦇e⦈" |
"⦇Hash h⦈ = Hash h" |
"⦇Hashed h e⦈ = Hash h"
using [[simproc del: alpha_lst]]
subgoal by (simp add: eqvt_def shallow_graph_aux_def)
subgoal by (erule shallow_graph.induct) (auto simp: fresh_star_def fresh_at_base)
subgoal for P a
by (rule term.strong_exhaust[of a P a]) (auto simp: fresh_star_def)
apply (simp_all add: fresh_star_def fresh_at_base)
subgoal
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff)
done
subgoal
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff)
done
subgoal
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff)
done
done
nominal_termination (eqvt)
by lexicographic_order
lemma fresh_shallow: "atom x ♯ e ⟹ atom x ♯ ⦇e⦈"
by (induct e rule: term.induct) auto
subsection ‹Small-step Semantics›