--- a/Nominal/Manual/Term4.thy Thu Apr 15 14:08:08 2010 +0200
+++ b/Nominal/Manual/Term4.thy Thu Apr 15 15:31:36 2010 +0200
@@ -29,7 +29,6 @@
thm permute_rtrm4_permute_rtrm4_list.simps
thm permute_rtrm4_permute_rtrm4_list.simps[simplified repaired]
-ML define_fv_alpha_export
local_setup {* snd o define_fv_alpha_export (Datatype.the_info @{theory} "Term4.rtrm4")
[[[], [], [(NONE, 0, 1, AlphaGen)]], [[], []] ] [] *}
print_theorems
@@ -79,14 +78,6 @@
(build_equivps [@{term alpha_rtrm4}, @{term alpha_rtrm4_list}] @{thms alpha4_reflp} @{thm alpha_rtrm4_alpha_rtrm4_list.induct} @{thms rtrm4.inject list.inject} @{thms alpha4_inj_no} @{thms rtrm4.distinct list.distinct} @{thms alpha_rtrm4_list.cases alpha_rtrm4.cases} @{thms alpha4_eqvt_no} ctxt)) ctxt)) *}
lemmas alpha4_equivp = alpha4_equivp_no[simplified fix2]
-(*lemma fv_rtrm4_rsp:
- "xa \<approx>4 ya \<Longrightarrow> fv_rtrm4 xa = fv_rtrm4 ya"
- "x \<approx>4l y \<Longrightarrow> fv_rtrm4_list x = fv_rtrm4_list y"
- apply (induct rule: alpha_rtrm4_alpha_rtrm4_list.inducts)
- apply (simp_all add: alpha_gen)
-done*)
-
-
quotient_type
trm4 = rtrm4 / alpha_rtrm4
(*and
@@ -96,30 +87,83 @@
local_setup {*
(fn ctxt => ctxt
|> snd o (Quotient_Def.quotient_lift_const [] ("Vr4", @{term rVr4}))
- |> snd o (Quotient_Def.quotient_lift_const [] ("Ap4", @{term rAp4}))
- |> snd o (Quotient_Def.quotient_lift_const [] ("Lm4", @{term rLm4})))
+ |> snd o (Quotient_Def.quotient_lift_const [@{typ "trm4"}] ("Ap4", @{term rAp4}))
+ |> snd o (Quotient_Def.quotient_lift_const [] ("Lm4", @{term rLm4}))
+ |> snd o (Quotient_Def.quotient_lift_const [] ("fv_trm4", @{term fv_rtrm4})))
*}
print_theorems
-local_setup {* snd o prove_const_rsp @{binding fv_rtrm4_rsp} [@{term fv_rtrm4}]
- (fn _ => fvbv_rsp_tac @{thm alpha_rtrm4_alpha_rtrm4_list.inducts(1)} @{thms fv_rtrm4_fv_rtrm4_list.simps} 1) *}
+
+
+lemma fv_rtrm4_rsp:
+ "xa \<approx>4 ya \<Longrightarrow> fv_rtrm4 xa = fv_rtrm4 ya"
+ "x \<approx>4l y \<Longrightarrow> fv_rtrm4_list x = fv_rtrm4_list y"
+ apply (induct rule: alpha_rtrm4_alpha_rtrm4_list.inducts)
+ apply (simp_all add: alpha_gen)
+done
+
+local_setup {* snd o prove_const_rsp [] @{binding fv_rtrm4_rsp'} [@{term fv_rtrm4}]
+ (fn _ => asm_full_simp_tac (@{simpset} addsimps @{thms fv_rtrm4_rsp}) 1) *}
+print_theorems
+
+ML constr_rsp_tac
+
+local_setup {* snd o prove_const_rsp [] @{binding rVr4_rsp} [@{term rVr4}]
+ (fn _ => constr_rsp_tac @{thms alpha4_inj} @{thms fv_rtrm4_rsp alpha4_equivp} 1) *}
+local_setup {* snd o prove_const_rsp [] @{binding rLm4_rsp} [@{term rLm4}]
+ (fn _ => constr_rsp_tac @{thms alpha4_inj} @{thms fv_rtrm4_rsp alpha4_equivp} 1) *}
+
+lemma [quot_respect]:
+ "(alpha_rtrm4 ===> list_rel alpha_rtrm4 ===> alpha_rtrm4) rAp4 rAp4"
+ by (simp add: alpha4_inj)
+
+(* Maybe also need: @{term "permute :: perm \<Rightarrow> rtrm4 list \<Rightarrow> rtrm4 list"} *)
+local_setup {* snd o prove_const_rsp [] @{binding permute_rtrm4_rsp}
+ [@{term "permute :: perm \<Rightarrow> rtrm4 \<Rightarrow> rtrm4"}]
+ (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha4_eqvt}) 1) *}
print_theorems
-local_setup {* snd o prove_const_rsp @{binding rVr4_rsp} [@{term rVr4}]
- (fn _ => constr_rsp_tac @{thms alpha4_inj} @{thms fv_rtrm4_rsp} @{thms alpha4_equivp} 1) *}
-lemma "(alpha_rtrm4 ===> list_rel alpha_rtrm4 ===> alpha_rtrm4) rAp4 rAp4"
-apply simp
-apply clarify
-apply (simp add: alpha4_inj)
+lemma list_rel_rsp:
+ "\<lbrakk>\<forall>x y. R x y \<longrightarrow> (\<forall>a b. R a b \<longrightarrow> S x a = T y b); list_rel R x y; list_rel R a b\<rbrakk>
+ \<Longrightarrow> list_rel S x a = list_rel T y b"
+ sorry
+
+lemma[quot_respect]:
+ "((R ===> R ===> op =) ===> list_rel R ===> list_rel R ===> op =) list_rel list_rel"
+ by (simp add: list_rel_rsp)
+lemma[quot_preserve]:
+ assumes a: "Quotient R abs1 rep1"
+ shows "((abs1 ---> abs1 ---> id) ---> map rep1 ---> map rep1 ---> id) list_rel = list_rel"
+ apply (simp add: expand_fun_eq)
+ apply clarify
+ apply (induct_tac xa xb rule: list_induct2')
+ apply (simp_all add: Quotient_abs_rep[OF a])
+ done
-local_setup {* snd o prove_const_rsp @{binding rLm4_rsp} [@{term rLm4}]
- (fn _ => constr_rsp_tac @{thms alpha4_inj} @{thms fv_rtrm4_rsp} @{thms alpha4_equivp} 1) *}
-local_setup {* snd o prove_const_rsp @{binding permute_rtrm4_rsp}
- [@{term "permute :: perm \<Rightarrow> rtrm4 \<Rightarrow> rtrm4"}, @{term "permute :: perm \<Rightarrow> rtrm4 list \<Rightarrow> rtrm4 list"}]
- (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha4_eqvt}) 1) *}
+lemma[quot_preserve]:
+ assumes a: "Quotient R abs1 rep1"
+ shows "(list_rel ((rep1 ---> rep1 ---> id) R) l m) = (l = m)"
+ by (induct l m rule: list_induct2') (simp_all add: Quotient_rel_rep[OF a])
+
+lemma bla: "(Ap4 trm4 list = Ap4 trm4a lista) =
+ (trm4 = trm4a \<and> list_rel (op =) list lista)"
+ by (lifting alpha4_inj(2))
+
+thm bla[simplified list_rel_eq]
-thm rtrm4.induct
-lemmas trm1_bp_induct = rtrm4.induct[quot_lifted]
+lemma " (Lm4 name rtrm4 = Lm4 namea rtrm4a) =
+ (\<exists>pi\<Colon>perm.
+ fv_trm4 rtrm4 - {atom name} = fv_trm4 rtrm4a - {atom namea} \<and>
+ (fv_trm4 rtrm4 - {atom name}) \<sharp>* pi \<and>
+ pi \<bullet> rtrm4 = rtrm4a \<and> pi \<bullet> {atom name} = {atom namea})"
+
+ML {* lift_thm [@{typ trm4}] @{context} @{thm alpha4_inj(1)} *}
+ML {* lift_thm [@{typ trm4}] @{context} @{thm alpha4_inj(2)} *}
+ML {* lift_thm [@{typ trm4}] @{context} @{thm alpha4_inj(3)[unfolded alpha_gen]} *}
+ML {* lift_thm [@{typ trm4}] @{context} @{thm rtrm4.induct} *}
+.
+
+(*lemmas trm1_bp_induct = rtrm4.induct[quot_lifted]*)
end