diff --git a/src/HOL/Library/Library.thy b/src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy +++ b/src/HOL/Library/Library.thy @@ -1,100 +1,101 @@ (*<*) theory Library imports AList Adhoc_Overloading BigO BNF_Axiomatization BNF_Corec Bourbaki_Witt_Fixpoint Char_ord Code_Cardinality Code_Lazy Code_Test Combine_PER Complete_Partial_Order2 Conditional_Parametricity Confluence Confluent_Quotient Countable Countable_Complete_Lattices Countable_Set_Type Debug Diagonal_Subsequence Discrete Disjoint_Sets Disjoint_FSets Dlist Dual_Ordered_Lattice Equipollence Extended Extended_Nat Extended_Nonnegative_Real Extended_Real Finite_Map Float FSet FuncSet Function_Division Fun_Lexorder Going_To_Filter Groups_Big_Fun Indicator_Function Infinite_Set Interval Interval_Float IArray Landau_Symbols Lattice_Algebras Lattice_Constructions Linear_Temporal_Logic_on_Streams ListVector Lub_Glb Mapping Monad_Syntax More_List Multiset_Order + NList Nonpos_Ints Numeral_Type Omega_Words_Fun Open_State_Syntax Option_ord Order_Continuity Parallel Pattern_Aliases Periodic_Fun Poly_Mapping Power_By_Squaring Preorder Product_Plus Quadratic_Discriminant Quotient_List Quotient_Option Quotient_Product Quotient_Set Quotient_Sum Quotient_Syntax Quotient_Type Ramsey Reflection Rewrite Saturated Set_Algebras Set_Idioms Signed_Division State_Monad Stream Sorting_Algorithms Sublist Sum_of_Squares Transitive_Closure_Table Tree_Multiset Tree_Real Type_Length Uprod While_Combinator Word Z2 begin end (*>*) diff --git a/src/HOL/Library/NList.thy b/src/HOL/Library/NList.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Library/NList.thy @@ -0,0 +1,98 @@ +(* Author: Tobias Nipkow + Copyright 2000 TUM +*) + +section \Fixed Length Lists\ + +theory NList +imports Main +begin + +definition nlists :: "nat \ 'a set \ 'a list set" + where "nlists n A = {xs. size xs = n \ set xs \ A}" + +lemma nlistsI: "\ size xs = n; set xs \ A \ \ xs \ nlists n A" + by (simp add: nlists_def) + +lemma nlistsE_length [simp](*rm simp del*): "xs \ nlists n A \ size xs = n" + by (simp add: nlists_def) + +lemma less_lengthI: "\ xs \ nlists n A; p < n \ \ p < size xs" +by (simp) + +lemma nlistsE_set[simp](*rm simp del*): "xs \ nlists n A \ set xs \ A" +unfolding nlists_def by (simp) + +lemma nlists_mono: +assumes "A \ B" shows "nlists n A \ nlists n B" +proof + fix xs assume "xs \ nlists n A" + then obtain size: "size xs = n" and inA: "set xs \ A" by simp + with assms have "set xs \ B" by simp + with size show "xs \ nlists n B" by(clarsimp intro!: nlistsI) +qed + +lemma nlists_n_0 [simp]: "nlists 0 A = {[]}" +unfolding nlists_def by (auto) + +lemma in_nlists_Suc_iff: "(xs \ nlists (Suc n) A) = (\y\A. \ys \ nlists n A. xs = y#ys)" +unfolding nlists_def by (cases "xs") auto + +lemma Cons_in_nlists_Suc [iff]: "(x#xs \ nlists (Suc n) A) \ (x\A \ xs \ nlists n A)" +unfolding nlists_def by (auto) + +lemma nlists_not_empty: "A\{} \ \xs. xs \ nlists n A" +by (induct "n") (auto simp: in_nlists_Suc_iff) + + +lemma nlistsE_nth_in: "\ xs \ nlists n A; i < n \ \ xs!i \ A" +unfolding nlists_def by (auto) + +lemma nlists_Cons_Suc [elim!]: + "l#xs \ nlists n A \ (\n'. n = Suc n' \ l \ A \ xs \ nlists n' A \ P) \ P" +unfolding nlists_def by (auto) + +lemma nlists_appendE [elim!]: + "a@b \ nlists n A \ (\n1 n2. n=n1+n2 \ a \ nlists n1 A \ b \ nlists n2 A \ P) \ P" +proof - + have "\n. a@b \ nlists n A \ \n1 n2. n=n1+n2 \ a \ nlists n1 A \ b \ nlists n2 A" + (is "\n. ?list a n \ \n1 n2. ?P a n n1 n2") + proof (induct a) + fix n assume "?list [] n" + hence "?P [] n 0 n" by simp + thus "\n1 n2. ?P [] n n1 n2" by fast + next + fix n l ls + assume "?list (l#ls) n" + then obtain n' where n: "n = Suc n'" "l \ A" and n': "ls@b \ nlists n' A" by fastforce + assume "\n. ls @ b \ nlists n A \ \n1 n2. n = n1 + n2 \ ls \ nlists n1 A \ b \ nlists n2 A" + from this and n' have "\n1 n2. n' = n1 + n2 \ ls \ nlists n1 A \ b \ nlists n2 A" . + then obtain n1 n2 where "n' = n1 + n2" "ls \ nlists n1 A" "b \ nlists n2 A" by fast + with n have "?P (l#ls) n (n1+1) n2" by simp + thus "\n1 n2. ?P (l#ls) n n1 n2" by fastforce + qed + moreover assume "a@b \ nlists n A" "\n1 n2. n=n1+n2 \ a \ nlists n1 A \ b \ nlists n2 A \ P" + ultimately show ?thesis by blast +qed + + +lemma nlists_update_in_list [simp, intro!]: + "\ xs \ nlists n A; x\A \ \ xs[i := x] \ nlists n A" + by (metis length_list_update nlistsE_length nlistsE_set nlistsI set_update_subsetI) + +lemma nlists_appendI [intro?]: + "\ a \ nlists n A; b \ nlists m A \ \ a @ b \ nlists (n+m) A" +unfolding nlists_def by (auto) + +lemma nlists_append: + "xs @ ys \ nlists k A \ + k = length(xs @ ys) \ xs \ nlists (length xs) A \ ys \ nlists (length ys) A" +unfolding nlists_def by (auto) + +lemma nlists_map [simp]: "(map f xs \ nlists (size xs) A) = (f ` set xs \ A)" +unfolding nlists_def by (auto) + +lemma nlists_replicateI [intro]: "x \ A \ replicate n x \ nlists n A" + by (induct n) auto + +end