| author | Cezary Kaliszyk <kaliszyk@in.tum.de> | 
| Thu, 25 Mar 2010 17:30:46 +0100 | |
| changeset 1650 | 4b949985cf57 | 
| parent 1644 | 0e705352bcef | 
| child 1651 | f731e9aff866 | 
| permissions | -rw-r--r-- | 
| 1600 | 1 | theory ExLet | 
| 2 | imports "Parser" | |
| 3 | begin | |
| 4 | ||
| 5 | text {* example 3 or example 5 from Terms.thy *}
 | |
| 6 | ||
| 7 | atom_decl name | |
| 8 | ||
| 9 | ML {* val _ = recursive := false  *}
 | |
| 1650 
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1644diff
changeset | 10 | |
| 
4b949985cf57
Gathering things to prove by induction together; removed cheat_bn_eqvt.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1644diff
changeset | 11 | |
| 1600 | 12 | nominal_datatype trm = | 
| 13 | Vr "name" | |
| 14 | | Ap "trm" "trm" | |
| 15 | | Lm x::"name" t::"trm" bind x in t | |
| 1602 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 16 | | Lt a::"lts" t::"trm" bind "bn a" in t | 
| 1600 | 17 | and lts = | 
| 1602 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 18 | Lnil | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 19 | | Lcons "name" "trm" "lts" | 
| 1600 | 20 | binder | 
| 21 | bn | |
| 22 | where | |
| 1602 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 23 |   "bn Lnil = {}"
 | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 24 | | "bn (Lcons x t l) = {atom x} \<union> (bn l)"
 | 
| 1600 | 25 | |
| 26 | thm trm_lts.fv | |
| 27 | thm trm_lts.eq_iff | |
| 28 | thm trm_lts.bn | |
| 29 | thm trm_lts.perm | |
| 1638 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 30 | thm trm_lts.induct[no_vars] | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 31 | thm trm_lts.inducts[no_vars] | 
| 1600 | 32 | thm trm_lts.distinct | 
| 33 | thm trm_lts.fv[simplified trm_lts.supp] | |
| 34 | ||
| 1644 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 35 | primrec | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 36 | permute_bn_raw | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 37 | where | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 38 | "permute_bn_raw pi (Lnil_raw) = Lnil_raw" | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 39 | | "permute_bn_raw pi (Lcons_raw a t l) = Lcons_raw (pi \<bullet> a) t (permute_bn_raw pi l)" | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 40 | |
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 41 | quotient_definition | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 42 | "permute_bn :: perm \<Rightarrow> lts \<Rightarrow> lts" | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 43 | is | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 44 | "permute_bn_raw" | 
| 1639 
a98d03fb9d53
added experiemental permute_bn
 Christian Urban <urbanc@in.tum.de> parents: 
1638diff
changeset | 45 | |
| 1644 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 46 | lemma [quot_respect]: "((op =) ===> alpha_lts_raw ===> alpha_lts_raw) permute_bn_raw permute_bn_raw" | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 47 | apply simp | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 48 | apply clarify | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 49 | apply (erule alpha_trm_raw_alpha_lts_raw_alpha_bn_raw.inducts) | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 50 | apply simp_all | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 51 | apply (rule alpha_trm_raw_alpha_lts_raw_alpha_bn_raw.intros) | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 52 | apply simp | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 53 | apply (rule alpha_trm_raw_alpha_lts_raw_alpha_bn_raw.intros) | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 54 | apply simp | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 55 | done | 
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 56 | |
| 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 57 | lemmas permute_bn = permute_bn_raw.simps[quot_lifted] | 
| 1639 
a98d03fb9d53
added experiemental permute_bn
 Christian Urban <urbanc@in.tum.de> parents: 
