| author | wenzelm | 
| Sat, 08 Jul 2006 12:54:38 +0200 | |
| changeset 20050 | a2fb9d553aad | 
| parent 19972 | 89c5afe4139a | 
| child 20503 | 503ac4c5ef91 | 
| permissions | -rw-r--r-- | 
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 1 | (* $Id$ *) | 
| 19496 | 2 | |
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 3 | theory Iteration | 
| 19496 | 4 | imports "../Nominal" | 
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 5 | begin | 
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 6 | |
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 7 | atom_decl name | 
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 8 | |
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 9 | nominal_datatype lam = Var "name" | 
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 10 | | App "lam" "lam" | 
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 11 |                      | Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
 | 
| 19858 | 12 | |
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 13 | types 'a f1_ty  = "name\<Rightarrow>('a::pt_name)"
 | 
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 14 |       'a f2_ty  = "'a\<Rightarrow>'a\<Rightarrow>('a::pt_name)"
 | 
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 15 |       'a f3_ty  = "name\<Rightarrow>'a\<Rightarrow>('a::pt_name)"
 | 
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 16 | |
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 17 | consts | 
| 19651 | 18 | it :: "'a f1_ty \<Rightarrow> 'a f2_ty \<Rightarrow> 'a f3_ty \<Rightarrow> (lam \<times> 'a::pt_name) set" | 
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 19 | |
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 20 | inductive "it f1 f2 f3" | 
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 21 | intros | 
| 19651 | 22 | it1: "(Var a, f1 a) \<in> it f1 f2 f3" | 
| 23 | it2: "\<lbrakk>(t1,r1) \<in> it f1 f2 f3; (t2,r2) \<in> it f1 f2 f3\<rbrakk> \<Longrightarrow> (App t1 t2, f2 r1 r2) \<in> it f1 f2 f3" | |
| 24 | it3: "\<lbrakk>a\<sharp>(f1,f2,f3); (t,r) \<in> it f1 f2 f3\<rbrakk> \<Longrightarrow> (Lam [a].t,f3 a r) \<in> it f1 f2 f3" | |
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 25 | |
| 19651 | 26 | lemma it_equiv: | 
| 27 | fixes pi::"name prm" | |
| 28 | assumes a: "(t,r) \<in> it f1 f2 f3" | |
| 29 | shows "(pi\<bullet>t,pi\<bullet>r) \<in> it (pi\<bullet>f1) (pi\<bullet>f2) (pi\<bullet>f3)" | |
| 30 | using a | |
| 31 | apply(induct) | |
| 32 | apply(perm_simp | auto intro!: it.intros simp add: fresh_right)+ | |
| 33 | done | |
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 34 | |
| 19651 | 35 | lemma it_fin_supp: | 
| 36 | assumes f: "finite ((supp (f1,f2,f3))::name set)" | |
| 37 | and a: "(t,r) \<in> it f1 f2 f3" | |
| 38 | shows "finite ((supp r)::name set)" | |
| 39 | using a f | |
| 40 | apply(induct) | |
| 41 | apply(finite_guess, simp add: supp_prod fs_name1)+ | |
| 42 | done | |
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 43 | |
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 44 | lemma it_total: | 
| 19651 | 45 | assumes a: "finite ((supp (f1,f2,f3))::name set)" | 
| 46 | and b: "\<exists>(a::name). a\<sharp>f3 \<and> (\<forall>(y::'a::pt_name). a\<sharp>f3 a y)" | |
| 47 | shows "\<exists>r. (t,r)\<in>it f1 f2 f3" | |
| 19858 | 48 | apply(rule_tac lam.induct'[of "\<lambda>_. (supp (f1,f2,f3))" "\<lambda>z. \<lambda>t. \<exists>r. (t,r)\<in>it f1 f2 f3", simplified]) | 
| 19651 | 49 | apply(fold fresh_def) | 
| 50 | apply(auto intro: it.intros a) | |
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 51 | done | 
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 52 | |
| 19651 | 53 | lemma it_unique: | 
| 54 | assumes a: "finite ((supp (f1,f2,f3))::name set)" | |
| 55 | and b: "\<exists>(a::name). a\<sharp>f3 \<and> (\<forall>(y::'a::pt_name). a\<sharp>f3 a y)" | |
| 56 | and c1: "(t,r)\<in>it f1 f2 f3" | |
| 57 | and c2: "(t,r')\<in>it f1 f2 f3" | |
| 58 | shows "r=r'" | |
| 59 | using c1 c2 | |
| 60 | proof (induct fixing: r') | |
| 61 | case it1 | |
| 62 | then show ?case by cases (simp_all add: lam.inject) | |
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 63 | next | 
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 64 | case (it2 r1 r2 t1 t2) | 
| 19651 | 65 | have ih1: "\<And>r'. (t1,r') \<in> it f1 f2 f3 \<Longrightarrow> r1 = r'" by fact | 
| 66 | have ih2: "\<And>r'. (t2,r') \<in> it f1 f2 f3 \<Longrightarrow> r2 = r'" by fact | |
| 67 | have "(App t1 t2, r') \<in>it f1 f2 f3" by fact | |
| 68 | then show ?case | |
| 69 | proof cases | |
| 70 | case it2 | |
| 71 | then show ?thesis using ih1 ih2 by (simp add: lam.inject) | |
| 72 | qed (simp_all (no_asm_use)) | |
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 73 | next | 
| 19651 | 74 | case (it3 a1 r1 t1) | 
| 75 | have f1: "a1\<sharp>(f1,f2,f3)" by fact | |
| 76 | have ih: "\<And>r'. (t1,r') \<in> it f1 f2 f3 \<Longrightarrow> r1 = r'" by fact | |
| 77 | have it1: "(t1,r1) \<in> it f1 f2 f3" by fact | |
| 78 | have "(Lam [a1].t1, r') \<in> it f1 f2 f3" by fact | |
| 79 | then show ?case | |
| 80 | proof cases | |
| 81 | case (it3 a2 r2 t2) | |
| 82 | then have f2: "a2\<sharp>(f1,f2,f3)" | |
| 83 | and it2: "(t2,r2) \<in> it f1 f2 f3" | |
| 84 | and eq1: "[a1].t1 = [a2].t2" and eq2: "r' = f3 a2 r2" by (simp_all add: lam.inject) | |
| 85 | have "\<exists>(c::name). c\<sharp>(f1,f2,f3,a1,a2,t1,t2,r1,r2)" using a it1 it2 | |
| 86 | by (auto intro!: at_exists_fresh[OF at_name_inst] simp add: supp_prod fs_name1 it_fin_supp[OF a]) | |
| 87 | then obtain c where fresh: "c\<sharp>f1" "c\<sharp>f2" "c\<sharp>f3" "c\<noteq>a1" "c\<noteq>a2" "c\<sharp>t1" "c\<sharp>t2" "c\<sharp>r1" "c\<sharp>r2" | |
| 88 | by (force simp add: fresh_prod fresh_atm) | |
| 89 | have eq3: "[(a1,c)]\<bullet>t1 = [(a2,c)]\<bullet>t2" using eq1 fresh | |
| 90 | apply(auto simp add: alpha) | |
| 91 | apply(rule trans) | |
| 92 | apply(rule perm_compose) | |
| 93 | apply(simp add: calc_atm perm_fresh_fresh) | |
| 94 | apply(rule pt_name3, rule at_ds5[OF at_name_inst]) | |
| 95 | done | |
| 96 | have eq4: "[(a1,c)]\<bullet>r1 = [(a2,c)]\<bullet>r2" using eq3 it2 f1 f2 fresh | |
| 97 | apply(drule_tac sym) | |
| 98 | apply(rule_tac pt_bij2[OF pt_name_inst, OF at_name_inst]) | |
| 99 | apply(rule ih) | |
| 100 | apply(drule_tac pi="[(a2,c)]" in it_equiv) | |
| 101 | apply(perm_simp only: fresh_prod) | |
| 102 | apply(drule_tac pi="[(a1,c)]" in it_equiv) | |
| 103 | apply(perm_simp) | |
| 104 | done | |
| 105 | have fs1: "a1\<sharp>f3 a1 r1" using b f1 | |
| 106 | apply(auto) | |
| 107 | apply(rule_tac pi="[(a1,a)]" in pt_fresh_bij2[OF pt_name_inst, OF at_name_inst]) | |
| 108 | apply(perm_simp add: calc_atm fresh_prod) | |
| 109 | done | |
| 110 | have fs2: "a2\<sharp>f3 a2 r2" using b f2 | |
| 111 | apply(auto) | |
| 112 | apply(rule_tac pi="[(a2,a)]" in pt_fresh_bij2[OF pt_name_inst, OF at_name_inst]) | |
| 113 | apply(perm_simp add: calc_atm fresh_prod) | |
| 114 | done | |
| 115 | have fs3: "c\<sharp>f3 a1 r1" using fresh it1 a | |
| 19858 | 116 | by (fresh_guess add: supp_prod fs_name1 it_fin_supp[OF a] fresh_atm) | 
| 19651 | 117 | have fs4: "c\<sharp>f3 a2 r2" using fresh it2 a | 
| 19858 | 118 | by (fresh_guess add: supp_prod fs_name1 it_fin_supp[OF a] fresh_atm) | 
| 19651 | 119 | have "f3 a1 r1 = [(a1,c)]\<bullet>(f3 a1 r1)" using fs1 fs3 by perm_simp | 
| 120 | also have "\<dots> = f3 c ([(a1,c)]\<bullet>r1)" using f1 fresh by (perm_simp add: calc_atm fresh_prod) | |
| 121 | also have "\<dots> = f3 c ([(a2,c)]\<bullet>r2)" using eq4 by simp | |
| 122 | also have "\<dots> = [(a2,c)]\<bullet>(f3 a2 r2)" using f2 fresh by (perm_simp add: calc_atm fresh_prod) | |
| 123 | also have "\<dots> = f3 a2 r2" using fs2 fs4 by perm_simp | |
| 124 | finally have eq4: "f3 a1 r1 = f3 a2 r2" by simp | |
| 125 | then show ?thesis using eq2 by simp | |
| 126 | qed (simp_all (no_asm_use)) | |
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 127 | qed | 
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 128 | |
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 129 | lemma it_function: | 
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 130 | assumes f: "finite ((supp (f1,f2,f3))::name set)" | 
| 19158 | 131 | and c: "\<exists>(a::name). a\<sharp>f3 \<and> (\<forall>(r::'a::pt_name). a\<sharp>f3 a r)" | 
| 132 | shows "\<exists>!r. (t,r) \<in> it f1 f2 f3" | |
| 19651 | 133 | proof (rule ex_ex1I, rule it_total[OF f, OF c]) | 
| 19225 | 134 | case (goal1 r1 r2) | 
| 19158 | 135 | have a1: "(t,r1) \<in> it f1 f2 f3" and a2: "(t,r2) \<in> it f1 f2 f3" by fact | 
| 19651 | 136 | thus "r1 = r2" using it_unique[OF f, OF c] by simp | 
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 137 | qed | 
| 19651 | 138 | |
| 139 | constdefs | |
| 140 |   itfun :: "'a f1_ty \<Rightarrow> 'a f2_ty \<Rightarrow> 'a f3_ty \<Rightarrow> lam \<Rightarrow> ('a::pt_name)" 
 | |
| 141 | "itfun f1 f2 f3 t \<equiv> (THE r. (t,r) \<in> it f1 f2 f3)" | |
| 142 | ||
| 143 | lemma itfun_eqvt: | |
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 144 | fixes pi::"name prm" | 
| 19158 | 145 | assumes f: "finite ((supp (f1,f2,f3))::name set)" | 
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 146 | and c: "\<exists>(a::name). a\<sharp>f3 \<and> (\<forall>(r::'a::pt_name). a\<sharp>f3 a r)" | 
| 19651 | 147 | shows "pi\<bullet>(itfun f1 f2 f3 t) = itfun (pi\<bullet>f1) (pi\<bullet>f2) (pi\<bullet>f3) (pi\<bullet>t)" | 
| 148 | proof - | |
| 149 | have f_pi: "finite ((supp (pi\<bullet>f1,pi\<bullet>f2,pi\<bullet>f3))::name set)" using f | |
| 150 | by (simp add: supp_prod pt_supp_finite_pi[OF pt_name_inst, OF at_name_inst]) | |
| 151 | have fs_pi: "\<exists>(a::name). a\<sharp>(pi\<bullet>f3) \<and> (\<forall>(r::'a::pt_name). a\<sharp>(pi\<bullet>f3) a r)" | |
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 152 | proof - | 
| 19651 | 153 | from c obtain a where fs1: "a\<sharp>f3" and fs2: "\<forall>(r::'a::pt_name). a\<sharp>f3 a r" by force | 
| 19972 
89c5afe4139a
added more infrastructure for the recursion combinator
 urbanc parents: 
19858diff
changeset | 154 | have "(pi\<bullet>a)\<sharp>(pi\<bullet>f3)" using fs1 by (simp add: fresh_bij) | 
| 19168 | 155 | moreover | 
| 19651 | 156 | have "\<forall>(r::'a::pt_name). (pi\<bullet>a)\<sharp>((pi\<bullet>f3) (pi\<bullet>a) r)" using fs2 by (perm_simp add: fresh_right) | 
| 157 | ultimately show "\<exists>(a::name). a\<sharp>(pi\<bullet>f3) \<and> (\<forall>(r::'a::pt_name). a\<sharp>(pi\<bullet>f3) a r)" by blast | |
| 19168 | 158 | qed | 
| 159 | show ?thesis | |
| 160 | apply(rule sym) | |
| 19651 | 161 | apply(auto simp add: itfun_def) | 
| 162 | apply(rule the1_equality[OF it_function, OF f_pi, OF fs_pi]) | |
| 163 | apply(rule it_equiv) | |
| 19168 | 164 | apply(rule theI'[OF it_function,OF f, OF c]) | 
| 165 | done | |
| 166 | qed | |
| 167 | ||
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 168 | lemma itfun_Var: | 
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 169 | assumes f: "finite ((supp (f1,f2,f3))::name set)" | 
| 19651 | 170 | and c: "\<exists>(a::name). a\<sharp>f3 \<and> (\<forall>(r::'a::pt_name). a\<sharp>f3 a r)" | 
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 171 | shows "itfun f1 f2 f3 (Var c) = (f1 c)" | 
| 19651 | 172 | using f c by (auto intro!: the1_equality it_function it.intros simp add: itfun_def) | 
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 173 | |
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 174 | lemma itfun_App: | 
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 175 | assumes f: "finite ((supp (f1,f2,f3))::name set)" | 
| 19651 | 176 | and c: "\<exists>(a::name). a\<sharp>f3 \<and> (\<forall>(r::'a::pt_name). a\<sharp>f3 a r)" | 
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 177 | shows "itfun f1 f2 f3 (App t1 t2) = (f2 (itfun f1 f2 f3 t1) (itfun f1 f2 f3 t2))" | 
| 19651 | 178 | by (auto intro!: the1_equality it_function[OF f, OF c] it.intros | 
| 179 | intro: theI'[OF it_function, OF f, OF c] simp add: itfun_def) | |
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 180 | |
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 181 | lemma itfun_Lam: | 
| 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 182 | assumes f: "finite ((supp (f1,f2,f3))::name set)" | 
| 19651 | 183 | and c: "\<exists>(a::name). a\<sharp>f3 \<and> (\<forall>(r::'a::pt_name). a\<sharp>f3 a r)" | 
| 184 | and a: "a\<sharp>(f1,f2,f3)" | |
| 185 | shows "itfun f1 f2 f3 (Lam [a].t) = f3 a (itfun f1 f2 f3 t)" | |
| 186 | using a | |
| 187 | by (auto intro!: the1_equality it_function[OF f, OF c] it.intros | |
| 188 | intro: theI'[OF it_function, OF f, OF c] simp add: itfun_def) | |
| 19496 | 189 | |
| 19157 
6e4ce7402dbe
initial commit (especially 2nd half needs to be cleaned up)
 urbanc parents: diff
changeset | 190 | end |