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