--- a/Quot/Examples/LamEx.thy Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,636 +0,0 @@
-theory LamEx
-imports Nominal "../Quotient" "../Quotient_List"
-begin
-
-atom_decl name
-
-datatype rlam =
- rVar "name"
-| rApp "rlam" "rlam"
-| rLam "name" "rlam"
-
-fun
- 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}"
-
-overloading
- perm_rlam \<equiv> "perm :: 'x prm \<Rightarrow> rlam \<Rightarrow> rlam" (unchecked)
-begin
-
-fun
- perm_rlam
-where
- "perm_rlam pi (rVar a) = rVar (pi \<bullet> a)"
-| "perm_rlam pi (rApp t1 t2) = rApp (perm_rlam pi t1) (perm_rlam pi t2)"
-| "perm_rlam pi (rLam a t) = rLam (pi \<bullet> a) (perm_rlam pi t)"
-
-end
-
-declare perm_rlam.simps[eqvt]
-
-instance rlam::pt_name
- apply(default)
- apply(induct_tac [!] x rule: rlam.induct)
- apply(simp_all add: pt_name2 pt_name3)
- done
-
-instance rlam::fs_name
- apply(default)
- apply(induct_tac [!] x rule: rlam.induct)
- apply(simp add: supp_def)
- apply(fold supp_def)
- apply(simp add: supp_atm)
- apply(simp add: supp_def Collect_imp_eq Collect_neg_eq)
- apply(simp add: supp_def)
- apply(simp add: supp_def Collect_imp_eq Collect_neg_eq[symmetric])
- apply(fold supp_def)
- apply(simp add: supp_atm)
- done
-
-declare set_diff_eqvt[eqvt]
-
-lemma rfv_eqvt[eqvt]:
- fixes pi::"name prm"
- shows "(pi\<bullet>rfv t) = rfv (pi\<bullet>t)"
-apply(induct t)
-apply(simp_all)
-apply(simp add: perm_set_eq)
-apply(simp add: union_eqvt)
-apply(simp add: set_diff_eqvt)
-apply(simp add: perm_set_eq)
-done
-
-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: "\<exists>pi::name prm. (rfv t - {a} = rfv s - {b} \<and> (rfv t - {a})\<sharp>* pi \<and> (pi \<bullet> t) \<approx> s \<and> (pi \<bullet> a) = b)
- \<Longrightarrow> rLam a t \<approx> rLam b s"
-
-
-(* should be automatic with new version of eqvt-machinery *)
-lemma alpha_eqvt:
- fixes pi::"name prm"
- shows "t \<approx> s \<Longrightarrow> (pi \<bullet> t) \<approx> (pi \<bullet> s)"
-apply(induct rule: alpha.induct)
-apply(simp add: a1)
-apply(simp add: a2)
-apply(simp)
-apply(rule a3)
-apply(erule conjE)
-apply(erule exE)
-apply(erule conjE)
-apply(rule_tac x="pi \<bullet> pia" in exI)
-apply(rule conjI)
-apply(rule_tac pi1="rev pi" in perm_bij[THEN iffD1])
-apply(perm_simp add: eqvts)
-apply(rule conjI)
-apply(rule_tac pi1="rev pi" in pt_fresh_star_bij(1)[OF pt_name_inst at_name_inst, THEN iffD1])
-apply(perm_simp add: eqvts)
-apply(rule conjI)
-apply(subst perm_compose[symmetric])
-apply(simp)
-apply(subst perm_compose[symmetric])
-apply(simp)
-done
-
-lemma alpha_refl:
- shows "t \<approx> t"
-apply(induct t rule: rlam.induct)
-apply(simp add: a1)
-apply(simp add: a2)
-apply(rule a3)
-apply(rule_tac x="[]" in exI)
-apply(simp_all add: fresh_star_def fresh_list_nil)
-done
-
-lemma alpha_sym:
- shows "t \<approx> s \<Longrightarrow> s \<approx> t"
-apply(induct rule: alpha.induct)
-apply(simp add: a1)
-apply(simp add: a2)
-apply(rule a3)
-apply(erule exE)
-apply(rule_tac x="rev pi" in exI)
-apply(simp)
-apply(simp add: fresh_star_def fresh_list_rev)
-apply(rule conjI)
-apply(erule conjE)+
-apply(rotate_tac 3)
-apply(drule_tac pi="rev pi" in alpha_eqvt)
-apply(perm_simp)
-apply(rule pt_bij2[OF pt_name_inst at_name_inst])
-apply(simp)
-done
-
-lemma alpha_trans:
- shows "t1 \<approx> t2 \<Longrightarrow> t2 \<approx> t3 \<Longrightarrow> t1 \<approx> t3"
-apply(induct arbitrary: t3 rule: alpha.induct)
-apply(erule alpha.cases)
-apply(simp_all)
-apply(simp add: a1)
-apply(rotate_tac 4)
-apply(erule alpha.cases)
-apply(simp_all)
-apply(simp add: a2)
-apply(rotate_tac 1)
-apply(erule alpha.cases)
-apply(simp_all)
-apply(erule conjE)+
-apply(erule exE)+
-apply(erule conjE)+
-apply(rule a3)
-apply(rule_tac x="pia @ pi" in exI)
-apply(simp add: fresh_star_def fresh_list_append)
-apply(simp add: pt_name2)
-apply(drule_tac x="rev pia \<bullet> sa" in spec)
-apply(drule mp)
-apply(rotate_tac 8)
-apply(drule_tac pi="rev pia" in alpha_eqvt)
-apply(perm_simp)
-apply(rotate_tac 11)
-apply(drule_tac pi="pia" in alpha_eqvt)
-apply(perm_simp)
-done
-
-lemma alpha_equivp:
- shows "equivp alpha"
-apply(rule equivpI)
-unfolding reflp_def symp_def transp_def
-apply(auto intro: alpha_refl alpha_sym alpha_trans)
-done
-
-lemma alpha_rfv:
- shows "t \<approx> s \<Longrightarrow> rfv t = rfv s"
-apply(induct rule: alpha.induct)
-apply(simp)
-apply(simp)
-apply(simp)
-done
-
-quotient_type lam = rlam / alpha
- by (rule alpha_equivp)
-
-
-quotient_definition
- "Var :: name \<Rightarrow> lam"
-is
- "rVar"
-
-quotient_definition
- "App :: lam \<Rightarrow> lam \<Rightarrow> lam"
-is
- "rApp"
-
-quotient_definition
- "Lam :: name \<Rightarrow> lam \<Rightarrow> lam"
-is
- "rLam"
-
-quotient_definition
- "fv :: lam \<Rightarrow> name set"
-is
- "rfv"
-
-(* definition of overloaded permutation function *)
-(* for the lifted type lam *)
-overloading
- perm_lam \<equiv> "perm :: 'x prm \<Rightarrow> lam \<Rightarrow> lam" (unchecked)
-begin
-
-quotient_definition
- "perm_lam :: 'x prm \<Rightarrow> lam \<Rightarrow> lam"
-is
- "perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam"
-
-end
-
-lemma perm_rsp[quot_respect]:
- "(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_respect]:
- "(op = ===> alpha) rVar rVar"
- by (auto intro: a1)
-
-lemma rApp_rsp[quot_respect]: "(alpha ===> alpha ===> alpha) rApp rApp"
- by (auto intro: a2)
-
-lemma rLam_rsp[quot_respect]: "(op = ===> alpha ===> alpha) rLam rLam"
- apply(auto)
- apply(rule a3)
- apply(rule_tac x="[]" in exI)
- unfolding fresh_star_def
- apply(simp add: fresh_list_nil)
- apply(simp add: alpha_rfv)
- done
-
-lemma rfv_rsp[quot_respect]:
- "(alpha ===> op =) rfv rfv"
-apply(simp add: alpha_rfv)
-done
-
-section {* lifted theorems *}
-
-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"
- by (lifting rlam.induct)
-
-ML {* show_all_types := true *}
-
-lemma perm_lam [simp]:
- fixes pi::"'a prm"
- shows "pi \<bullet> Var a = Var (pi \<bullet> a)"
- and "pi \<bullet> App t1 t2 = App (pi \<bullet> t1) (pi \<bullet> t2)"
- and "pi \<bullet> Lam a t = Lam (pi \<bullet> a) (pi \<bullet> t)"
-apply(lifting perm_rlam.simps)
-ML_prf {*
- List.last (map (symmetric o #def) (Quotient_Info.qconsts_dest @{context}));
- List.last (map (Thm.varifyT o symmetric o #def) (Quotient_Info.qconsts_dest @{context}))
-*}
-done
-
-instance lam::pt_name
-apply(default)
-apply(induct_tac [!] x rule: lam_induct)
-apply(simp_all add: pt_name2 pt_name3)
-done
-
-lemma fv_lam [simp]:
- shows "fv (Var a) = {a}"
- and "fv (App t1 t2) = fv t1 \<union> fv t2"
- and "fv (Lam a t) = fv t - {a}"
-apply(lifting rfv_var rfv_app rfv_lam)
-done
-
-
-lemma a1:
- "a = b \<Longrightarrow> Var a = Var b"
- by (lifting a1)
-
-lemma a2:
- "\<lbrakk>x = xa; xb = xc\<rbrakk> \<Longrightarrow> App x xb = App xa xc"
- by (lifting a2)
-
-lemma a3:
- "\<lbrakk>\<exists>pi::name prm. (fv t - {a} = fv s - {b} \<and> (fv t - {a})\<sharp>* pi \<and> (pi \<bullet> t) = s \<and> (pi \<bullet> a) = b)\<rbrakk>
- \<Longrightarrow> Lam a t = Lam b s"
- by (lifting a3)
-
-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>t a s b. \<lbrakk>a1 = Lam a t; a2 = Lam b s;
- \<exists>pi::name prm. fv t - {a} = fv s - {b} \<and> (fv t - {a}) \<sharp>* pi \<and> (pi \<bullet> t) = s \<and> pi \<bullet> a = b\<rbrakk> \<Longrightarrow> P\<rbrakk>
- \<Longrightarrow> P"
- by (lifting alpha.cases)
-
-lemma alpha_induct:
- "\<lbrakk>qx = qxa; \<And>a b. a = b \<Longrightarrow> qxb (Var a) (Var b);
- \<And>x xa xb xc. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc);
- \<And>t a s b.
- \<lbrakk>\<exists>pi::name prm. fv t - {a} = fv s - {b} \<and>
- (fv t - {a}) \<sharp>* pi \<and> ((pi \<bullet> t) = s \<and> qxb (pi \<bullet> t) s) \<and> pi \<bullet> a = b\<rbrakk> \<Longrightarrow> qxb (Lam a t) (Lam b s)\<rbrakk>
- \<Longrightarrow> qxb qx qxa"
- by (lifting alpha.induct)
-
-lemma lam_inject [simp]:
- shows "(Var a = Var b) = (a = b)"
- and "(App t1 t2 = App s1 s2) = (t1 = s1 \<and> t2 = s2)"
-apply(lifting rlam.inject(1) rlam.inject(2))
-apply(auto)
-apply(drule alpha.cases)
-apply(simp_all)
-apply(simp add: alpha.a1)
-apply(drule alpha.cases)
-apply(simp_all)
-apply(drule alpha.cases)
-apply(simp_all)
-apply(rule alpha.a2)
-apply(simp_all)
-done
-
-lemma rlam_distinct:
- shows "\<not>(rVar nam \<approx> rApp rlam1' rlam2')"
- and "\<not>(rApp rlam1' rlam2' \<approx> rVar nam)"
- and "\<not>(rVar nam \<approx> rLam nam' rlam')"
- and "\<not>(rLam nam' rlam' \<approx> rVar nam)"
- and "\<not>(rApp rlam1 rlam2 \<approx> rLam nam' rlam')"
- and "\<not>(rLam nam' rlam' \<approx> rApp rlam1 rlam2)"
-apply auto
-apply(erule alpha.cases)
-apply simp_all
-apply(erule alpha.cases)
-apply simp_all
-apply(erule alpha.cases)
-apply simp_all
-apply(erule alpha.cases)
-apply simp_all
-apply(erule alpha.cases)
-apply simp_all
-apply(erule alpha.cases)
-apply simp_all
-done
-
-lemma lam_distinct[simp]:
- shows "Var nam \<noteq> App lam1' lam2'"
- and "App lam1' lam2' \<noteq> Var nam"
- and "Var nam \<noteq> Lam nam' lam'"
- and "Lam nam' lam' \<noteq> Var nam"
- and "App lam1 lam2 \<noteq> Lam nam' lam'"
- and "Lam nam' lam' \<noteq> App lam1 lam2"
-apply(lifting rlam_distinct(1) rlam_distinct(2) rlam_distinct(3) rlam_distinct(4) rlam_distinct(5) rlam_distinct(6))
-done
-
-lemma var_supp1:
- shows "(supp (Var a)) = ((supp a)::name set)"
- by (simp add: supp_def)
-
-lemma var_supp:
- shows "(supp (Var a)) = {a::name}"
- using var_supp1 by (simp add: supp_atm)
-
-lemma app_supp:
- shows "supp (App t1 t2) = (supp t1) \<union> ((supp t2)::name set)"
-apply(simp only: perm_lam supp_def lam_inject)
-apply(simp add: Collect_imp_eq Collect_neg_eq)
-done
-
-lemma lam_supp:
- shows "supp (Lam x t) = ((supp ([x].t))::name set)"
-apply(simp add: supp_def)
-apply(simp add: abs_perm)
-sorry
-
-
-instance lam::fs_name
-apply(default)
-apply(induct_tac x rule: lam_induct)
-apply(simp add: var_supp)
-apply(simp add: app_supp)
-apply(simp add: lam_supp abs_supp)
-done
-
-lemma fresh_lam:
- "(a \<sharp> Lam b t) \<longleftrightarrow> (a = b) \<or> (a \<noteq> b \<and> a \<sharp> t)"
-apply(simp add: fresh_def)
-apply(simp add: lam_supp abs_supp)
-apply(auto)
-done
-
-lemma lam_induct_strong:
- fixes a::"'a::fs_name"
- assumes a1: "\<And>name b. P b (Var name)"
- and a2: "\<And>lam1 lam2 b. \<lbrakk>\<And>c. P c lam1; \<And>c. P c lam2\<rbrakk> \<Longrightarrow> P b (App lam1 lam2)"
- and a3: "\<And>name lam b. \<lbrakk>\<And>c. P c lam; name \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lam name lam)"
- shows "P a lam"
-proof -
- have "\<And>(pi::name prm) a. P a (pi \<bullet> lam)"
- proof (induct lam rule: lam_induct)
- case (1 name pi)
- show "P a (pi \<bullet> Var name)"
- apply (simp)
- apply (rule a1)
- done
- next
- case (2 lam1 lam2 pi)
- have b1: "\<And>(pi::name prm) a. P a (pi \<bullet> lam1)" by fact
- have b2: "\<And>(pi::name prm) a. P a (pi \<bullet> lam2)" by fact
- show "P a (pi \<bullet> App lam1 lam2)"
- apply (simp)
- apply (rule a2)
- apply (rule b1)
- apply (rule b2)
- done
- next
- case (3 name lam pi a)
- have b: "\<And>(pi::name prm) a. P a (pi \<bullet> lam)" by fact
- obtain c::name where fr: "c\<sharp>(a, pi\<bullet>name, pi\<bullet>lam)"
- apply(rule exists_fresh[of "(a, pi\<bullet>name, pi\<bullet>lam)"])
- apply(simp_all add: fs_name1)
- done
- from b fr have p: "P a (Lam c (([(c, pi\<bullet>name)]@pi)\<bullet>lam))"
- apply -
- apply(rule a3)
- apply(blast)
- apply(simp)
- done
- have eq: "[(c, pi\<bullet>name)] \<bullet> Lam (pi \<bullet> name) (pi \<bullet> lam) = Lam (pi \<bullet> name) (pi \<bullet> lam)"
- apply(rule perm_fresh_fresh)
- using fr
- apply(simp add: fresh_lam)
- apply(simp add: fresh_lam)
- done
- show "P a (pi \<bullet> Lam name lam)"
- apply (simp)
- apply(subst eq[symmetric])
- using p
- apply(simp only: perm_lam pt_name2 swap_simps)
- done
- qed
- then have "P a (([]::name prm) \<bullet> lam)" by blast
- then show "P a lam" by simp
-qed
-
-
-lemma var_fresh:
- fixes a::"name"
- shows "(a \<sharp> (Var b)) = (a \<sharp> b)"
- apply(simp add: fresh_def)
- apply(simp add: var_supp1)
- done
-
-(* lemma hom_reg: *)
-
-lemma rlam_rec_eqvt:
- fixes pi::"name prm"
- and f1::"name \<Rightarrow> ('a::pt_name)"
- shows "(pi\<bullet>rlam_rec f1 f2 f3 t) = rlam_rec (pi\<bullet>f1) (pi\<bullet>f2) (pi\<bullet>f3) (pi\<bullet>t)"
-apply(induct t)
-apply(simp_all)
-apply(simp add: perm_fun_def)
-apply(perm_simp)
-apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
-back
-apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
-apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
-apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
-apply(simp)
-apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
-back
-apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
-apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
-apply(simp)
-done
-
-
-lemma rlam_rec_respects:
- assumes f1: "f_var \<in> Respects (op= ===> op=)"
- and f2: "f_app \<in> Respects (alpha ===> alpha ===> op= ===> op= ===> op=)"
- and f3: "f_lam \<in> Respects (op= ===> alpha ===> op= ===> op=)"
- shows "rlam_rec f_var f_app f_lam \<in> Respects (alpha ===> op =)"
-apply(simp add: mem_def)
-apply(simp add: Respects_def)
-apply(rule allI)
-apply(rule allI)
-apply(rule impI)
-apply(erule alpha.induct)
-apply(simp)
-apply(simp)
-using f2
-apply(simp add: mem_def)
-apply(simp add: Respects_def)
-using f3[simplified mem_def Respects_def]
-apply(simp)
-apply(case_tac "a=b")
-apply(clarify)
-apply(simp)
-(* probably true *)
-sorry
-
-function
- term1_hom :: "(name \<Rightarrow> 'a) \<Rightarrow>
- (rlam \<Rightarrow> rlam \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow>
- ((name \<Rightarrow> rlam) \<Rightarrow> (name \<Rightarrow> 'a) \<Rightarrow> 'a) \<Rightarrow> rlam \<Rightarrow> 'a"
-where
- "term1_hom var app abs' (rVar x) = (var x)"
-| "term1_hom var app abs' (rApp t u) =
- app t u (term1_hom var app abs' t) (term1_hom var app abs' u)"
-| "term1_hom var app abs' (rLam x u) =
- abs' (\<lambda>y. [(x, y)] \<bullet> u) (\<lambda>y. term1_hom var app abs' ([(x, y)] \<bullet> u))"
-apply(pat_completeness)
-apply(auto)
-done
-
-lemma pi_size:
- fixes pi::"name prm"
- and t::"rlam"
- shows "size (pi \<bullet> t) = size t"
-apply(induct t)
-apply(auto)
-done
-
-termination term1_hom
- apply(relation "measure (\<lambda>(f1, f2, f3, t). size t)")
-apply(auto simp add: pi_size)
-done
-
-lemma lam_exhaust:
- "\<lbrakk>\<And>name. y = Var name \<Longrightarrow> P; \<And>rlam1 rlam2. y = App rlam1 rlam2 \<Longrightarrow> P; \<And>name rlam. y = Lam name rlam \<Longrightarrow> P\<rbrakk>
- \<Longrightarrow> P"
-apply(lifting rlam.exhaust)
-done
-
-(* THIS IS NOT TRUE, but it lets prove the existence of the hom function *)
-lemma lam_inject':
- "(Lam a x = Lam b y) = ((\<lambda>c. [(a, c)] \<bullet> x) = (\<lambda>c. [(b, c)] \<bullet> y))"
-sorry
-
-function
- hom :: "(name \<Rightarrow> 'a) \<Rightarrow>
- (lam \<Rightarrow> lam \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow>
- ((name \<Rightarrow> lam) \<Rightarrow> (name \<Rightarrow> 'a) \<Rightarrow> 'a) \<Rightarrow> lam \<Rightarrow> 'a"
-where
- "hom f_var f_app f_lam (Var x) = f_var x"
-| "hom f_var f_app f_lam (App l r) = f_app l r (hom f_var f_app f_lam l) (hom f_var f_app f_lam r)"
-| "hom f_var f_app f_lam (Lam a x) = f_lam (\<lambda>b. ([(a,b)] \<bullet> x)) (\<lambda>b. hom f_var f_app f_lam ([(a,b)] \<bullet> x))"
-defer
-apply(simp_all add: lam_inject') (* inject, distinct *)
-apply(tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
-apply(rule refl)
-apply(rule ext)
-apply(tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
-apply simp_all
-apply(erule conjE)+
-apply(rule_tac x="b" in cong)
-apply simp_all
-apply auto
-apply(rule_tac y="b" in lam_exhaust)
-apply simp_all
-apply auto
-apply meson
-apply(simp_all add: lam_inject')
-apply metis
-done
-
-termination hom
- apply -
-(*
-ML_prf {* Size.size_thms @{theory} "LamEx.lam" *}
-*)
-sorry
-
-thm hom.simps
-
-lemma term1_hom_rsp:
- "\<lbrakk>(alpha ===> alpha ===> op =) f_app f_app; ((op = ===> alpha) ===> op =) f_lam f_lam\<rbrakk>
- \<Longrightarrow> (alpha ===> op =) (term1_hom f_var f_app f_lam) (term1_hom f_var f_app f_lam)"
-apply(simp)
-apply(rule allI)+
-apply(rule impI)
-apply(erule alpha.induct)
-apply(auto)[1]
-apply(auto)[1]
-apply(simp)
-apply(erule conjE)+
-apply(erule exE)+
-apply(erule conjE)+
-apply(clarify)
-sorry
-
-lemma hom: "
-\<forall>f_var. \<forall>f_app \<in> Respects(alpha ===> alpha ===> op =).
-\<forall>f_lam \<in> Respects((op = ===> alpha) ===> op =).
-\<exists>hom\<in>Respects (alpha ===> op =).
- ((\<forall>x. hom (rVar x) = f_var x) \<and>
- (\<forall>l r. hom (rApp l r) = f_app l r (hom l) (hom r)) \<and>
- (\<forall>x a. hom (rLam a x) = f_lam (\<lambda>b. ([(a,b)]\<bullet> x)) (\<lambda>b. hom ([(a,b)] \<bullet> x))))"
-apply(rule allI)
-apply(rule ballI)+
-apply(rule_tac x="term1_hom f_var f_app f_lam" in bexI)
-apply(simp_all)
-apply(simp only: in_respects)
-apply(rule term1_hom_rsp)
-apply(assumption)+
-done
-
-lemma hom':
-"\<exists>hom.
- ((\<forall>x. hom (Var x) = f_var x) \<and>
- (\<forall>l r. hom (App l r) = f_app l r (hom l) (hom r)) \<and>
- (\<forall>x a. hom (Lam a x) = f_lam (\<lambda>b. ([(a,b)] \<bullet> x)) (\<lambda>b. hom ([(a,b)] \<bullet> x))))"
-apply (lifting hom)
-done
-
-(* test test
-lemma raw_hom_correct:
- assumes f1: "f_var \<in> Respects (op= ===> op=)"
- and f2: "f_app \<in> Respects (alpha ===> alpha ===> op= ===> op= ===> op=)"
- and f3: "f_lam \<in> Respects ((op= ===> alpha) ===> (op= ===> op=) ===> op=)"
- shows "\<exists>!hom\<in>Respects (alpha ===> op =).
- ((\<forall>x. hom (rVar x) = f_var x) \<and>
- (\<forall>l r. hom (rApp l r) = f_app l r (hom l) (hom r)) \<and>
- (\<forall>x a. hom (rLam a x) = f_lam (\<lambda>b. ([(a,b)]\<bullet> x)) (\<lambda>b. hom ([(a,b)] \<bullet> x))))"
-unfolding Bex1_def
-apply(rule ex1I)
-sorry
-*)
-
-
-end
-