| author | traytel | 
| Thu, 25 Jul 2013 12:25:07 +0200 | |
| changeset 52730 | 6bf02eb4ddf7 | 
| parent 51772 | d2b265ebc1fa | 
| child 52992 | abd760a19e22 | 
| permissions | -rw-r--r-- | 
| 50530 | 1 | (* Title: HOL/BNF/Examples/Koenig.thy | 
| 50517 | 2 | Author: Dmitriy Traytel, TU Muenchen | 
| 3 | Author: Andrei Popescu, TU Muenchen | |
| 4 | Copyright 2012 | |
| 5 | ||
| 6 | Koenig's lemma. | |
| 7 | *) | |
| 8 | ||
| 9 | header {* Koenig's lemma *}
 | |
| 10 | ||
| 11 | theory Koenig | |
| 50518 | 12 | imports TreeFI Stream | 
| 50517 | 13 | begin | 
| 14 | ||
| 15 | (* selectors for streams *) | |
| 16 | lemma shd_def': "shd as = fst (stream_dtor as)" | |
| 17 | unfolding shd_def stream_case_def fst_def by (rule refl) | |
| 18 | ||
| 19 | lemma stl_def': "stl as = snd (stream_dtor as)" | |
| 20 | unfolding stl_def stream_case_def snd_def by (rule refl) | |
| 21 | ||
| 22 | lemma unfold_pair_fun_shd[simp]: "shd (stream_dtor_unfold (f \<odot> g) t) = f t" | |
| 23 | unfolding shd_def' pair_fun_def stream.dtor_unfold by simp | |
| 24 | ||
| 25 | lemma unfold_pair_fun_stl[simp]: "stl (stream_dtor_unfold (f \<odot> g) t) = | |
| 26 | stream_dtor_unfold (f \<odot> g) (g t)" | |
| 27 | unfolding stl_def' pair_fun_def stream.dtor_unfold by simp | |
| 28 | ||
| 29 | (* infinite trees: *) | |
| 30 | coinductive infiniteTr where | |
| 31 | "\<lbrakk>tr' \<in> listF_set (sub tr); infiniteTr tr'\<rbrakk> \<Longrightarrow> infiniteTr tr" | |
| 32 | ||
| 33 | lemma infiniteTr_strong_coind[consumes 1, case_names sub]: | |
| 34 | assumes *: "phi tr" and | |
| 35 | **: "\<And> tr. phi tr \<Longrightarrow> \<exists> tr' \<in> listF_set (sub tr). phi tr' \<or> infiniteTr tr'" | |
| 36 | shows "infiniteTr tr" | |
| 37 | using assms by (elim infiniteTr.coinduct) blast | |
| 38 | ||
| 39 | lemma infiniteTr_coind[consumes 1, case_names sub, induct pred: infiniteTr]: | |
| 40 | assumes *: "phi tr" and | |
| 41 | **: "\<And> tr. phi tr \<Longrightarrow> \<exists> tr' \<in> listF_set (sub tr). phi tr'" | |
| 42 | shows "infiniteTr tr" | |
| 43 | using assms by (elim infiniteTr.coinduct) blast | |
| 44 | ||
| 45 | lemma infiniteTr_sub[simp]: | |
| 46 | "infiniteTr tr \<Longrightarrow> (\<exists> tr' \<in> listF_set (sub tr). infiniteTr tr')" | |
| 47 | by (erule infiniteTr.cases) blast | |
| 48 | ||
| 49 | definition "konigPath \<equiv> stream_dtor_unfold | |
| 50 | (lab \<odot> (\<lambda>tr. SOME tr'. tr' \<in> listF_set (sub tr) \<and> infiniteTr tr'))" | |
| 51 | ||
| 52 | lemma konigPath_simps[simp]: | |
| 53 | "shd (konigPath t) = lab t" | |
| 54 | "stl (konigPath t) = konigPath (SOME tr. tr \<in> listF_set (sub t) \<and> infiniteTr tr)" | |
| 55 | unfolding konigPath_def by simp+ | |
| 56 | ||
| 57 | (* proper paths in trees: *) | |
| 58 | coinductive properPath where | |
| 59 | "\<lbrakk>shd as = lab tr; tr' \<in> listF_set (sub tr); properPath (stl as) tr'\<rbrakk> \<Longrightarrow> | |
| 60 | properPath as tr" | |
| 61 | ||
| 62 | lemma properPath_strong_coind[consumes 1, case_names shd_lab sub]: | |
| 63 | assumes *: "phi as tr" and | |
| 64 | **: "\<And> as tr. phi as tr \<Longrightarrow> shd as = lab tr" and | |
| 65 | ***: "\<And> as tr. | |
| 66 | phi as tr \<Longrightarrow> | |
| 67 | \<exists> tr' \<in> listF_set (sub tr). phi (stl as) tr' \<or> properPath (stl as) tr'" | |
| 68 | shows "properPath as tr" | |
| 69 | using assms by (elim properPath.coinduct) blast | |
| 70 | ||
| 71 | lemma properPath_coind[consumes 1, case_names shd_lab sub, induct pred: properPath]: | |
| 72 | assumes *: "phi as tr" and | |
| 73 | **: "\<And> as tr. phi as tr \<Longrightarrow> shd as = lab tr" and | |
| 74 | ***: "\<And> as tr. | |
| 75 | phi as tr \<Longrightarrow> | |
| 76 | \<exists> tr' \<in> listF_set (sub tr). phi (stl as) tr'" | |
| 77 | shows "properPath as tr" | |
| 78 | using properPath_strong_coind[of phi, OF * **] *** by blast | |
| 79 | ||
| 80 | lemma properPath_shd_lab: | |
| 81 | "properPath as tr \<Longrightarrow> shd as = lab tr" | |
| 82 | by (erule properPath.cases) blast | |
| 83 | ||
| 84 | lemma properPath_sub: | |
| 85 | "properPath as tr \<Longrightarrow> | |
| 86 | \<exists> tr' \<in> listF_set (sub tr). phi (stl as) tr' \<or> properPath (stl as) tr'" | |
| 87 | by (erule properPath.cases) blast | |
| 88 | ||
| 89 | (* prove the following by coinduction *) | |
| 90 | theorem Konig: | |
| 91 | assumes "infiniteTr tr" | |
| 92 | shows "properPath (konigPath tr) tr" | |
| 93 | proof- | |
| 94 |   {fix as
 | |
| 95 | assume "infiniteTr tr \<and> as = konigPath tr" hence "properPath as tr" | |
| 96 | proof (induct rule: properPath_coind, safe) | |
| 97 | fix t | |
| 98 | let ?t = "SOME t'. t' \<in> listF_set (sub t) \<and> infiniteTr t'" | |
| 99 | assume "infiniteTr t" | |
| 100 | hence "\<exists>t' \<in> listF_set (sub t). infiniteTr t'" by simp | |
| 101 | hence "\<exists>t'. t' \<in> listF_set (sub t) \<and> infiniteTr t'" by blast | |
| 102 | hence "?t \<in> listF_set (sub t) \<and> infiniteTr ?t" by (elim someI_ex) | |
| 103 | moreover have "stl (konigPath t) = konigPath ?t" by simp | |
| 104 | ultimately show "\<exists>t' \<in> listF_set (sub t). | |
| 105 | infiniteTr t' \<and> stl (konigPath t) = konigPath t'" by blast | |
| 106 | qed simp | |
| 107 | } | |
| 108 | thus ?thesis using assms by blast | |
| 109 | qed | |
| 110 | ||
| 111 | (* some more stream theorems *) | |
| 112 | ||
| 51772 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 traytel parents: 
50530diff
changeset | 113 | lemma stream_map[simp]: "smap f = stream_dtor_unfold (f o shd \<odot> stl)" | 
| 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 traytel parents: 
50530diff
changeset | 114 | unfolding smap_def pair_fun_def shd_def'[abs_def] stl_def'[abs_def] | 
| 50517 | 115 | map_pair_def o_def prod_case_beta by simp | 
| 116 | ||
| 117 | definition plus :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream" (infixr "\<oplus>" 66) where | |
| 118 | [simp]: "plus xs ys = | |
| 119 | stream_dtor_unfold ((%(xs, ys). shd xs + shd ys) \<odot> (%(xs, ys). (stl xs, stl ys))) (xs, ys)" | |
| 120 | ||
| 121 | definition scalar :: "nat \<Rightarrow> nat stream \<Rightarrow> nat stream" (infixr "\<cdot>" 68) where | |
| 51772 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 traytel parents: 
50530diff
changeset | 122 | [simp]: "scalar n = smap (\<lambda>x. n * x)" | 
| 50517 | 123 | |
| 124 | definition ones :: "nat stream" where [simp]: "ones = stream_dtor_unfold ((%x. 1) \<odot> id) ()" | |
| 125 | definition twos :: "nat stream" where [simp]: "twos = stream_dtor_unfold ((%x. 2) \<odot> id) ()" | |
| 126 | definition ns :: "nat \<Rightarrow> nat stream" where [simp]: "ns n = scalar n ones" | |
| 127 | ||
| 128 | lemma "ones \<oplus> ones = twos" | |
| 129 | by (rule stream.coinduct[of "%x1 x2. \<exists>x. x1 = ones \<oplus> ones \<and> x2 = twos"]) auto | |
| 130 | ||
| 131 | lemma "n \<cdot> twos = ns (2 * n)" | |
| 132 | by (rule stream.coinduct[of "%x1 x2. \<exists>n. x1 = n \<cdot> twos \<and> x2 = ns (2 * n)"]) force+ | |
| 133 | ||
| 134 | lemma prod_scalar: "(n * m) \<cdot> xs = n \<cdot> m \<cdot> xs" | |
| 135 | by (rule stream.coinduct[of "%x1 x2. \<exists>n m xs. x1 = (n * m) \<cdot> xs \<and> x2 = n \<cdot> m \<cdot> xs"]) force+ | |
| 136 | ||
| 137 | lemma scalar_plus: "n \<cdot> (xs \<oplus> ys) = n \<cdot> xs \<oplus> n \<cdot> ys" | |
| 138 | by (rule stream.coinduct[of "%x1 x2. \<exists>n xs ys. x1 = n \<cdot> (xs \<oplus> ys) \<and> x2 = n \<cdot> xs \<oplus> n \<cdot> ys"]) | |
| 139 | (force simp: add_mult_distrib2)+ | |
| 140 | ||
| 141 | lemma plus_comm: "xs \<oplus> ys = ys \<oplus> xs" | |
| 142 | by (rule stream.coinduct[of "%x1 x2. \<exists>xs ys. x1 = xs \<oplus> ys \<and> x2 = ys \<oplus> xs"]) force+ | |
| 143 | ||
| 144 | lemma plus_assoc: "(xs \<oplus> ys) \<oplus> zs = xs \<oplus> ys \<oplus> zs" | |
| 145 | by (rule stream.coinduct[of "%x1 x2. \<exists>xs ys zs. x1 = (xs \<oplus> ys) \<oplus> zs \<and> x2 = xs \<oplus> ys \<oplus> zs"]) force+ | |
| 146 | ||
| 147 | end |