| 76032 |      1 | (*  Author:     Tobias Nipkow
 | 
| 76031 |      2 |     Copyright   2000 TUM
 | 
|  |      3 | *)
 | 
|  |      4 | 
 | 
|  |      5 | section \<open>Fixed Length Lists\<close>
 | 
|  |      6 | 
 | 
|  |      7 | theory NList
 | 
|  |      8 | imports Main
 | 
|  |      9 | begin
 | 
|  |     10 | 
 | 
|  |     11 | definition nlists :: "nat \<Rightarrow> 'a set \<Rightarrow> 'a list set"
 | 
|  |     12 |   where "nlists n A = {xs. size xs = n \<and> set xs \<subseteq> A}"
 | 
|  |     13 | 
 | 
|  |     14 | lemma nlistsI: "\<lbrakk> size xs = n; set xs \<subseteq> A \<rbrakk> \<Longrightarrow> xs \<in> nlists n A"
 | 
|  |     15 |   by (simp add: nlists_def)
 | 
|  |     16 | 
 | 
|  |     17 | text \<open>These [simp] attributes are double-edged.
 | 
|  |     18 |  Many proofs in Jinja rely on it but they can degrade performance.\<close>
 | 
|  |     19 | 
 | 
|  |     20 | lemma nlistsE_length [simp]: "xs \<in> nlists n A \<Longrightarrow> size xs = n"
 | 
|  |     21 |   by (simp add: nlists_def)
 | 
|  |     22 | 
 | 
|  |     23 | lemma less_lengthI: "\<lbrakk> xs \<in> nlists n A; p < n \<rbrakk> \<Longrightarrow> p < size xs"
 | 
|  |     24 | by (simp)
 | 
|  |     25 | 
 | 
|  |     26 | lemma nlistsE_set[simp]: "xs \<in> nlists n A \<Longrightarrow> set xs \<subseteq> A"
 | 
|  |     27 | unfolding nlists_def by (simp)
 | 
|  |     28 | 
 | 
|  |     29 | lemma nlists_mono:
 | 
|  |     30 | assumes "A \<subseteq> B" shows "nlists n A \<subseteq> nlists n B"
 | 
|  |     31 | proof
 | 
|  |     32 |   fix xs assume "xs \<in> nlists n A"
 | 
|  |     33 |   then obtain size: "size xs = n" and inA: "set xs \<subseteq> A" by (simp)
 | 
|  |     34 |   with assms have "set xs \<subseteq> B" by simp
 | 
|  |     35 |   with size show "xs \<in> nlists n B" by(clarsimp intro!: nlistsI)
 | 
|  |     36 | qed
 | 
|  |     37 | 
 | 
|  |     38 | lemma nlists_n_0 [simp]: "nlists 0 A = {[]}"
 | 
|  |     39 | unfolding nlists_def by (auto)
 | 
|  |     40 | 
 | 
|  |     41 | lemma in_nlists_Suc_iff: "(xs \<in> nlists (Suc n) A) = (\<exists>y\<in>A. \<exists>ys \<in> nlists n A. xs = y#ys)"
 | 
|  |     42 | unfolding nlists_def by (cases "xs") auto
 | 
|  |     43 | 
 | 
|  |     44 | lemma Cons_in_nlists_Suc [iff]: "(x#xs \<in> nlists (Suc n) A) \<longleftrightarrow> (x\<in>A \<and> xs \<in> nlists n A)"
 | 
|  |     45 | unfolding nlists_def by (auto)
 | 
|  |     46 | 
 | 
|  |     47 | lemma nlists_not_empty: "A\<noteq>{} \<Longrightarrow> \<exists>xs. xs \<in> nlists n A"
 | 
|  |     48 | by (induct "n") (auto simp: in_nlists_Suc_iff)
 | 
|  |     49 | 
 | 
|  |     50 | 
 | 
|  |     51 | lemma nlistsE_nth_in: "\<lbrakk> xs \<in> nlists n A; i < n \<rbrakk> \<Longrightarrow> xs!i \<in> A"
 | 
|  |     52 | unfolding nlists_def by (auto)
 | 
|  |     53 | 
 | 
|  |     54 | lemma nlists_Cons_Suc [elim!]:
 | 
|  |     55 |   "l#xs \<in> nlists n A \<Longrightarrow> (\<And>n'. n = Suc n' \<Longrightarrow> l \<in> A \<Longrightarrow> xs \<in> nlists n' A \<Longrightarrow> P) \<Longrightarrow> P"
 | 
|  |     56 | unfolding nlists_def by (auto)
 | 
|  |     57 | 
 | 
|  |     58 | lemma nlists_appendE [elim!]:
 | 
