diff -r b320da14e63c -r 8d2667ebe26c Nominal/Term1.thy --- a/Nominal/Term1.thy Fri Mar 05 10:23:40 2010 +0100 +++ b/Nominal/Term1.thy Fri Mar 05 13:04:47 2010 +0100 @@ -142,12 +142,16 @@ apply default apply (rule rtrm1_bp_fs)+ done - -lemma fv_eq_bv: "fv_bp bp = bv1 bp" +lemma fv_eq_bv_pre: "fv_bp bp = bv1 bp" apply(induct bp rule: trm1_bp_inducts(2)) apply(simp_all) done +lemma fv_eq_bv: "fv_bp = bv1" +apply(rule ext) +apply(rule fv_eq_bv_pre) +done + lemma helper2: "{b. \pi. pi \ (a \ b) \ bp \ bp} = {}" apply auto apply (rule_tac x="(x \ a)" in exI) @@ -165,6 +169,35 @@ apply (rule alpha_bp_eq_eq) done +lemma ex_out: + "(\x. Z x \ Q) = (Q \ (\x. Z x))" + "(\x. Q \ Z x) = (Q \ (\x. Z x))" + "(\x. P x \ Q \ Z x) = (Q \ (\x. P x \ Z x))" + "(\x. Q \ P x \ Z x) = (Q \ (\x. P x \ Z x))" + "(\x. Q \ P x \ Z x \ W x) = (Q \ (\x. P x \ Z x \ W x))" +apply (blast)+ +done + +lemma "(Abs bs (x, x') = Abs cs (y, y')) = (\p. (bs, x) \gen op = supp p (cs, y) \ (bs, x') \gen op = supp p (cs, y'))" +thm Abs_eq_iff +apply (simp add: Abs_eq_iff) +apply (rule arg_cong[of _ _ "Ex"]) +apply (rule ext) +apply (simp only: alpha_gen) +apply (simp only: supp_Pair eqvts) +apply rule +apply (erule conjE)+ +oops + +lemma "(f (p \ bp), p \ bp) \gen op = f pi (f bp, bp) = False" +apply (simp add: alpha_gen fresh_star_def) +oops + +(* TODO: permute_ABS should be in eqvt? *) + +lemma Collect_neg_conj: "{x. \(P x \ Q x)} = {x. \(P x)} \ {x. \(Q x)}" +by (simp add: Collect_imp_eq Collect_neg_eq[symmetric]) + lemma supp_fv: "supp t = fv_trm1 t" "supp b = fv_bp b" @@ -188,21 +221,23 @@ apply(simp add: supp_at_base) apply(simp (no_asm) add: supp_def Collect_imp_eq Collect_neg_eq) apply(simp add: Collect_imp_eq[symmetric] Collect_neg_eq[symmetric] supp_def[symmetric]) -(*apply(subgoal_tac "supp (Lt1 bp rtrm11 rtrm12) = supp (Abs (bv1 bp) (rtrm12)) \ supp(rtrm11)*) -apply(subgoal_tac "supp (Lt1 bp rtrm11 rtrm12) = supp (Abs (bv1 bp) (bp, rtrm12)) \ supp(rtrm11)") +apply(subgoal_tac "supp (Lt1 bp rtrm11 rtrm12) = supp (Abs (bv1 bp) (rtrm12)) \ supp(rtrm11)") +(*apply(subgoal_tac "supp (Lt1 bp rtrm11 rtrm12) = supp (Abs (bv1 bp) (bp, rtrm12)) \ supp(rtrm11)")*) apply(simp add: supp_Abs fv_trm1 supp_Pair Un_Diff Un_assoc fv_eq_bv) apply(blast) (* Un_commute in a good place *) -apply(simp (no_asm) only: supp_def permute_trm1) -apply(simp only: alpha1_INJ) +apply(simp (no_asm) only: supp_def permute_set_eq atom_eqvt permute_trm1) +apply(simp only: alpha1_INJ permute_ABS permute_prod.simps Abs_eq_iff) +apply(simp only: ex_out) apply(simp only: Un_commute) +apply(simp only: alpha_bp_eq fv_eq_bv) +apply(simp only: alpha_gen fv_eq_bv supp_Pair) +apply(simp only: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric] bv1_eqvt fv_eq_bv) +apply(simp only: ex_out) +apply(simp only: Collect_neg_conj finite_Un Diff_cancel) +apply(simp) apply(simp add: Collect_imp_eq) -apply(simp add: Collect_neg_eq[symmetric]) -apply(simp only: Abs_eq_iff) -apply(simp add: alpha_bp_eq) -apply(simp add: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric] bv1_eqvt fv_eq_bv) -apply(simp add: alpha_gen fv_eq_bv supp_Pair) -apply(simp add: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric] bv1_eqvt fv_eq_bv) -apply(simp add: Collect_imp_eq Collect_neg_eq Collect_disj_eq fresh_star_def helper2) +apply(simp add: Collect_neg_eq[symmetric] fresh_star_def) +apply(fold supp_def) sorry lemma trm1_supp: