theory LamEx
imports Nominal QuotMain
begin
atom_decl name
thm abs_fresh(1)
nominal_datatype rlam =
rVar "name"
| rApp "rlam" "rlam"
| rLam "name" "rlam"
print_theorems
function
rfv :: "rlam \<Rightarrow> name set"
where
rfv_var: "rfv (rVar a) = {a}"
| rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)"
| rfv_lam: "rfv (rLam a t) = (rfv t) - {a}"
sorry
termination rfv sorry
inductive
alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100)
where
a1: "a = b \<Longrightarrow> (rVar a) \<approx> (rVar b)"
| a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2"
| a3: "\<lbrakk>t \<approx> ([(a,b)]\<bullet>s); a \<notin> rfv (rLam b t)\<rbrakk> \<Longrightarrow> rLam a t \<approx> rLam b s"
print_theorems
lemma alpha_refl:
fixes t::"rlam"
shows "t \<approx> t"
apply(induct t rule: rlam.induct)
apply(simp add: a1)
apply(simp add: a2)
apply(rule a3)
apply(subst pt_swap_bij'')
apply(rule pt_name_inst)
apply(rule at_name_inst)
apply(simp)
apply(simp)
done
lemma alpha_EQUIV:
shows "EQUIV alpha"
sorry
quotient lam = rlam / alpha
apply(rule alpha_EQUIV)
done
print_quotients
quotient_def
Var :: "name \<Rightarrow> lam"
where
"Var \<equiv> rVar"
quotient_def
App :: "lam \<Rightarrow> lam \<Rightarrow> lam"
where
"App \<equiv> rApp"
quotient_def
Lam :: "name \<Rightarrow> lam \<Rightarrow> lam"
where
"Lam \<equiv> rLam"
thm Var_def
thm App_def
thm Lam_def
quotient_def
fv :: "lam \<Rightarrow> name set"
where
"fv \<equiv> rfv"
thm fv_def
(* definition of overloaded permutation function *)
(* for the lifted type lam *)
overloading
perm_lam \<equiv> "perm :: 'x prm \<Rightarrow> lam \<Rightarrow> lam" (unchecked)
begin
quotient_def
perm_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam"
where
"perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"
end
(*quotient_def (for lam)
abs_fun_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam"
where
"perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"*)
thm perm_lam_def
(* lemmas that need to lift *)
lemma pi_var_com:
fixes pi::"'x prm"
shows "(pi\<bullet>rVar a) \<approx> rVar (pi\<bullet>a)"
sorry
lemma pi_app_com:
fixes pi::"'x prm"
shows "(pi\<bullet>rApp t1 t2) \<approx> rApp (pi\<bullet>t1) (pi\<bullet>t2)"
sorry
lemma pi_lam_com:
fixes pi::"'x prm"
shows "(pi\<bullet>rLam a t) \<approx> rLam (pi\<bullet>a) (pi\<bullet>t)"
sorry
lemma real_alpha:
assumes a: "t = [(a,b)]\<bullet>s" "a\<sharp>[b].s"
shows "Lam a t = Lam b s"
using a
unfolding fresh_def supp_def
sorry
lemma perm_rsp[quot_rsp]:
"(op = ===> alpha ===> alpha) op \<bullet> op \<bullet>"
apply(auto)
(* this is propably true if some type conditions are imposed ;o) *)
sorry
lemma fresh_rsp:
"(op = ===> alpha ===> op =) fresh fresh"
apply(auto)
(* this is probably only true if some type conditions are imposed *)
sorry
lemma rVar_rsp[quot_rsp]:
"(op = ===> alpha) rVar rVar"
by (auto intro:a1)
lemma rApp_rsp[quot_rsp]: "(alpha ===> alpha ===> alpha) rApp rApp"
by (auto intro:a2)
lemma rLam_rsp[quot_rsp]: "(op = ===> alpha ===> alpha) rLam rLam"
apply(auto)
apply(rule a3)
apply(rule_tac t="[(x,x)]\<bullet>y" and s="y" in subst)
apply(rule sym)
apply(rule trans)
apply(rule pt_name3)
apply(rule at_ds1[OF at_name_inst])
apply(simp add: pt_name1)
apply(assumption)
apply(simp add: abs_fresh)
done
lemma rfv_rsp[quot_rsp]: "(alpha ===> op =) rfv rfv"
sorry
lemma rvar_inject: "rVar a \<approx> rVar b = (a = b)"
apply (auto)
apply (erule alpha.cases)
apply (simp_all add: rlam.inject alpha_refl)
done
ML {* val qty = @{typ "lam"} *}
ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp rfv_rsp} *}
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "lam" *}
ML {* fun lift_tac_lam lthy t = lift_tac lthy t [rel_eqv] [quot] *}
lemma pi_var: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Var a = Var (pi \<bullet> a)"
apply (tactic {* lift_tac_lam @{context} @{thm pi_var_com} 1 *})
done
lemma pi_app: "(pi\<Colon>('x \<times> 'x) list) \<bullet> App (x\<Colon>lam) (xa\<Colon>lam) = App (pi \<bullet> x) (pi \<bullet> xa)"
apply (tactic {* lift_tac_lam @{context} @{thm pi_app_com} 1 *})
done
lemma pi_lam: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Lam (a\<Colon>name) (x\<Colon>lam) = Lam (pi \<bullet> a) (pi \<bullet> x)"
apply (tactic {* lift_tac_lam @{context} @{thm pi_lam_com} 1 *})
done
lemma fv_var: "fv (Var (a\<Colon>name)) = {a}"
apply (tactic {* lift_tac_lam @{context} @{thm rfv_var} 1 *})
done
lemma fv_app: "fv (App (x\<Colon>lam) (xa\<Colon>lam)) = fv x \<union> fv xa"
apply (tactic {* lift_tac_lam @{context} @{thm rfv_app} 1 *})
done
lemma fv_lam: "fv (Lam (a\<Colon>name) (x\<Colon>lam)) = fv x - {a}"
apply (tactic {* lift_tac_lam @{context} @{thm rfv_lam} 1 *})
done
lemma a1: "(a\<Colon>name) = (b\<Colon>name) \<Longrightarrow> Var a = Var b"
apply (tactic {* lift_tac_lam @{context} @{thm a1} 1 *})
done
lemma a2: "\<lbrakk>(x\<Colon>lam) = (xa\<Colon>lam); (xb\<Colon>lam) = (xc\<Colon>lam)\<rbrakk> \<Longrightarrow> App x xb = App xa xc"
apply (tactic {* lift_tac_lam @{context} @{thm a2} 1 *})
done
lemma a3: "\<lbrakk>(x\<Colon>lam) = [(a\<Colon>name, b\<Colon>name)] \<bullet> (xa\<Colon>lam); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> Lam a x = Lam b xa"
apply (tactic {* lift_tac_lam @{context} @{thm a3} 1 *})
done
lemma alpha_cases: "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P;
\<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P;
\<And>x a b xa. \<lbrakk>a1 = Lam a x; a2 = Lam b xa; x = [(a, b)] \<bullet> xa; a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> P\<rbrakk>
\<Longrightarrow> P"
apply (tactic {* lift_tac_lam @{context} @{thm alpha.cases} 1 *})
done
lemma alpha_induct: "\<lbrakk>(qx\<Colon>lam) = (qxa\<Colon>lam); \<And>(a\<Colon>name) b\<Colon>name. a = b \<Longrightarrow> (qxb\<Colon>lam \<Rightarrow> lam \<Rightarrow> bool) (Var a) (Var b);
\<And>(x\<Colon>lam) (xa\<Colon>lam) (xb\<Colon>lam) xc\<Colon>lam. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc);
\<And>(x\<Colon>lam) (a\<Colon>name) (b\<Colon>name) xa\<Colon>lam.
\<lbrakk>x = [(a, b)] \<bullet> xa; qxb x ([(a, b)] \<bullet> xa); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> qxb (Lam a x) (Lam b xa)\<rbrakk>
\<Longrightarrow> qxb qx qxa"
apply (tactic {* lift_tac_lam @{context} @{thm alpha.induct} 1 *})
done
lemma var_inject: "(Var a = Var b) = (a = b)"
apply (tactic {* lift_tac_lam @{context} @{thm rvar_inject} 1 *})
done
lemma lam_induct:" \<lbrakk>\<And>name. P (Var name); \<And>lam1 lam2. \<lbrakk>P lam1; P lam2\<rbrakk> \<Longrightarrow> P (App lam1 lam2);
\<And>name lam. P lam \<Longrightarrow> P (Lam name lam)\<rbrakk> \<Longrightarrow> P lam"
apply (tactic {* lift_tac_lam @{context} @{thm rlam.induct} 1 *})
done
lemma var_supp:
shows "supp (Var a) = ((supp a)::name set)"
apply(simp add: supp_def)
apply(simp add: pi_var)
apply(simp add: var_inject)
done
lemma var_fresh:
fixes a::"name"
shows "(a\<sharp>(Var b)) = (a\<sharp>b)"
apply(simp add: fresh_def)
apply(simp add: var_supp)
done
(* Construction Site code *)
fun
option_map::"('a \<Rightarrow> 'b) \<Rightarrow> ('a noption) \<Rightarrow> ('b noption)"
where
"option_map f (nSome x) = nSome (f x)"
| "option_map f nNone = nNone"
fun
option_rel
where
"option_rel r (nSome x) (nSome y) = r x y"
| "option_rel r _ _ = False"
declare [[map noption = (option_map, option_rel)]]
lemma "option_map id = id"
sorry
lemma OPT_QUOTIENT:
assumes q: "QUOTIENT R Abs Rep"
shows "QUOTIENT (option_rel R) (option_map Abs) (option_map Rep)"
apply (unfold QUOTIENT_def)
apply (auto)
using q
apply (unfold QUOTIENT_def)
apply (case_tac "a :: 'b noption")
apply (simp)
apply (simp)
apply (case_tac "a :: 'b noption")
apply (simp only: option_map.simps)
apply (subst option_rel.simps)
(* Simp starts hanging so don't know how to continue *)
sorry
lemma real_alpha_pre:
assumes a: "t = [(a,b)]\<bullet>s" "a\<sharp>[b].s"
shows "rLam a t \<approx> rLam b s"
sorry
lemma perm_prs : "(id ---> option_map ABS_lam) ([b].REP_lam s) = [b].s"
sorry
lemma perm_prs2: "(id ---> option_map ABS_lam) ((id ---> option_map REP_lam) ([b].(s::lam))) = [b].s"
sorry
lemma real_alpha_lift:
"\<lbrakk>t = [(a, b)] \<bullet> s; a \<sharp> [b].s\<rbrakk> \<Longrightarrow> Lam a t = Lam b s"
apply (tactic {* procedure_tac @{context} @{thm real_alpha_pre} 1 *})
prefer 2
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* all_inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (tactic {* inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
apply (simp only: perm_prs)
prefer 2
apply (tactic {* all_inj_repabs_tac @{context} [quot] [rel_refl] [trans2] 1 *})
prefer 3
apply (tactic {* clean_tac @{context} [quot] 1 *})
thm all_prs
thm REP_ABS_RSP
thm ball_reg_eqv_range
thm perm_lam_def
apply (simp only: perm_prs)
(*apply (tactic {* regularize_tac @{context} [quot] 1 *})*)
sorry
done