1638diff
changeset | 58 | |
| 1642 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 59 | lemma permute_bn_zero: | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 60 | "permute_bn 0 a = a" | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 61 | apply(induct a rule: trm_lts.inducts(2)) | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 62 | apply(rule TrueI) | 
| 1644 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 63 | apply(simp_all add:permute_bn eqvts) | 
| 1642 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 64 | done | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 65 | |
| 1640 
cd5a6db05540
trying to prove the string induction for let.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1639diff
changeset | 66 | lemma permute_bn_add: | 
| 
cd5a6db05540
trying to prove the string induction for let.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1639diff
changeset | 67 | "permute_bn (p + q) a = permute_bn p (permute_bn q a)" | 
| 
cd5a6db05540
trying to prove the string induction for let.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1639diff
changeset | 68 | oops | 
| 
cd5a6db05540
trying to prove the string induction for let.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1639diff
changeset | 69 | |
| 1643 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 70 | lemma permute_bn_alpha_bn: "alpha_bn lts (permute_bn q lts)" | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 71 | apply(induct lts rule: trm_lts.inducts(2)) | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 72 | apply(rule TrueI) | 
| 1644 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 73 | apply(simp_all add:permute_bn eqvts trm_lts.eq_iff) | 
| 1643 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 74 | done | 
| 1641 | 75 | |
| 1642 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 76 | lemma perm_bn: | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 77 | "p \<bullet> bn l = bn(permute_bn p l)" | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 78 | apply(induct l rule: trm_lts.inducts(2)) | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 79 | apply(rule TrueI) | 
| 1644 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 80 | apply(simp_all add:permute_bn eqvts) | 
| 1642 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 81 | done | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 82 | |
| 1643 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 83 | lemma Lt_subst: | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 84 | "supp (Abs (bn lts) trm) \<sharp>* q \<Longrightarrow> (Lt lts trm) = Lt (permute_bn q lts) (q \<bullet> trm)" | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 85 | apply (simp only: trm_lts.eq_iff) | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 86 | apply (rule_tac x="q" in exI) | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 87 | apply (simp add: alphas) | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 88 | apply (simp add: permute_bn_alpha_bn) | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 89 | apply (simp add: perm_bn[symmetric]) | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 90 | apply (simp add: eqvts[symmetric]) | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 91 | apply (simp add: supp_Abs) | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 92 | apply (simp add: trm_lts.supp) | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 93 | apply (rule supp_perm_eq[symmetric]) | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 94 | apply (subst supp_finite_atom_set) | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 95 | apply (rule finite_Diff) | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 96 | apply (simp add: finite_supp) | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 97 | apply (assumption) | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 98 | done | 
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 99 | |
| 
953403c5faa0
Showed Let substitution.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1642diff
changeset | 100 | |
| 1642 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 101 | lemma fin_bn: | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 102 | "finite (bn l)" | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 103 | apply(induct l rule: trm_lts.inducts(2)) | 
| 1644 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 104 | apply(simp_all add:permute_bn eqvts) | 
| 1642 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 105 | done | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 106 | |
| 1638 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 107 | lemma | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 108 | fixes t::trm | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 109 | and l::lts | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 110 | and c::"'a::fs" | 
| 1640 
cd5a6db05540
trying to prove the string induction for let.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1639diff
changeset | 111 | assumes a1: "\<And>name c. P1 c (Vr name)" | 
| 1638 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 112 | and a2: "\<And>trm1 trm2 c. \<lbrakk>\<And>d. P1 d trm1; \<And>d. P1 d trm2\<rbrakk> \<Longrightarrow> P1 c (Ap trm1 trm2)" | 
| 1640 
cd5a6db05540
trying to prove the string induction for let.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1639diff
changeset | 113 | and a3: "\<And>name trm c. \<lbrakk>atom name \<sharp> c; \<And>d. P1 d trm\<rbrakk> \<Longrightarrow> P1 c (Lm name trm)" | 
| 1641 | 114 | and a4: "\<And>lts trm c. \<lbrakk>bn lts \<sharp>* c; \<And>d. P2 d lts; \<And>d. P1 d trm\<rbrakk> \<Longrightarrow> P1 c (Lt lts trm)" | 
| 1638 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 115 | and a5: "\<And>c. P2 c Lnil" | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 116 | and a6: "\<And>name trm lts c. \<lbrakk>\<And>d. P1 d trm; \<And>d. P2 d lts\<rbrakk> \<Longrightarrow> P2 c (Lcons name trm lts)" | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 117 | shows "P1 c t" and "P2 c l" | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 118 | proof - | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 119 | have "(\<And>(p::perm) (c::'a::fs). P1 c (p \<bullet> t))" and | 
| 1642 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 120 | b': "(\<And>(p::perm) (q::perm) (c::'a::fs). P2 c (permute_bn p (q \<bullet> l)))" | 
| 1638 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 121 | apply(induct rule: trm_lts.inducts) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 122 | apply(simp) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 123 | apply(rule a1) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 124 | apply(simp) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 125 | apply(rule a2) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 126 | apply(simp) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 127 | apply(simp) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 128 | apply(simp) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 129 | apply(subgoal_tac "\<exists>q. (q \<bullet> (atom (p \<bullet> name))) \<sharp> c \<and> supp (Lm (p \<bullet> name) (p \<bullet> trm)) \<sharp>* q") | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 130 | apply(erule exE) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 131 | apply(rule_tac t="Lm (p \<bullet> name) (p \<bullet> trm)" | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 132 | and s="q\<bullet> Lm (p \<bullet> name) (p \<bullet> trm)" in subst) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 133 | apply(rule supp_perm_eq) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 134 | apply(simp) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 135 | apply(simp) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 136 | apply(rule a3) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 137 | apply(simp add: atom_eqvt) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 138 | apply(subst permute_plus[symmetric]) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 139 | apply(blast) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 140 | apply(rule at_set_avoiding2_atom) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 141 | apply(simp add: finite_supp) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 142 | apply(simp add: finite_supp) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 143 | apply(simp add: fresh_def) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 144 | apply(simp add: trm_lts.fv[simplified trm_lts.supp]) | 
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 145 | apply(simp) | 
| 1642 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 146 | apply(subgoal_tac "\<exists>q. (q \<bullet> bn (p \<bullet> lts)) \<sharp>* c \<and> supp (Abs (bn (p \<bullet> lts)) (p \<bullet> trm)) \<sharp>* q") | 
| 1638 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 147 | apply(erule exE) | 
| 1641 | 148 | apply(erule conjE) | 
| 149 | apply(subst Lt_subst) | |
| 150 | apply assumption | |
| 1638 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 151 | apply(rule a4) | 
| 1642 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 152 | apply(simp add:perm_bn) | 
| 1641 | 153 | apply assumption | 
| 154 | apply (simp add: fresh_star_def fresh_def) | |
| 1640 
cd5a6db05540
trying to prove the string induction for let.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1639diff
changeset | 155 | apply(rotate_tac 1) | 
| 
cd5a6db05540
trying to prove the string induction for let.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1639diff
changeset | 156 | apply(drule_tac x="q + p" in meta_spec) | 
| 
cd5a6db05540
trying to prove the string induction for let.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1639diff
changeset | 157 | apply(simp) | 
| 1642 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 158 | apply(rule at_set_avoiding2) | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 159 | apply(rule fin_bn) | 
| 1641 | 160 | apply(simp add: finite_supp) | 
| 161 | apply(simp add: supp_Abs) | |
| 162 | apply(rule finite_Diff) | |
| 163 | apply(simp add: finite_supp) | |
| 1642 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 164 | apply(simp add: fresh_star_def fresh_def supp_Abs) | 
| 1644 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 165 | apply(simp add: eqvts permute_bn) | 
| 1640 
cd5a6db05540
trying to prove the string induction for let.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1639diff
changeset | 166 | apply(rule a5) | 
| 1644 
0e705352bcef
Properly defined permute_bn. No more sorry's in Let strong induction.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1643diff
changeset | 167 | apply(simp add: permute_bn) | 
| 1640 
cd5a6db05540
trying to prove the string induction for let.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1639diff
changeset | 168 | apply(rule a6) | 
| 
cd5a6db05540
trying to prove the string induction for let.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1639diff
changeset | 169 | apply simp | 
| 
cd5a6db05540
trying to prove the string induction for let.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1639diff
changeset | 170 | apply simp | 
| 1642 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 171 | done | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 172 | then have a: "P1 c (0 \<bullet> t)" by blast | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 173 | have "P2 c (permute_bn 0 (0 \<bullet> l))" using b' by blast | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 174 | then show "P1 c t" and "P2 c l" using a permute_bn_zero by simp_all | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 175 | qed | 
| 
06f44d498cef
Only let substitution is left.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1641diff
changeset | 176 | |
| 1638 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 177 | |
| 
36798cdbc452
first attempt of strong induction for lets with assignments
 Christian Urban <urbanc@in.tum.de> parents: 
