# HG changeset patch # User Cezary Kaliszyk # Date 1270104513 -7200 # Node ID 00680cea0dde083000ce604879d23861fcba66a0 # Parent 925a5e9aa8323723dac83528b266396042e33313 Let with multiple bindings. diff -r 925a5e9aa832 -r 00680cea0dde Nominal/Abs.thy --- a/Nominal/Abs.thy Thu Apr 01 08:06:01 2010 +0200 +++ b/Nominal/Abs.thy Thu Apr 01 08:48:33 2010 +0200 @@ -444,6 +444,12 @@ \ pi \ bsl = csl)" by (simp_all add: alphas) +lemma alphas3: + "(bsl, x1, x2, x3) \lst (\(x1, y1, z1) (x2, y2, z2). R1 x1 x2 \ R2 y1 y2 \ R3 z1 z2) (\(a, b, c). f1 a \ (f2 b \ f3 c)) pi (csl, y1, y2, y3) = (f1 x1 \ (f2 x2 \ f3 x3) - set bsl = f1 y1 \ (f2 y2 \ f3 y3) - set csl \ + (f1 x1 \ (f2 x2 \ f3 x3) - set bsl) \* pi \ + R1 (pi \ x1) y1 \ R2 (pi \ x2) y2 \ R3 (pi \ x3) y3 \ pi \ bsl = csl)" +by (simp add: alphas) + lemma alpha_gen_compose_sym: fixes pi assumes b: "(aa, t) \gen (\x1 x2. R x1 x2 \ R x2 x1) f pi (ab, s)" diff -r 925a5e9aa832 -r 00680cea0dde Nominal/ExLetMult.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/ExLetMult.thy Thu Apr 01 08:48:33 2010 +0200 @@ -0,0 +1,31 @@ +theory ExLetMult +imports "Parser" +begin + +atom_decl name + +ML {* val _ = recursive := true *} +ML {* val _ = alpha_type := AlphaLst *} +ML {* val _ = cheat_equivp := true *} +nominal_datatype trm = + Vr "name" +| Ap "trm" "trm" +| Lm "trm" "trm" +| Lt l::"lts" t::"trm" s::"trm" bind "bn l" in t, bind "bn l" in s +and lts = + Lnil +| Lcons "name" "trm" "lts" +binder + bn +where + "bn Lnil = []" +| "bn (Lcons x t l) = (atom x) # (bn l)" + +thm trm_lts.eq_iff +thm trm_lts.induct +thm trm_lts.fv[simplified trm_lts.supp] + +end + + + diff -r 925a5e9aa832 -r 00680cea0dde Nominal/Fv.thy --- a/Nominal/Fv.thy Thu Apr 01 08:06:01 2010 +0200 +++ b/Nominal/Fv.thy Thu Apr 01 08:48:33 2010 +0200 @@ -931,7 +931,7 @@ simp_tac (HOL_basic_ss addsimps @{thms supp_def}) THEN_ALL_NEW simp_tac (HOL_basic_ss addsimps (@{thms permute_abs} @ perm)) THEN_ALL_NEW simp_tac (HOL_basic_ss addsimps (@{thms Abs_eq_iff} @ eqiff)) THEN_ALL_NEW - simp_tac (HOL_basic_ss addsimps @{thms alphas2}) THEN_ALL_NEW + simp_tac (HOL_basic_ss addsimps @{thms alphas3 alphas2}) THEN_ALL_NEW simp_tac (HOL_basic_ss addsimps @{thms alphas}) THEN_ALL_NEW asm_full_simp_tac (HOL_basic_ss addsimps (@{thm supp_Pair} :: sym_eqvts ctxt)) THEN_ALL_NEW asm_full_simp_tac (HOL_basic_ss addsimps (@{thm Pair_eq} :: all_eqvts ctxt)) THEN_ALL_NEW diff -r 925a5e9aa832 -r 00680cea0dde Nominal/Parser.thy --- a/Nominal/Parser.thy Thu Apr 01 08:06:01 2010 +0200 +++ b/Nominal/Parser.thy Thu Apr 01 08:48:33 2010 +0200 @@ -456,9 +456,12 @@ val q_bn = map (lift_thm qtys lthy16) raw_bn_eqs; val lthy17 = note_simp_suffix "bn" q_bn lthy16; val _ = tracing "Lifting eq-iff"; - val eq_iff_unfolded1 = map (Local_Defs.unfold lthy17 @{thms alphas2}) alpha_eq_iff - val eq_iff_unfolded2 = map (Local_Defs.unfold lthy17 @{thms alphas}) eq_iff_unfolded1 - val q_eq_iff_pre1 = map (lift_thm qtys lthy17) eq_iff_unfolded2; + val _ = map tracing (map PolyML.makestring alpha_eq_iff); + val eq_iff_unfolded0 = map (Local_Defs.unfold lthy17 @{thms alphas3}) alpha_eq_iff + val eq_iff_unfolded1 = map (Local_Defs.unfold lthy17 @{thms alphas2}) eq_iff_unfolded0 + val eq_iff_unfolded2 = map (Local_Defs.unfold lthy17 @{thms alphas} ) eq_iff_unfolded1 + val q_eq_iff_pre0 = map (lift_thm qtys lthy17) eq_iff_unfolded2; + val q_eq_iff_pre1 = map (Local_Defs.fold lthy17 @{thms alphas3}) q_eq_iff_pre0 val q_eq_iff_pre2 = map (Local_Defs.fold lthy17 @{thms alphas2}) q_eq_iff_pre1 val q_eq_iff = map (Local_Defs.fold lthy17 @{thms alphas}) q_eq_iff_pre2 val (_, lthy18) = Local_Theory.note ((suffix_bind "eq_iff", []), q_eq_iff) lthy17; diff -r 925a5e9aa832 -r 00680cea0dde Nominal/Rsp.thy --- a/Nominal/Rsp.thy Thu Apr 01 08:06:01 2010 +0200 +++ b/Nominal/Rsp.thy Thu Apr 01 08:48:33 2010 +0200 @@ -83,7 +83,7 @@ REPEAT o rtac @{thm exI[of _ "0 :: perm"]} THEN_ALL_NEW simp_tac (HOL_basic_ss addsimps @{thms alphas2}) THEN_ALL_NEW asm_full_simp_tac (HOL_ss addsimps (rsp @ - @{thms alphas fresh_star_def fresh_zero_perm permute_zero ball_triv add_0_left})) + @{thms split_conv alphas fresh_star_def fresh_zero_perm permute_zero ball_triv add_0_left})) )) *}