# HG changeset patch # User Cezary Kaliszyk # Date 1268845144 -3600 # Node ID 44e68ab6841eaa5098eb728bee2e5cc8760d1376 # Parent b55b78e6391382e3ad32e93a1b2fc9b12d1b8823 Updated Term1, including statement of strong induction. diff -r b55b78e63913 -r 44e68ab6841e Nominal/Term1.thy --- a/Nominal/Term1.thy Wed Mar 17 17:40:14 2010 +0100 +++ b/Nominal/Term1.thy Wed Mar 17 17:59:04 2010 +0100 @@ -12,9 +12,9 @@ | rLm1 "name" "rtrm1" --"name is bound in trm1" | rLt1 "bp" "rtrm1" "rtrm1" --"all variables in bp are bound in the 2nd trm1" and bp = -(* BUnit*) - BVr "name" -(*| BPr "bp" "bp"*) + BUnit +| BVr "name" +| BPr "bp" "bp" print_theorems @@ -23,9 +23,9 @@ primrec bv1 where -(* "bv1 (BUnit) = {}"*) - "bv1 (BVr x) = {atom x}" -(*| "bv1 (BPr bp1 bp2) = (bv1 bp1) \ (bv1 bp2)"*) + "bv1 (BUnit) = {}" +| "bv1 (BVr x) = {atom x}" +| "bv1 (BPr bp1 bp2) = (bv1 bp1) \ (bv1 bp2)" setup {* snd o define_raw_perms (Datatype.the_info @{theory} "Term1.rtrm1") 2 *} thm permute_rtrm1_permute_bp.simps @@ -33,7 +33,7 @@ local_setup {* snd o define_fv_alpha (Datatype.the_info @{theory} "Term1.rtrm1") [[[], [], [(NONE, 0, 1)], [(SOME (@{term bv1}, true), 0, 2)]], - [[](*, [], []*)]] [(@{term bv1}, 1, [(*[],*) [0](*, [0, 1]*)])] *} + [[], [], []]] [(@{term bv1}, 1, [[], [0], [0, 1]])] *} notation alpha_rtrm1 ("_ \1 _" [100, 100] 100) and @@ -52,8 +52,7 @@ snd o build_eqvts @{binding fv_rtrm1_fv_bp_eqvt} [@{term fv_rtrm1}, @{term fv_bp}] (build_eqvts_tac @{thm rtrm1_bp.induct} @{thms fv_rtrm1_fv_bp.simps permute_rtrm1_permute_bp.simps} @{context}) *} -(* -local_setup {* +(*local_setup {* snd o build_eqvts @{binding fv_rtrm1_fv_bv1_eqvt} [@{term fv_rtrm1}, @{term fv_bv1}] (build_eqvts_tac @{thm rtrm1_bp.induct} @{thms fv_rtrm1_fv_bv1.simps permute_rtrm1_permute_bp.simps} @{context}) *} print_theorems @@ -64,16 +63,9 @@ print_theorems *) -lemma alpha1_eqvt: - "(rtrm1 \1 rtrm1a \ (p \ rtrm1) \1 (p \ rtrm1a)) \ - (bp \1b bpa \ (p \ bp) \1b (p \ bpa)) \ - (alpha_bv1 a b c \ alpha_bv1 (p \ a) (p \ b) (p \ c))" -by (tactic {* alpha_eqvt_tac @{thm alpha_rtrm1_alpha_bp_alpha_bv1.induct} @{thms alpha1_inj permute_rtrm1_permute_bp.simps} @{context} 1 *}) - -(* local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha1_eqvt}, []), -build_alpha_eqvts [@{term alpha_rtrm1}, @{term alpha_bp}] [@{term "permute :: perm \ rtrm1 \ rtrm1"},@{term "permute :: perm \ bp \ bp"}] @{thms permute_rtrm1_permute_bp.simps alpha1_inj} @{thm alpha_rtrm1_alpha_bp_alpha_bv1.induct} ctxt) ctxt)) *}*) +build_alpha_eqvts [@{term alpha_rtrm1}, @{term alpha_bp}, @{term alpha_bv1}] (fn _ => alpha_eqvt_tac @{thm alpha_rtrm1_alpha_bp_alpha_bv1.induct} @{thms permute_rtrm1_permute_bp.simps alpha1_inj} ctxt 1) ctxt) ctxt)) *} lemma alpha1_eqvt_proper[eqvt]: "pi \ (t \1 s) = ((pi \ t) \1 (pi \ s))" @@ -92,6 +84,10 @@ done thm eqvts_raw(1) +lemma "(b \1 a \ a \1 b) \ (x \1b y \ y \1b x) \ (alpha_bv1 x y \ alpha_bv1 y x)" +apply (tactic {* symp_tac @{thm alpha_rtrm1_alpha_bp_alpha_bv1.induct} @{thms alpha1_inj} @{thms alpha1_eqvt} @{context} 1 *}) +done + lemma alpha1_equivp: "equivp alpha_rtrm1" "equivp alpha_bp" @@ -99,9 +95,8 @@ (* local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha1_equivp}, []), - (build_equivps [@{term alpha_rtrm1}, @{term alpha_bp}] @{thm rtrm1_bp.induct} @{thm alpha_rtrm1_alpha_bp_alpha_bv1.induct} @{thms rtrm1.inject bp.inject} @{thms alpha1_inj} @{thms rtrm1.distinct bp.distinct} @{thms alpha_rtrm1.cases alpha_bp.cases} @{thms alpha1_eqvt} ctxt)) ctxt)) *} -thm alpha1_equivp -*) + (build_equivps [@{term alpha_rtrm1}, @{term alpha_bp}, @{term alpha_bv1}] @{thm rtrm1_bp.induct} @{thm alpha_rtrm1_alpha_bp_alpha_bv1.induct} @{thms rtrm1.inject bp.inject} @{thms alpha1_inj} @{thms rtrm1.distinct bp.distinct} @{thms alpha_rtrm1.cases alpha_bp.cases} @{thms alpha1_eqvt} ctxt)) ctxt)) *} +thm alpha1_equivp*) local_setup {* define_quotient_type [(([], @{binding trm1}, NoSyn), (@{typ rtrm1}, @{term alpha_rtrm1}))] (rtac @{thm alpha1_equivp(1)} 1) *} @@ -138,17 +133,17 @@ lemmas permute_trm1 = permute_rtrm1_permute_bp.simps[quot_lifted] and fv_trm1 = fv_rtrm1_fv_bp.simps[quot_lifted] -and fv_trm1_eqvt = fv_rtrm1_fv_bp_eqvt(1)[quot_lifted] -and alpha1_INJ = alpha1_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen] +and fv_trm1_eqvt = fv_rtrm1_fv_bp_eqvt[quot_lifted] +and alpha1_INJ = alpha1_inj[unfolded alpha_gen2, unfolded alpha_gen, quot_lifted, folded alpha_gen2, folded alpha_gen] lemma supports: "(supp (atom x)) supports (Vr1 x)" "(supp t \ supp s) supports (Ap1 t s)" "(supp (atom x) \ supp t) supports (Lm1 x t)" "(supp b \ supp t \ supp s) supports (Lt1 b t s)" -(* "{} supports BUnit"*) + "{} supports BUnit" "(supp (atom x)) supports (BVr x)" -(* "(supp a \ supp b) supports (BPr a b)"*) + "(supp a \ supp b) supports (BPr a b)" apply(tactic {* ALLGOALS (supports_tac @{thms permute_trm1}) *}) done @@ -197,84 +192,32 @@ 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 " -{a\atom. infinite ({b\atom. \ (\pi\perm. P pi a b \ Q pi a b)})} = -{a\atom. infinite {b\atom. \ (\p\perm. P p a b)}} \ -{a\atom. infinite {b\atom. \ (\p\perm. Q p a b)}}" -oops - lemma inf_or: "(infinite x \ infinite y) = infinite (x \ y)" by (simp add: finite_Un) - - lemma supp_fv_let: assumes sa : "fv_bp bp = supp bp" - shows "\fv_trm1 ta = supp ta; fv_trm1 tb = supp tb\ - \ supp (Lt1 bp ta tb) = fv_trm1 (Lt1 bp ta tb)" -apply(simp only: fv_trm1 fv_eq_bv sa[simplified fv_eq_bv]) -apply simp + shows "\fv_trm1 ta = supp ta; fv_trm1 tb = supp tb; fv_bp bp = supp bp\ + \ supp (Lt1 bp ta tb) = supp ta \ (supp (bp, tb) - supp bp)" apply(fold supp_Abs) -apply(simp only: fv_trm1 fv_eq_bv sa[simplified fv_eq_bv,symmetric]) +apply(simp (no_asm) only: fv_trm1 fv_eq_bv sa[simplified fv_eq_bv,symmetric]) apply(simp (no_asm) only: supp_def) apply(simp only: permute_set_eq permute_trm1) -apply(simp only: alpha1_INJ) +apply(simp only: alpha1_INJ alpha_bp_eq) apply(simp only: ex_out) apply(simp only: Collect_neg_conj) apply(simp only: permute_ABS) apply(simp only: Abs_eq_iff) -apply(simp only: alpha_gen fv_eq_bv supp_Pair) +apply(simp only: alpha_gen supp_Pair split_conv eqvts) apply(simp only: inf_or[symmetric]) apply(simp only: Collect_disj_eq) apply(tactic {* Cong_Tac.cong_tac @{thm cong} 1 *}) apply(rule refl) -apply(simp only: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric] bv1_eqvt fv_eq_bv sa[simplified fv_eq_bv,symmetric]) -apply(induct bp) -apply(simp_all only: TrueI) -apply(simp_all only: permute_trm1) -apply(simp_all only: bv1.simps) -apply(simp_all only: alpha1_INJ) (* Doesn't look true... *) -apply(simp) -sorry - -lemma -"(\ (\p. (a \ b) \ supp tb - {atom ((a \ b) \ name)} = supp tb - {atom name} \ - ({atom (p \ (a \ b) \ name)} = {atom name}) \ - ((a \ b) \ supp tb - {atom ((a \ b) \ name)}) \* p \ - p \ (a \ b) \ tb = tb)) = - (\ (\p. (a \ b) \ supp tb - {atom ((a \ b) \ name)} = supp tb - {atom name} \ - ((a \ b) \ supp tb - {atom ((a \ b) \ name)}) \* p \ - p \ (a \ b) \ tb = tb))" -apply simp -apply rule -prefer 2 -apply (meson)[2] -apply clarify -apply (erule_tac x="p" in allE) -apply simp -apply (simp add: atom_eqvt[symmetric]) -sorry - -thm trm1_bp_inducts +apply (simp only: eqvts[symmetric] fv_trm1_eqvt[symmetric]) +apply (simp only: eqvts Pair_eq) +done lemma supp_fv: "supp t = fv_trm1 t" @@ -292,28 +235,16 @@ apply(simp add: Abs_eq_iff) apply(simp add: alpha_gen.simps) apply(simp add: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric]) -defer -apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt) +apply(simp add: supp_fv_let fv_trm1 fv_eq_bv supp_Pair) +apply blast +apply(simp add: supp_def permute_trm1 alpha1_INJ fv_trm1) +apply(simp add: supp_def permute_trm1 alpha1_INJ fv_trm1) apply(simp only: supp_at_base[simplified supp_def]) -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(rule supp_fv_let) apply(simp_all)*) -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_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 (no_asm) only: supp_def permute_set_eq atom_eqvt permute_trm1 alpha1_INJ[simplified alpha_bp_eq]) +apply(simp add: Collect_imp_eq Collect_neg_eq[symmetric]) apply(fold supp_def) -sorry +apply simp +done lemma trm1_supp: "supp (Vr1 x) = {atom x}"