1602diff
changeset | 178 | |
| 1602 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 179 | lemma lets_bla: | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 180 | "x \<noteq> z \<Longrightarrow> y \<noteq> z \<Longrightarrow> x \<noteq> y \<Longrightarrow>(Lt (Lcons x (Vr y) Lnil) (Vr x)) \<noteq> (Lt (Lcons x (Vr z) Lnil) (Vr x))" | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 181 | by (simp add: trm_lts.eq_iff) | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 182 | |
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 183 | lemma lets_ok: | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 184 | "(Lt (Lcons x (Vr y) Lnil) (Vr x)) = (Lt (Lcons y (Vr y) Lnil) (Vr y))" | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 185 | apply (simp add: trm_lts.eq_iff) | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 186 | apply (rule_tac x="(x \<leftrightarrow> y)" in exI) | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 187 | apply (simp_all add: alphas) | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 188 | apply (simp add: fresh_star_def eqvts) | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 189 | done | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 190 | |
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 191 | lemma lets_ok3: | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 192 | "x \<noteq> y \<Longrightarrow> | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 193 | (Lt (Lcons x (Ap (Vr y) (Vr x)) (Lcons y (Vr y) Lnil)) (Ap (Vr x) (Vr y))) \<noteq> | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 194 | (Lt (Lcons y (Ap (Vr x) (Vr y)) (Lcons x (Vr x) Lnil)) (Ap (Vr x) (Vr y)))" | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 195 | apply (simp add: alphas trm_lts.eq_iff) | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 196 | done | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 197 | |
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 198 | |
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 199 | lemma lets_not_ok1: | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 200 | "(Lt (Lcons x (Vr x) (Lcons y (Vr y) Lnil)) (Ap (Vr x) (Vr y))) = | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 201 | (Lt (Lcons y (Vr x) (Lcons x (Vr y) Lnil)) (Ap (Vr x) (Vr y)))" | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 202 | apply (simp add: alphas trm_lts.eq_iff) | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 203 | apply (rule_tac x="0::perm" in exI) | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 204 | apply (simp add: fresh_star_def eqvts) | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 205 | apply blast | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 206 | done | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 207 | |
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 208 | lemma lets_nok: | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 209 | "x \<noteq> y \<Longrightarrow> x \<noteq> z \<Longrightarrow> z \<noteq> y \<Longrightarrow> | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 210 | (Lt (Lcons x (Ap (Vr z) (Vr z)) (Lcons y (Vr z) Lnil)) (Ap (Vr x) (Vr y))) \<noteq> | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 211 | (Lt (Lcons y (Vr z) (Lcons x (Ap (Vr z) (Vr z)) Lnil)) (Ap (Vr x) (Vr y)))" | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 212 | apply (simp add: alphas trm_lts.eq_iff fresh_star_def) | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 213 | done | 
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 214 | |
| 
a7e60da429e2
Move Let properties to ExLet
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1600diff
changeset | 215 | |
| 1600 | 216 | end | 
| 217 | ||
| 218 | ||
| 219 |