|  |     59 |   "a@b \<in> nlists n A \<Longrightarrow> (\<And>n1 n2. n=n1+n2 \<Longrightarrow> a \<in> nlists n1 A \<Longrightarrow> b \<in> nlists n2 A \<Longrightarrow> P) \<Longrightarrow> P"
 | 
|  |     60 | proof -
 | 
|  |     61 |   have "\<And>n. a@b \<in> nlists n A \<Longrightarrow> \<exists>n1 n2. n=n1+n2 \<and> a \<in> nlists n1 A \<and> b \<in> nlists n2 A"
 | 
|  |     62 |     (is "\<And>n. ?list a n \<Longrightarrow> \<exists>n1 n2. ?P a n n1 n2")
 | 
|  |     63 |   proof (induct a)
 | 
|  |     64 |     fix n assume "?list [] n"
 | 
|  |     65 |     hence "?P [] n 0 n" by simp
 | 
|  |     66 |     thus "\<exists>n1 n2. ?P [] n n1 n2" by fast
 | 
|  |     67 |   next
 | 
|  |     68 |     fix n l ls
 | 
|  |     69 |     assume "?list (l#ls) n"
 | 
|  |     70 |     then obtain n' where n: "n = Suc n'" "l \<in> A" and n': "ls@b \<in> nlists n' A" by fastforce
 | 
|  |     71 |     assume "\<And>n. ls @ b \<in> nlists n A \<Longrightarrow> \<exists>n1 n2. n = n1 + n2 \<and> ls \<in> nlists n1 A \<and> b \<in> nlists n2 A"
 | 
|  |     72 |     from this and n' have "\<exists>n1 n2. n' = n1 + n2 \<and> ls \<in> nlists n1 A \<and> b \<in> nlists n2 A" .
 | 
|  |     73 |     then obtain n1 n2 where "n' = n1 + n2" "ls \<in> nlists n1 A" "b \<in> nlists n2 A" by fast
 | 
|  |     74 |     with n have "?P (l#ls) n (n1+1) n2" by simp
 | 
|  |     75 |     thus "\<exists>n1 n2. ?P (l#ls) n n1 n2" by fastforce
 | 
|  |     76 |   qed
 | 
|  |     77 |   moreover assume "a@b \<in> nlists n A" "\<And>n1 n2. n=n1+n2 \<Longrightarrow> a \<in> nlists n1 A \<Longrightarrow> b \<in> nlists n2 A \<Longrightarrow> P"
 | 
|  |     78 |   ultimately show ?thesis by blast
 | 
|  |     79 | qed
 | 
|  |     80 | 
 | 
|  |     81 | 
 | 
|  |     82 | lemma nlists_update_in_list [simp, intro!]:
 | 
|  |     83 |   "\<lbrakk> xs \<in> nlists n A; x\<in>A \<rbrakk> \<Longrightarrow> xs[i := x] \<in> nlists n A"
 | 
|  |     84 |   by (metis length_list_update nlistsE_length nlistsE_set nlistsI set_update_subsetI)
 | 
|  |     85 | 
 | 
|  |     86 | lemma nlists_appendI [intro?]:
 | 
|  |     87 |   "\<lbrakk> a \<in> nlists n A; b \<in> nlists m A \<rbrakk> \<Longrightarrow> a @ b \<in> nlists (n+m) A"
 | 
|  |     88 | unfolding nlists_def by (auto)
 | 
|  |     89 | 
 | 
|  |     90 | lemma nlists_append:
 | 
|  |     91 |   "xs @ ys \<in> nlists k A \<longleftrightarrow>
 | 
|  |     92 |    k = length(xs @ ys) \<and> xs \<in> nlists (length xs) A \<and> ys \<in> nlists (length ys) A"
 | 
|  |     93 | unfolding nlists_def by (auto)
 | 
|  |     94 | 
 | 
|  |     95 | lemma nlists_map [simp]: "(map f xs \<in> nlists (size xs) A) = (f ` set xs \<subseteq> A)"
 | 
|  |     96 | unfolding nlists_def by (auto)
 | 
|  |     97 | 
 | 
|  |     98 | lemma nlists_replicateI [intro]: "x \<in> A \<Longrightarrow> replicate n x \<in> nlists n A"
 | 
|  |     99 |  by (induct n) auto
 | 
|  |    100 | 
 | 
|  |    101 | lemma nlists_set[code]: "nlists n (set xs) = set (List.n_lists n xs)"
 | 
|  |    102 | unfolding nlists_def by (rule sym, induct n) (auto simp: image_iff length_Suc_conv)
 | 
|  |    103 | 
 | 
|  |    104 | end
 |