# HG changeset patch # User Christian Urban # Date 1269500703 -3600 # Node ID 36798cdbc452397bcf7f522226552012cf1e78a3 # Parent a5501c9fad9b778ded0d78fe229fd99af5d837a4 first attempt of strong induction for lets with assignments diff -r a5501c9fad9b -r 36798cdbc452 Nominal/ExLet.thy --- a/Nominal/ExLet.thy Thu Mar 25 07:21:41 2010 +0100 +++ b/Nominal/ExLet.thy Thu Mar 25 08:05:03 2010 +0100 @@ -25,10 +25,82 @@ thm trm_lts.eq_iff thm trm_lts.bn thm trm_lts.perm -thm trm_lts.induct +thm trm_lts.induct[no_vars] +thm trm_lts.inducts[no_vars] thm trm_lts.distinct thm trm_lts.fv[simplified trm_lts.supp] +lemma + fixes t::trm + and l::lts + and c::"'a::fs" + assumes a1: "\name c. P1 c (Vr name)" + and a2: "\trm1 trm2 c. \\d. P1 d trm1; \d. P1 d trm2\ \ P1 c (Ap trm1 trm2)" + and a3: "\name trm c. \atom name \ c; \d. P1 d trm\ \ P1 c (Lm name trm)" + and a4: "\lts trm c. \bn lts \* c; \d. P2 d lts; \d. P1 d trm\ \ P1 c (Lt lts trm)" + and a5: "\c. P2 c Lnil" + and a6: "\name trm lts c. \\d. P1 d trm; \d. P2 d lts\ \ P2 c (Lcons name trm lts)" + shows "P1 c t" and "P2 c l" +proof - + have "(\(p::perm) (c::'a::fs). P1 c (p \ t))" and + "(\(p::perm) (c::'a::fs). P2 c (p \ l))" + apply(induct rule: trm_lts.inducts) + apply(simp) + apply(rule a1) + apply(simp) + apply(rule a2) + apply(simp) + apply(simp) + apply(simp) + apply(subgoal_tac "\q. (q \ (atom (p \ name))) \ c \ supp (Lm (p \ name) (p \ trm)) \* q") + apply(erule exE) + apply(rule_tac t="Lm (p \ name) (p \ trm)" + and s="q\ Lm (p \ name) (p \ trm)" in subst) + apply(rule supp_perm_eq) + apply(simp) + apply(simp) + apply(rule a3) + apply(simp add: atom_eqvt) + apply(subst permute_plus[symmetric]) + apply(blast) + apply(rule at_set_avoiding2_atom) + apply(simp add: finite_supp) + apply(simp add: finite_supp) + apply(simp add: fresh_def) + apply(simp add: trm_lts.fv[simplified trm_lts.supp]) + apply(simp) + apply(subgoal_tac "\q. (q \ bn (p \ lts)) \* c \ supp (Lt (p \ lts) (p \ trm)) \* q") + apply(erule exE) + apply(rule_tac t="Lt (p \ lts) (p \ trm)" + and s="q \ Lt (p \ lts) (p \ trm)" in subst) + apply(rule supp_perm_eq) + apply(simp) + apply(simp) + apply(rule a4) + apply(simp add: eqvts) + apply(subst permute_plus[symmetric]) + apply(blast) + apply(subst permute_plus[symmetric]) + apply(blast) + apply(rule at_set_avoiding2) + apply(simp add: finite_supp) + defer + apply(simp add: finite_supp) + apply(simp add: finite_supp) + apply(simp add: fresh_star_def) + apply(simp add: fresh_def) + apply(simp add: trm_lts.fv[simplified trm_lts.supp]) + defer + apply(simp) + apply(rule a5) + apply(simp) + apply(rule a6) + apply(simp) + apply(simp) + oops + + + lemma lets_bla: "x \ z \ y \ z \ x \ y \(Lt (Lcons x (Vr y) Lnil) (Vr x)) \ (Lt (Lcons x (Vr z) Lnil) (Vr x))" by (simp add: trm_lts.eq_iff)